| author | wenzelm | 
| Thu, 31 May 2001 20:53:49 +0200 | |
| changeset 11356 | 8fbb19b84f94 | 
| parent 10646 | 37b9897dbf3a | 
| child 11701 | 3d51fbf81c17 | 
| 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 | |
| 12 | (*** Abel_Cancel simproc on the integers ***) | |
| 13 | ||
| 14 | (* Lemmas needed for the simprocs *) | |
| 15 | ||
| 16 | (*Deletion of other terms in the formula, seeking the -x at the front of z*) | |
| 17 | Goal "((x::int) + (y + z) = y + u) = ((x + z) = u)"; | |
| 18 | by (stac zadd_left_commute 1); | |
| 19 | by (rtac zadd_left_cancel 1); | |
| 20 | qed "zadd_cancel_21"; | |
| 21 | ||
| 22 | (*A further rule to deal with the case that | |
| 23 | everything gets cancelled on the right.*) | |
| 24 | Goal "((x::int) + (y + z) = y) = (x = -z)"; | |
| 25 | by (stac zadd_left_commute 1); | |
| 26 | by (res_inst_tac [("t", "y")] (zadd_int0_right RS subst) 1
 | |
| 27 | THEN stac zadd_left_cancel 1); | |
| 28 | by (simp_tac (simpset() addsimps [eq_zdiff_eq RS sym]) 1); | |
| 29 | qed "zadd_cancel_end"; | |
| 30 | ||
| 31 | ||
| 32 | structure Int_Cancel_Data = | |
| 33 | struct | |
| 34 | val ss = HOL_ss | |
| 35 | val eq_reflection = eq_reflection | |
| 36 | ||
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
8785diff
changeset | 37 | val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) | 
| 7707 | 38 | val T = HOLogic.intT | 
| 39 |   val zero	= Const ("IntDef.int", HOLogic.natT --> T) $ HOLogic.zero
 | |
| 40 | val restrict_to_left = restrict_to_left | |
| 41 | val add_cancel_21 = zadd_cancel_21 | |
| 42 | val add_cancel_end = zadd_cancel_end | |
| 43 | val add_left_cancel = zadd_left_cancel | |
| 44 | val add_assoc = zadd_assoc | |
| 45 | val add_commute = zadd_commute | |
| 46 | val add_left_commute = zadd_left_commute | |
| 47 | val add_0 = zadd_int0 | |
| 48 | val add_0_right = zadd_int0_right | |
| 49 | ||
| 50 | val eq_diff_eq = eq_zdiff_eq | |
| 51 | val eqI_rules = [zless_eqI, zeq_eqI, zle_eqI] | |
| 52 | fun dest_eqI th = | |
| 53 | #1 (HOLogic.dest_bin "op =" HOLogic.boolT | |
| 54 | (HOLogic.dest_Trueprop (concl_of th))) | |
| 55 | ||
| 56 | val diff_def = zdiff_def | |
| 57 | val minus_add_distrib = zminus_zadd_distrib | |
| 58 | val minus_minus = zminus_zminus | |
| 59 | val minus_0 = zminus_int0 | |
| 60 | val add_inverses = [zadd_zminus_inverse, zadd_zminus_inverse2]; | |
| 61 | val cancel_simps = [zadd_zminus_cancel, zminus_zadd_cancel] | |
| 62 | end; | |
| 63 | ||
| 64 | structure Int_Cancel = Abel_Cancel (Int_Cancel_Data); | |
| 65 | ||
| 66 | Addsimprocs [Int_Cancel.sum_conv, Int_Cancel.rel_conv]; | |
| 67 | ||
| 68 | ||
| 69 | ||
| 70 | (*** misc ***) | |
| 71 | ||
| 8785 | 72 | Goal "- (z - y) = y - (z::int)"; | 
| 73 | by (Simp_tac 1); | |
| 74 | qed "zminus_zdiff_eq"; | |
| 75 | Addsimps [zminus_zdiff_eq]; | |
| 76 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 77 | Goal "(w<z) = neg(w-z)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 78 | 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 | 79 | qed "zless_eq_neg"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 80 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 81 | Goal "(w=z) = iszero(w-z)"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 82 | 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 | 83 | qed "eq_eq_iszero"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 84 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 85 | Goal "(w<=z) = (~ neg(z-w))"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 86 | 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 | 87 | qed "zle_eq_not_neg"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 88 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 89 | |
| 5593 | 90 | (*** Inequality lemmas involving int (Suc m) ***) | 
| 91 | ||
| 92 | Goal "(w < z + int (Suc m)) = (w < z + int m | w = z + int m)"; | |
| 93 | by (auto_tac (claset(), | |
| 94 | simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, zadd_int])); | |
| 6717 
70b251dc7055
int_Suc->int_Suc_int_1 avoiding confusion with the more useful Bin.int_Suc
 paulson parents: 
5593diff
changeset | 95 | by (cut_inst_tac [("m","m")] int_Suc_int_1 1);
 | 
| 
70b251dc7055
int_Suc->int_Suc_int_1 avoiding confusion with the more useful Bin.int_Suc
 paulson parents: 
5593diff
changeset | 96 | by (cut_inst_tac [("m","n")] int_Suc_int_1 1);
 | 
| 5593 | 97 | by (Asm_full_simp_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 98 | by (case_tac "n" 1); | 
| 5593 | 99 | by Auto_tac; | 
| 6717 
70b251dc7055
int_Suc->int_Suc_int_1 avoiding confusion with the more useful Bin.int_Suc
 paulson parents: 
5593diff
changeset | 100 | by (cut_inst_tac [("m","m")] int_Suc_int_1 1);
 | 
| 5593 | 101 | by (full_simp_tac (simpset() addsimps zadd_ac) 1); | 
| 102 | by (asm_full_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1); | |
| 103 | by (auto_tac (claset(), | |
| 104 | simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, zadd_int])); | |
| 105 | qed "zless_add_int_Suc_eq"; | |
| 106 | ||
| 107 | Goal "(w + int (Suc m) <= z) = (w + int m < z)"; | |
| 108 | by (simp_tac (simpset() addsimps [zle_def, zless_add_int_Suc_eq]) 1); | |
| 109 | by (auto_tac (claset() addIs [zle_anti_sym] addEs [zless_asym], | |
| 10646 
37b9897dbf3a
greater use of overloaded rules (order_less_imp_le not zless_imp_zle, ...)
 paulson parents: 
10472diff
changeset | 110 | simpset() addsimps [order_less_imp_le, symmetric zle_def])); | 
| 5593 | 111 | qed "add_int_Suc_zle_eq"; | 
| 112 | ||
| 113 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 114 | (*** Monotonicity results ***) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 115 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 116 | 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 | 117 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 118 | qed "zadd_right_cancel_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 119 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 120 | 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 | 121 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 122 | qed "zadd_left_cancel_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 123 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 124 | 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 | 125 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 126 | 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 | 127 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 128 | qed "zadd_right_cancel_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 129 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 130 | 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 | 131 | by (Simp_tac 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 132 | qed "zadd_left_cancel_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 133 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 134 | 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 | 135 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 136 | (*"v<=w ==> v+z <= w+z"*) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 137 | 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 | 138 | |
| 6998 | 139 | (*"v<=w ==> z+v <= z+w"*) | 
| 140 | bind_thm ("zadd_zless_mono2", zadd_left_cancel_zless RS iffD2);
 | |
| 141 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 142 | (*"v<=w ==> v+z <= w+z"*) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 143 | 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 | 144 | |
| 6998 | 145 | (*"v<=w ==> z+v <= z+w"*) | 
| 146 | bind_thm ("zadd_zle_mono2", zadd_left_cancel_zle RS iffD2);
 | |
| 147 | ||
| 7081 | 148 | 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 | 149 | 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 | 150 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 151 | qed "zadd_zle_mono"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 152 | |
| 7081 | 153 | 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 | 154 | 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 | 155 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 156 | qed "zadd_zless_mono"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 157 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 158 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 159 | (*** Comparison laws ***) | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 160 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 161 | Goal "(- x < - y) = (y < (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 162 | 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 | 163 | qed "zminus_zless_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 164 | Addsimps [zminus_zless_zminus]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 165 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 166 | Goal "(- x <= - y) = (y <= (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 167 | 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 | 168 | qed "zminus_zle_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 169 | Addsimps [zminus_zle_zminus]; | 
| 
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 | (** 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 | 172 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 173 | Goal "(x < - y) = (y < - (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 174 | 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 | 175 | qed "zless_zminus"; | 
| 
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 "(- x < y) = (- y < (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 178 | 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 | 179 | qed "zminus_zless"; | 
| 
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 | Goal "(x <= - y) = (y <= - (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 182 | 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 | 183 | qed "zle_zminus"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 184 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 185 | Goal "(- x <= y) = (- y <= (x::int))"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 186 | 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 | 187 | qed "zminus_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 188 | |
| 6917 | 189 | Goal "(x = - y) = (y = - (x::int))"; | 
| 190 | by Auto_tac; | |
| 191 | qed "equation_zminus"; | |
| 192 | ||
| 193 | Goal "(- x = y) = (- (y::int) = x)"; | |
| 194 | by Auto_tac; | |
| 195 | qed "zminus_equation"; | |
| 196 | ||
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 197 | Goal "- (int (Suc n)) < int 0"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 198 | 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 | 199 | qed "negative_zless_0"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 200 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 201 | 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 | 202 | 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 | 203 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 204 | qed "negative_zless"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 205 | AddIffs [negative_zless]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 206 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 207 | Goal "- int n <= int 0"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 208 | 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 | 209 | qed "negative_zle_0"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 210 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 211 | Goal "- int n <= int m"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 212 | 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 | 213 | qed "negative_zle"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 214 | AddIffs [negative_zle]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 215 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 216 | Goal "~(int 0 <= - (int (Suc n)))"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 217 | by (stac zle_zminus 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 218 | by (Simp_tac 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 219 | qed "not_zle_0_negative"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 220 | Addsimps [not_zle_0_negative]; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 221 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 222 | 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 | 223 | by Safe_tac; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 224 | by (Simp_tac 3); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 225 | by (dtac (zle_zminus RS iffD1) 2); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 226 | 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 | 227 | by (ALLGOALS Asm_full_simp_tac); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 228 | qed "int_zle_neg"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 229 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 230 | Goal "~(int n < - int m)"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 231 | 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 | 232 | qed "not_int_zless_negative"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 233 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 234 | 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 | 235 | by (rtac iffI 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 236 | 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 | 237 | by (dtac sym 1); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 238 | by (ALLGOALS Asm_simp_tac); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 239 | qed "negative_eq_positive"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 240 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 241 | 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 | 242 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 243 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 244 | Goal "(w <= z) = (EX n. z = w + int n)"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 245 | by (auto_tac (claset() addSIs [not_sym RS not0_implies_Suc], | 
| 10472 | 246 | 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 | 247 | qed "zle_iff_zadd"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 248 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 249 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 250 | Goalw [zdiff_def,zless_def] "neg x = (x < int 0)"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 251 | by Auto_tac; | 
| 6917 | 252 | qed "neg_eq_less_int0"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 253 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 254 | Goalw [zle_def] "(~neg x) = (int 0 <= x)"; | 
| 6917 | 255 | by (simp_tac (simpset() addsimps [neg_eq_less_int0]) 1); | 
| 256 | qed "not_neg_eq_ge_int0"; | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 257 | |
| 9945 | 258 | Goal "abs (int m) = int m"; | 
| 259 | by (simp_tac (simpset() addsimps [zabs_def]) 1); | |
| 260 | qed "abs_int_eq"; | |
| 261 | Addsimps [abs_int_eq]; | |
| 262 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 263 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 264 | (**** 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 | 265 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 266 | 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 | 267 | by Auto_tac; | 
| 7009 
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6998diff
changeset | 268 | qed "nat_int"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 269 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 270 | 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 | 271 | by (auto_tac (claset(), | 
| 6917 | 272 | simpset() addsimps [neg_eq_less_int0, zminus_zless])); | 
| 7009 
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6998diff
changeset | 273 | qed "nat_zminus_int"; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 274 | |
| 7009 
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6998diff
changeset | 275 | Addsimps [nat_int, nat_zminus_int]; | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 276 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 277 | Goal "~ neg z ==> int (nat z) = z"; | 
| 6917 | 278 | by (dtac (not_neg_eq_ge_int0 RS iffD1) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 279 | 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 | 280 | 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 | 281 | qed "not_neg_nat"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 282 | |
| 7081 | 283 | 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 | 284 | by (auto_tac (claset(), | 
| 6917 | 285 | simpset() addsimps [neg_eq_less_int0, zless_iff_Suc_zadd, | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 286 | zdiff_eq_eq RS sym, zdiff_def])); | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 287 | qed "negD"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 288 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 289 | 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 | 290 | by Auto_tac; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 291 | qed "neg_nat"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 292 | |
| 7518 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 293 | 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 | 294 | by (case_tac "neg z" 1); | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 295 | 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 | 296 | 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 | 297 | by (auto_tac (claset() addDs [order_less_trans], | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 298 | simpset() addsimps [neg_eq_less_int0])); | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 299 | qed "zless_nat_eq_int_zless"; | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 300 | |
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 301 | Goal "z <= int 0 ==> nat z = 0"; | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 302 | by (auto_tac (claset(), | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 303 | simpset() addsimps [order_le_less, neg_eq_less_int0, | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 304 | zle_def, neg_nat])); | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 305 | qed "nat_le_int0"; | 
| 
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
 paulson parents: 
7081diff
changeset | 306 | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 307 | (*An alternative condition is int 0 <= w *) | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 308 | Goal "int 0 < z ==> (nat w < nat z) = (w < z)"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 309 | by (stac (zless_int RS sym) 1); | 
| 6917 | 310 | by (asm_simp_tac (simpset() addsimps [not_neg_nat, not_neg_eq_ge_int0, | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 311 | order_le_less]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 312 | by (case_tac "neg w" 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 313 | by (asm_simp_tac (simpset() addsimps [not_neg_nat]) 2); | 
| 6917 | 314 | by (asm_full_simp_tac (simpset() addsimps [neg_eq_less_int0, neg_nat]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 315 | by (blast_tac (claset() addIs [order_less_trans]) 1); | 
| 6917 | 316 | val lemma = result(); | 
| 317 | ||
| 318 | Goal "(nat w < nat z) = (int 0 < z & w < z)"; | |
| 319 | by (case_tac "int 0 < z" 1); | |
| 320 | by (auto_tac (claset(), | |
| 7009 
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6998diff
changeset | 321 | simpset() addsimps [lemma, nat_le_int0, linorder_not_less])); | 
| 6917 | 322 | qed "zless_nat_conj"; | 
| 323 | ||
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 324 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 325 | (* 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 | 326 | |
| 6942 | 327 | val prems = Goal | 
| 328 | "[|!! 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 | 329 | by (case_tac "neg z" 1); | 
| 6942 | 330 | by (fast_tac (claset() addSDs [negD] addSEs prems) 1); | 
| 331 | by (dtac (not_neg_nat RS sym) 1); | |
| 332 | by (eresolve_tac prems 1); | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 333 | qed "int_cases"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 334 | |
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: diff
changeset | 335 | 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 | 336 | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 337 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 338 | (*** Monotonicity of Multiplication ***) | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 339 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 340 | Goal "i <= (j::int) ==> i * int k <= j * int k"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 341 | by (induct_tac "k" 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 342 | by (stac int_Suc_int_1 2); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 343 | by (ALLGOALS | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 344 | (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2, zadd_zle_mono]))); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 345 | val lemma = result(); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 346 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 347 | Goal "[| i <= j; int 0 <= k |] ==> i*k <= j*k"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 348 | 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 | 349 | by (etac lemma 2); | 
| 6917 | 350 | by (full_simp_tac (simpset() addsimps [not_neg_eq_ge_int0]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 351 | qed "zmult_zle_mono1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 352 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 353 | Goal "[| i <= j; k <= int 0 |] ==> j*k <= i*k"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 354 | by (rtac (zminus_zle_zminus RS iffD1) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 355 | by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 356 | zmult_zle_mono1, zle_zminus]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 357 | qed "zmult_zle_mono1_neg"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 358 | |
| 6942 | 359 | Goal "[| i <= j; int 0 <= k |] ==> k*i <= k*j"; | 
| 360 | by (dtac zmult_zle_mono1 1); | |
| 361 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); | |
| 362 | qed "zmult_zle_mono2"; | |
| 363 | ||
| 364 | Goal "[| i <= j; k <= int 0 |] ==> k*j <= k*i"; | |
| 365 | by (dtac zmult_zle_mono1_neg 1); | |
| 366 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); | |
| 367 | qed "zmult_zle_mono2_neg"; | |
| 368 | ||
| 6990 | 369 | (* <= monotonicity, BOTH arguments*) | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 370 | Goal "[| i <= j; k <= l; int 0 <= j; int 0 <= k |] ==> i*k <= j*l"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 371 | by (etac (zmult_zle_mono1 RS order_trans) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 372 | by (assume_tac 1); | 
| 6990 | 373 | by (etac zmult_zle_mono2 1); | 
| 374 | by (assume_tac 1); | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 375 | qed "zmult_zle_mono"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 376 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 377 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 378 | (** strict, in 1st argument; proof is by induction on k>0 **) | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 379 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 380 | Goal "i<j ==> 0<k --> int k * i < int k * j"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 381 | by (induct_tac "k" 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 382 | by (stac int_Suc_int_1 2); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 383 | by (case_tac "n=0" 2); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 384 | by (ALLGOALS (asm_full_simp_tac | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 385 | (simpset() addsimps [zadd_zmult_distrib, zadd_zless_mono, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 386 | order_le_less]))); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 387 | val lemma = result() RS mp; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 388 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 389 | Goal "[| i<j; int 0 < k |] ==> k*i < k*j"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 390 | 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 | 391 | by (etac lemma 2); | 
| 6917 | 392 | by (asm_simp_tac (simpset() addsimps [not_neg_eq_ge_int0, | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 393 | order_le_less]) 1); | 
| 6917 | 394 | 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 | 395 | by Auto_tac; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 396 | qed "zmult_zless_mono2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 397 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 398 | Goal "[| i<j; int 0 < k |] ==> i*k < j*k"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 399 | by (dtac zmult_zless_mono2 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 400 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 401 | qed "zmult_zless_mono1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 402 | |
| 6990 | 403 | (* < monotonicity, BOTH arguments*) | 
| 404 | Goal "[| i < j; k < l; int 0 < j; int 0 < k |] ==> i*k < j*l"; | |
| 405 | by (etac (zmult_zless_mono1 RS order_less_trans) 1); | |
| 406 | by (assume_tac 1); | |
| 407 | by (etac zmult_zless_mono2 1); | |
| 408 | by (assume_tac 1); | |
| 409 | qed "zmult_zless_mono"; | |
| 410 | ||
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 411 | Goal "[| i<j; k < int 0 |] ==> j*k < i*k"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 412 | by (rtac (zminus_zless_zminus RS iffD1) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 413 | by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 414 | zmult_zless_mono1, zless_zminus]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 415 | qed "zmult_zless_mono1_neg"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 416 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 417 | Goal "[| i<j; k < int 0 |] ==> k*j < k*i"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 418 | by (rtac (zminus_zless_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 RS sym, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 420 | zmult_zless_mono2, zless_zminus]) 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 421 | qed "zmult_zless_mono2_neg"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 422 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 423 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 424 | Goal "(m*n = int 0) = (m = int 0 | n = int 0)"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 425 | by (case_tac "m < int 0" 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 426 | by (auto_tac (claset(), | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 427 | simpset() addsimps [linorder_not_less, order_le_less, | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 428 | linorder_neq_iff])); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 429 | by (REPEAT | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 430 | (force_tac (claset() addDs [zmult_zless_mono1_neg, zmult_zless_mono1], | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 431 | simpset()) 1)); | 
| 6990 | 432 | qed "zmult_eq_int0_iff"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 433 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 434 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 435 | (** 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 | 436 | but not (yet?) for k*m < n*k. **) | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 437 | |
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 438 | Goal "(m*k < n*k) = ((int 0 < k & m<n) | (k < int 0 & n<m))"; | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 439 | by (case_tac "k = int 0" 1); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 440 | by (auto_tac (claset(), simpset() addsimps [linorder_neq_iff, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 441 | zmult_zless_mono1, zmult_zless_mono1_neg])); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 442 | by (auto_tac (claset(), | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 443 | simpset() addsimps [linorder_not_less, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 444 | inst "y1" "m*k" (linorder_not_le RS sym), | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 445 | inst "y1" "m" (linorder_not_le RS sym)])); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 446 | by (ALLGOALS (etac notE)); | 
| 10646 
37b9897dbf3a
greater use of overloaded rules (order_less_imp_le not zless_imp_zle, ...)
 paulson parents: 
10472diff
changeset | 447 | 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 | 448 | zmult_zle_mono1_neg])); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 449 | qed "zmult_zless_cancel2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 450 | |
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 451 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 452 | Goal "(k*m < k*n) = ((int 0 < k & m<n) | (k < int 0 & n<m))"; | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 453 | by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 454 | zmult_zless_cancel2]) 1); | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 455 | qed "zmult_zless_cancel1"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 456 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 457 | Goal "(m*k <= n*k) = ((int 0 < k --> m<=n) & (k < int 0 --> n<=m))"; | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 458 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 459 | zmult_zless_cancel2]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 460 | qed "zmult_zle_cancel2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 461 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 462 | Goal "(k*m <= k*n) = ((int 0 < k --> m<=n) & (k < int 0 --> n<=m))"; | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 463 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 464 | zmult_zless_cancel1]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 465 | qed "zmult_zle_cancel1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 466 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 467 | Goal "(m*k = n*k) = (k = int 0 | m=n)"; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 468 | by (cut_facts_tac [linorder_less_linear] 1); | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 469 | by Safe_tac; | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 470 | by Auto_tac; | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 471 | by (REPEAT | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 472 |     (force_tac (claset() addD2 ("mono_neg", zmult_zless_mono1_neg)
 | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 473 |                          addD2 ("mono_pos", zmult_zless_mono1), 
 | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 474 | simpset() addsimps [linorder_neq_iff]) 1)); | 
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 475 | |
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 476 | qed "zmult_cancel2"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 477 | |
| 9633 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 478 | Goal "(k*m = k*n) = (k = int 0 | m=n)"; | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 479 | by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, | 
| 
a71a83253997
better rules for cancellation of common factors across comparisons
 paulson parents: 
9582diff
changeset | 480 | zmult_cancel2]) 1); | 
| 6866 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 481 | qed "zmult_cancel1"; | 
| 
f795b63139ec
many new theorems concerning multiplication and (in)equations
 paulson parents: 
6717diff
changeset | 482 | Addsimps [zmult_cancel1, zmult_cancel2]; |