| author | wenzelm | 
| Mon, 04 Mar 2002 19:06:52 +0100 | |
| changeset 13012 | f8bfc61ee1b5 | 
| parent 12613 | 279facb4253a | 
| permissions | -rw-r--r-- | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 1 | (* Title: HOL/Integ/Int.ML | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 4 | Copyright 1998 University of Cambridge | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 5 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 6 | Type "int" is a linear order | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 7 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 8 | And many further lemmas | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 9 | *) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 10 | |
| 7707 | 11 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 12 | Goal "int 0 = (0::int)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 13 | 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: 
11701diff
changeset | 14 | qed "int_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 15 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 16 | Goal "int 1 = 1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 17 | by (simp_tac (simpset() addsimps [One_int_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 18 | qed "int_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 19 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 20 | Goal "int (Suc 0) = 1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 21 | by (simp_tac (simpset() addsimps [One_int_def, One_nat_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 22 | qed "int_Suc0_eq_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 23 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 24 | Goalw [zdiff_def,zless_def] "neg x = (x < 0)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 25 | by Auto_tac; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 26 | qed "neg_eq_less_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 27 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 28 | Goalw [zle_def] "(~neg x) = (0 <= x)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 29 | by (simp_tac (simpset() addsimps [neg_eq_less_0]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 30 | qed "not_neg_eq_ge_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 31 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 32 | (** Needed to simplify inequalities when Numeral1 can get simplified to 1 **) | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 33 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 34 | Goal "~ neg 0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 35 | by (simp_tac (simpset() addsimps [One_int_def, neg_eq_less_0]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 36 | qed "not_neg_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 37 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 38 | Goal "~ neg 1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 39 | by (simp_tac (simpset() addsimps [One_int_def, neg_eq_less_0]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 40 | qed "not_neg_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 41 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 42 | Goal "iszero 0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 43 | by (simp_tac (simpset() addsimps [iszero_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 44 | qed "iszero_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 45 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 46 | Goal "~ iszero 1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 47 | by (simp_tac (simpset() addsimps [Zero_int_def, One_int_def, One_nat_def, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 48 | iszero_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 49 | qed "not_iszero_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 50 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 51 | Goal "0 < (1::int)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 52 | by (simp_tac (simpset() addsimps [Zero_int_def, One_int_def, One_nat_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 53 | qed "int_0_less_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 54 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 55 | Goal "0 \\<noteq> (1::int)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 56 | by (simp_tac (simpset() addsimps [Zero_int_def, One_int_def, One_nat_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 57 | qed "int_0_neq_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 58 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 59 | Addsimps [int_0, int_1, int_0_neq_1]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 60 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 61 | |
| 7707 | 62 | (*** Abel_Cancel simproc on the integers ***) | 
| 63 | ||
| 64 | (* Lemmas needed for the simprocs *) | |
| 65 | ||
| 66 | (*Deletion of other terms in the formula, seeking the -x at the front of z*) | |
| 67 | Goal "((x::int) + (y + z) = y + u) = ((x + z) = u)"; | |
| 68 | by (stac zadd_left_commute 1); | |
| 69 | by (rtac zadd_left_cancel 1); | |
| 70 | qed "zadd_cancel_21"; | |
| 71 | ||
| 72 | (*A further rule to deal with the case that | |
| 73 | everything gets cancelled on the right.*) | |
| 74 | Goal "((x::int) + (y + z) = y) = (x = -z)"; | |
| 75 | by (stac zadd_left_commute 1); | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 76 | by (res_inst_tac [("t", "y")] (zadd_0_right RS subst) 1
 | 
| 7707 | 77 | THEN stac zadd_left_cancel 1); | 
| 78 | by (simp_tac (simpset() addsimps [eq_zdiff_eq RS sym]) 1); | |
| 79 | qed "zadd_cancel_end"; | |
| 80 | ||
| 81 | ||
| 82 | structure Int_Cancel_Data = | |
| 83 | struct | |
| 84 | val ss = HOL_ss | |
| 85 | val eq_reflection = eq_reflection | |
| 86 | ||
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
8785diff
changeset | 87 | val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) | 
| 7707 | 88 | val T = HOLogic.intT | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 89 |   val zero		= Const ("0", HOLogic.intT)
 | 
| 7707 | 90 | val restrict_to_left = restrict_to_left | 
| 91 | val add_cancel_21 = zadd_cancel_21 | |
| 92 | val add_cancel_end = zadd_cancel_end | |
| 93 | val add_left_cancel = zadd_left_cancel | |
| 94 | val add_assoc = zadd_assoc | |
| 95 | val add_commute = zadd_commute | |
| 96 | val add_left_commute = zadd_left_commute | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 97 | val add_0 = zadd_0 | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 98 | val add_0_right = zadd_0_right | 
| 7707 | 99 | |
| 100 | val eq_diff_eq = eq_zdiff_eq | |
| 101 | val eqI_rules = [zless_eqI, zeq_eqI, zle_eqI] | |
| 102 | fun dest_eqI th = | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
10646diff
changeset | 103 | #1 (HOLogic.dest_bin "op =" HOLogic.boolT | 
| 7707 | 104 | (HOLogic.dest_Trueprop (concl_of th))) | 
| 105 | ||
| 106 | val diff_def = zdiff_def | |
| 107 | val minus_add_distrib = zminus_zadd_distrib | |
| 108 | val minus_minus = zminus_zminus | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 109 | val minus_0 = zminus_0 | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 110 | val add_inverses = [zadd_zminus_inverse, zadd_zminus_inverse2] | 
| 7707 | 111 | val cancel_simps = [zadd_zminus_cancel, zminus_zadd_cancel] | 
| 112 | end; | |
| 113 | ||
| 114 | structure Int_Cancel = Abel_Cancel (Int_Cancel_Data); | |
| 115 | ||
| 116 | Addsimprocs [Int_Cancel.sum_conv, Int_Cancel.rel_conv]; | |
| 117 | ||
| 118 | ||
| 119 | ||
| 120 | (*** misc ***) | |
| 121 | ||
| 8785 | 122 | Goal "- (z - y) = y - (z::int)"; | 
| 123 | by (Simp_tac 1); | |
| 124 | qed "zminus_zdiff_eq"; | |
| 125 | Addsimps [zminus_zdiff_eq]; | |
| 126 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 127 | Goal "(w<z) = neg(w-z)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 128 | by (simp_tac (simpset() addsimps [zless_def]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 129 | qed "zless_eq_neg"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 130 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 131 | Goal "(w=z) = iszero(w-z)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 132 | by (simp_tac (simpset() addsimps [iszero_def, zdiff_eq_eq]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 133 | qed "eq_eq_iszero"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 134 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 135 | Goal "(w<=z) = (~ neg(z-w))"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 136 | by (simp_tac (simpset() addsimps [zle_def, zless_def]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 137 | qed "zle_eq_not_neg"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 138 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 139 | (** Inequality reasoning **) | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 140 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 141 | Goal "(w < z + (1::int)) = (w<z | w=z)"; | 
| 5593 | 142 | by (auto_tac (claset(), | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 143 | simpset() addsimps [zless_iff_Suc_zadd, int_Suc, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 144 | gr0_conv_Suc, zero_reorient])); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 145 | by (res_inst_tac [("x","Suc n")] exI 1); 
 | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 146 | by (simp_tac (simpset() addsimps [int_Suc]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 147 | qed "zless_add1_eq"; | 
| 5593 | 148 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 149 | Goal "(w + (1::int) <= z) = (w<z)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 150 | by (asm_full_simp_tac (simpset() addsimps [zle_def, zless_add1_eq]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 151 | by (auto_tac (claset() addIs [zle_anti_sym], | 
| 10646 
37b9897dbf3a
greater use of overloaded rules (order_less_imp_le not zless_imp_zle, ...)
 paulson parents: 
10472diff
changeset | 152 | simpset() addsimps [order_less_imp_le, symmetric zle_def])); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 153 | qed "add1_zle_eq"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 154 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 155 | Goal "((1::int) + w <= z) = (w<z)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 156 | by (stac zadd_commute 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 157 | by (rtac add1_zle_eq 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 158 | qed "add1_left_zle_eq"; | 
| 5593 | 159 | |
| 160 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 161 | (*** Monotonicity results ***) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 162 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 163 | Goal "(v+z < w+z) = (v < (w::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 164 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 165 | qed "zadd_right_cancel_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 166 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 167 | Goal "(z+v < z+w) = (v < (w::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 168 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 169 | qed "zadd_left_cancel_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 170 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 171 | Addsimps [zadd_right_cancel_zless, zadd_left_cancel_zless]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 172 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 173 | Goal "(v+z <= w+z) = (v <= (w::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 174 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 175 | qed "zadd_right_cancel_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 176 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 177 | Goal "(z+v <= z+w) = (v <= (w::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 178 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 179 | qed "zadd_left_cancel_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 180 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 181 | Addsimps [zadd_right_cancel_zle, zadd_left_cancel_zle]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 182 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 183 | (*"v<=w ==> v+z <= w+z"*) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 184 | bind_thm ("zadd_zless_mono1", zadd_right_cancel_zless RS iffD2);
 | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 185 | |
| 6998 | 186 | (*"v<=w ==> z+v <= z+w"*) | 
| 187 | bind_thm ("zadd_zless_mono2", zadd_left_cancel_zless RS iffD2);
 | |
| 188 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 189 | (*"v<=w ==> v+z <= w+z"*) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 190 | bind_thm ("zadd_zle_mono1", zadd_right_cancel_zle RS iffD2);
 | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 191 | |
| 6998 | 192 | (*"v<=w ==> z+v <= z+w"*) | 
| 193 | bind_thm ("zadd_zle_mono2", zadd_left_cancel_zle RS iffD2);
 | |
| 194 | ||
| 7081 | 195 | Goal "[| w'<=w; z'<=z |] ==> w' + z' <= w + (z::int)"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 196 | by (etac (zadd_zle_mono1 RS zle_trans) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 197 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 198 | qed "zadd_zle_mono"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 199 | |
| 7081 | 200 | Goal "[| w'<w; z'<=z |] ==> w' + z' < w + (z::int)"; | 
| 10646 
37b9897dbf3a
greater use of overloaded rules (order_less_imp_le not zless_imp_zle, ...)
 paulson parents: 
10472diff
changeset | 201 | by (etac (zadd_zless_mono1 RS order_less_le_trans) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 202 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 203 | qed "zadd_zless_mono"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 204 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 205 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 206 | (*** Comparison laws ***) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 207 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 208 | Goal "(- x < - y) = (y < (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 209 | by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 210 | qed "zminus_zless_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 211 | Addsimps [zminus_zless_zminus]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 212 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 213 | Goal "(- x <= - y) = (y <= (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 214 | by (simp_tac (simpset() addsimps [zle_def]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 215 | qed "zminus_zle_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 216 | Addsimps [zminus_zle_zminus]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 217 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 218 | (** The next several equations can make the simplifier loop! **) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 219 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 220 | Goal "(x < - y) = (y < - (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 221 | by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 222 | qed "zless_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 223 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 224 | Goal "(- x < y) = (- y < (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 225 | by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 226 | qed "zminus_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 227 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 228 | Goal "(x <= - y) = (y <= - (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 229 | by (simp_tac (simpset() addsimps [zle_def, zminus_zless]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 230 | qed "zle_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 231 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 232 | Goal "(- x <= y) = (- y <= (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 233 | by (simp_tac (simpset() addsimps [zle_def, zless_zminus]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 234 | qed "zminus_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 235 | |
| 6917 | 236 | Goal "(x = - y) = (y = - (x::int))"; | 
| 237 | by Auto_tac; | |
| 238 | qed "equation_zminus"; | |
| 239 | ||
| 240 | Goal "(- x = y) = (- (y::int) = x)"; | |
| 241 | by Auto_tac; | |
| 242 | qed "zminus_equation"; | |
| 243 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 244 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 245 | (** Instances of the equations above, for zero **) | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 246 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 247 | (*instantiate a variable to zero and simplify*) | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 248 | fun zero_instance v th = simplify (simpset()) (inst v "0" th); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 249 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 250 | Addsimps [zero_instance "x" zless_zminus, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 251 | zero_instance "y" zminus_zless, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 252 | zero_instance "x" zle_zminus, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 253 | zero_instance "y" zminus_zle, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 254 | zero_instance "x" equation_zminus, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 255 | zero_instance "y" zminus_equation]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 256 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 257 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 258 | Goal "- (int (Suc n)) < 0"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 259 | by (simp_tac (simpset() addsimps [zless_def]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 260 | qed "negative_zless_0"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 261 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 262 | Goal "- (int (Suc n)) < int m"; | 
| 10646 
37b9897dbf3a
greater use of overloaded rules (order_less_imp_le not zless_imp_zle, ...)
 paulson parents: 
10472diff
changeset | 263 | by (rtac (negative_zless_0 RS order_less_le_trans) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 264 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 265 | qed "negative_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 266 | AddIffs [negative_zless]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 267 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 268 | Goal "- int n <= 0"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 269 | by (simp_tac (simpset() addsimps [zminus_zle]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 270 | qed "negative_zle_0"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 271 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 272 | Goal "- int n <= int m"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 273 | by (simp_tac (simpset() addsimps [zless_def, zle_def, zdiff_def, zadd_int]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 274 | qed "negative_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 275 | AddIffs [negative_zle]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 276 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 277 | Goal "~(0 <= - (int (Suc n)))"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 278 | by (stac zle_zminus 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 279 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 280 | qed "not_zle_0_negative"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 281 | Addsimps [not_zle_0_negative]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 282 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 283 | Goal "(int n <= - int m) = (n = 0 & m = 0)"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 284 | by Safe_tac; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 285 | by (Simp_tac 3); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 286 | by (dtac (zle_zminus RS iffD1) 2); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 287 | by (ALLGOALS (dtac (negative_zle_0 RSN(2,zle_trans)))); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 288 | by (ALLGOALS Asm_full_simp_tac); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 289 | qed "int_zle_neg"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 290 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 291 | Goal "~(int n < - int m)"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 292 | by (simp_tac (simpset() addsimps [symmetric zle_def]) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 293 | qed "not_int_zless_negative"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 294 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 295 | Goal "(- int n = int m) = (n = 0 & m = 0)"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 296 | by (rtac iffI 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 297 | by (rtac (int_zle_neg RS iffD1) 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 298 | by (dtac sym 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 299 | by (ALLGOALS Asm_simp_tac); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 300 | qed "negative_eq_positive"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 301 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 302 | Addsimps [negative_eq_positive, not_int_zless_negative]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 303 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 304 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 305 | Goal "(w <= z) = (EX n. z = w + int n)"; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 306 | by (auto_tac (claset() addIs [inst "x" "0::nat" exI] | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 307 | addSIs [not_sym RS not0_implies_Suc], | 
| 10472 | 308 | simpset() addsimps [zless_iff_Suc_zadd, int_le_less])); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 309 | qed "zle_iff_zadd"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 310 | |
| 9945 | 311 | Goal "abs (int m) = int m"; | 
| 312 | by (simp_tac (simpset() addsimps [zabs_def]) 1); | |
| 313 | qed "abs_int_eq"; | |
| 314 | Addsimps [abs_int_eq]; | |
| 315 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 316 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 317 | (**** nat: magnitide of an integer, as a natural number ****) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 318 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 319 | Goalw [nat_def] "nat(int n) = n"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 320 | by Auto_tac; | 
| 7009 
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6998diff
changeset | 321 | qed "nat_int"; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 322 | Addsimps [nat_int]; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 323 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 324 | Goalw [nat_def] "nat(- (int n)) = 0"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 325 | by (auto_tac (claset(), | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 326 | simpset() addsimps [neg_eq_less_0, zero_reorient, zminus_zless])); | 
| 7009 
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6998diff
changeset | 327 | qed "nat_zminus_int"; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 328 | Addsimps [nat_zminus_int]; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 329 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 330 | Goalw [Zero_int_def] "nat 0 = 0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 331 | by (rtac nat_int 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 332 | qed "nat_zero"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 333 | Addsimps [nat_zero]; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 334 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 335 | Goal "~ neg z ==> int (nat z) = z"; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 336 | by (dtac (not_neg_eq_ge_0 RS iffD1) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 337 | by (dtac zle_imp_zless_or_eq 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 338 | by (auto_tac (claset(), simpset() addsimps [zless_iff_Suc_zadd])); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 339 | qed "not_neg_nat"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 340 | |
| 7081 | 341 | Goal "neg x ==> EX n. x = - (int (Suc n))"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 342 | by (auto_tac (claset(), | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 343 | simpset() addsimps [neg_eq_less_0, zless_iff_Suc_zadd, | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 344 | zdiff_eq_eq RS sym, zdiff_def])); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 345 | qed "negD"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 346 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 347 | Goalw [nat_def] "neg z ==> nat z = 0"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 348 | by Auto_tac; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 349 | qed "neg_nat"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 350 | |
| 7518 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 351 | Goal "(m < nat z) = (int m < z)"; | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 352 | by (case_tac "neg z" 1); | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 353 | by (etac (not_neg_nat RS subst) 2); | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 354 | by (auto_tac (claset(), simpset() addsimps [neg_nat])); | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 355 | by (auto_tac (claset() addDs [order_less_trans], | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 356 | simpset() addsimps [neg_eq_less_0])); | 
| 7518 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 357 | qed "zless_nat_eq_int_zless"; | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 358 | |
| 12613 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 359 | Goal "0 <= z ==> int (nat z) = z"; | 
| 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 360 | by (asm_full_simp_tac | 
| 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 361 | (simpset() addsimps [neg_eq_less_0, zle_def, not_neg_nat]) 1); | 
| 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 362 | qed "nat_0_le"; | 
| 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 363 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 364 | Goal "z <= 0 ==> nat z = 0"; | 
| 7518 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 365 | by (auto_tac (claset(), | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 366 | simpset() addsimps [order_le_less, neg_eq_less_0, | 
| 7518 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 367 | zle_def, neg_nat])); | 
| 12613 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 368 | qed "nat_le_0"; | 
| 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 369 | Addsimps [nat_0_le, nat_le_0]; | 
| 7518 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 370 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 371 | (*An alternative condition is 0 <= w *) | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 372 | Goal "0 < z ==> (nat w < nat z) = (w < z)"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 373 | by (stac (zless_int RS sym) 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 374 | by (asm_simp_tac (simpset() addsimps [not_neg_nat, not_neg_eq_ge_0, | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 375 | order_le_less]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 376 | by (case_tac "neg w" 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 377 | by (asm_simp_tac (simpset() addsimps [not_neg_nat]) 2); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 378 | by (asm_full_simp_tac (simpset() addsimps [neg_eq_less_0, neg_nat]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 379 | by (blast_tac (claset() addIs [order_less_trans]) 1); | 
| 6917 | 380 | val lemma = result(); | 
| 381 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 382 | Goal "(nat w < nat z) = (0 < z & w < z)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 383 | by (case_tac "0 < z" 1); | 
| 12613 
279facb4253a
Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
 paulson parents: 
11868diff
changeset | 384 | by (auto_tac (claset(), simpset() addsimps [lemma, linorder_not_less])); | 
| 6917 | 385 | qed "zless_nat_conj"; | 
| 386 | ||
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 387 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 388 | (* a case theorem distinguishing non-negative and negative int *) | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 389 | |
| 6942 | 390 | val prems = Goal | 
| 391 | "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P"; | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 392 | by (case_tac "neg z" 1); | 
| 6942 | 393 | by (fast_tac (claset() addSDs [negD] addSEs prems) 1); | 
| 394 | by (dtac (not_neg_nat RS sym) 1); | |
| 395 | by (eresolve_tac prems 1); | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 396 | qed "int_cases"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 397 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 398 | fun int_case_tac x = res_inst_tac [("z",x)] int_cases; 
 | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 399 | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 400 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 401 | (*** Monotonicity of Multiplication ***) | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 402 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 403 | Goal "i <= (j::int) ==> i * int k <= j * int k"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 404 | by (induct_tac "k" 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 405 | by (stac int_Suc 2); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 406 | by (ALLGOALS | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 407 | (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2, zadd_zle_mono, | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 408 | int_Suc0_eq_1]))); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 409 | val lemma = result(); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 410 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 411 | Goal "[| i <= j; (0::int) <= k |] ==> i*k <= j*k"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 412 | by (res_inst_tac [("t", "k")] (not_neg_nat RS subst) 1);
 | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 413 | by (etac lemma 2); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 414 | by (full_simp_tac (simpset() addsimps [not_neg_eq_ge_0]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 415 | qed "zmult_zle_mono1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 416 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 417 | Goal "[| i <= j; k <= (0::int) |] ==> j*k <= i*k"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 418 | by (rtac (zminus_zle_zminus RS iffD1) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 419 | by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 420 | zmult_zle_mono1, zle_zminus]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 421 | qed "zmult_zle_mono1_neg"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 422 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 423 | Goal "[| i <= j; (0::int) <= k |] ==> k*i <= k*j"; | 
| 6942 | 424 | by (dtac zmult_zle_mono1 1); | 
| 425 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); | |
| 426 | qed "zmult_zle_mono2"; | |
| 427 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 428 | Goal "[| i <= j; k <= (0::int) |] ==> k*j <= k*i"; | 
| 6942 | 429 | by (dtac zmult_zle_mono1_neg 1); | 
| 430 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); | |
| 431 | qed "zmult_zle_mono2_neg"; | |
| 432 | ||
| 6990 | 433 | (* <= monotonicity, BOTH arguments*) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 434 | Goal "[| i <= j; k <= l; (0::int) <= j; (0::int) <= k |] ==> i*k <= j*l"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 435 | by (etac (zmult_zle_mono1 RS order_trans) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 436 | by (assume_tac 1); | 
| 6990 | 437 | by (etac zmult_zle_mono2 1); | 
| 438 | by (assume_tac 1); | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 439 | qed "zmult_zle_mono"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 440 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 441 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 442 | (** strict, in 1st argument; proof is by induction on k>0 **) | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 443 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 444 | Goal "i<j ==> 0<k --> int k * i < int k * j"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 445 | by (induct_tac "k" 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 446 | by (stac int_Suc 2); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 447 | by (case_tac "n=0" 2); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 448 | by (ALLGOALS (asm_full_simp_tac | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 449 | (simpset() addsimps [zadd_zmult_distrib, zadd_zless_mono, | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 450 | int_Suc0_eq_1, order_le_less]))); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 451 | val lemma = result(); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 452 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 453 | Goal "[| i<j; (0::int) < k |] ==> k*i < k*j"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 454 | by (res_inst_tac [("t", "k")] (not_neg_nat RS subst) 1);
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 455 | by (etac (lemma RS mp) 2); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 456 | by (asm_simp_tac (simpset() addsimps [not_neg_eq_ge_0, | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 457 | order_le_less]) 1); | 
| 6917 | 458 | by (forward_tac [conjI RS (zless_nat_conj RS iffD2)] 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 459 | by Auto_tac; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 460 | qed "zmult_zless_mono2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 461 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 462 | Goal "[| i<j; (0::int) < k |] ==> i*k < j*k"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 463 | by (dtac zmult_zless_mono2 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 464 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 465 | qed "zmult_zless_mono1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 466 | |
| 6990 | 467 | (* < monotonicity, BOTH arguments*) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 468 | Goal "[| i < j; k < l; (0::int) < j; (0::int) < k |] ==> i*k < j*l"; | 
| 6990 | 469 | by (etac (zmult_zless_mono1 RS order_less_trans) 1); | 
| 470 | by (assume_tac 1); | |
| 471 | by (etac zmult_zless_mono2 1); | |
| 472 | by (assume_tac 1); | |
| 473 | qed "zmult_zless_mono"; | |
| 474 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 475 | Goal "[| i<j; k < (0::int) |] ==> j*k < i*k"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 476 | by (rtac (zminus_zless_zminus RS iffD1) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 477 | by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 478 | zmult_zless_mono1, zless_zminus]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 479 | qed "zmult_zless_mono1_neg"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 480 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 481 | Goal "[| i<j; k < (0::int) |] ==> k*j < k*i"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 482 | by (rtac (zminus_zless_zminus RS iffD1) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 483 | by (asm_simp_tac (simpset() addsimps [zmult_zminus RS sym, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 484 | zmult_zless_mono2, zless_zminus]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 485 | qed "zmult_zless_mono2_neg"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 486 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 487 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 488 | Goal "(m*n = (0::int)) = (m = 0 | n = 0)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 489 | by (case_tac "m < (0::int)" 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 490 | by (auto_tac (claset(), | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 491 | simpset() addsimps [linorder_not_less, order_le_less, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 492 | linorder_neq_iff])); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 493 | by (REPEAT | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 494 | (force_tac (claset() addDs [zmult_zless_mono1_neg, zmult_zless_mono1], | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 495 | simpset()) 1)); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 496 | qed "zmult_eq_0_iff"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 497 | AddIffs [zmult_eq_0_iff]; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 498 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 499 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 500 | (** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 501 | but not (yet?) for k*m < n*k. **) | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 502 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 503 | Goal "(m*k < n*k) = (((0::int) < k & m<n) | (k < 0 & n<m))"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 504 | by (case_tac "k = (0::int)" 1); | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 505 | by (auto_tac (claset(), simpset() addsimps [linorder_neq_iff, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 506 | zmult_zless_mono1, zmult_zless_mono1_neg])); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 507 | by (auto_tac (claset(), | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 508 | simpset() addsimps [linorder_not_less, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 509 | inst "y1" "m*k" (linorder_not_le RS sym), | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 510 | inst "y1" "m" (linorder_not_le RS sym)])); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 511 | by (ALLGOALS (etac notE)); | 
| 10646 
37b9897dbf3a
greater use of overloaded rules (order_less_imp_le not zless_imp_zle, ...)
 paulson parents: 
10472diff
changeset | 512 | by (auto_tac (claset(), simpset() addsimps [order_less_imp_le, zmult_zle_mono1, | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 513 | zmult_zle_mono1_neg])); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 514 | qed "zmult_zless_cancel2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 515 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 516 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 517 | Goal "(k*m < k*n) = (((0::int) < k & m<n) | (k < 0 & n<m))"; | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 518 | by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 519 | zmult_zless_cancel2]) 1); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 520 | qed "zmult_zless_cancel1"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 521 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 522 | Goal "(m*k <= n*k) = (((0::int) < k --> m<=n) & (k < 0 --> n<=m))"; | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 523 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 524 | zmult_zless_cancel2]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 525 | qed "zmult_zle_cancel2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 526 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 527 | Goal "(k*m <= k*n) = (((0::int) < k --> m<=n) & (k < 0 --> n<=m))"; | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 528 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 529 | zmult_zless_cancel1]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 530 | qed "zmult_zle_cancel1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 531 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 532 | Goal "(m*k = n*k) = (k = (0::int) | m=n)"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 533 | by (cut_facts_tac [linorder_less_linear] 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 534 | by Safe_tac; | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 535 | by Auto_tac; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 536 | by (REPEAT | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 537 |     (force_tac (claset() addD2 ("mono_neg", zmult_zless_mono1_neg)
 | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 538 |                          addD2 ("mono_pos", zmult_zless_mono1), 
 | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 539 | simpset() addsimps [linorder_neq_iff]) 1)); | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 540 | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 541 | qed "zmult_cancel2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 542 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 543 | Goal "(k*m = k*n) = (k = (0::int) | m=n)"; | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 544 | by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 545 | zmult_cancel2]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 546 | qed "zmult_cancel1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 547 | Addsimps [zmult_cancel1, zmult_cancel2]; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 548 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 549 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 550 | (*Analogous to zadd_int*) | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 551 | Goal "n<=m --> int m - int n = int (m-n)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 552 | by (induct_thm_tac diff_induct "m n" 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 553 | by (auto_tac (claset(), simpset() addsimps [int_Suc, symmetric zdiff_def])); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 554 | qed_spec_mp "zdiff_int"; |