author | wenzelm |
Mon, 29 Nov 1999 15:52:49 +0100 | |
changeset 8039 | a901bafe4578 |
parent 7707 | 1f4b67fdfdae |
child 8423 | 3c19160b6432 |
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:
6717
diff
changeset
|
7 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
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 |
||
37 |
val thy = IntDef.thy |
|
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 |
||
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
72 |
Goal "(w<z) = neg(w-z)"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
73 |
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
|
74 |
qed "zless_eq_neg"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
75 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
76 |
Goal "(w=z) = iszero(w-z)"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
77 |
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
|
78 |
qed "eq_eq_iszero"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
79 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
80 |
Goal "(w<=z) = (~ neg(z-w))"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
81 |
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
|
82 |
qed "zle_eq_not_neg"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
83 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
84 |
|
5593 | 85 |
(*** Inequality lemmas involving int (Suc m) ***) |
86 |
||
87 |
Goal "(w < z + int (Suc m)) = (w < z + int m | w = z + int m)"; |
|
88 |
by (auto_tac (claset(), |
|
89 |
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:
5593
diff
changeset
|
90 |
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:
5593
diff
changeset
|
91 |
by (cut_inst_tac [("m","n")] int_Suc_int_1 1); |
5593 | 92 |
by (Asm_full_simp_tac 1); |
93 |
by (exhaust_tac "n" 1); |
|
94 |
by Auto_tac; |
|
6717
70b251dc7055
int_Suc->int_Suc_int_1 avoiding confusion with the more useful Bin.int_Suc
paulson
parents:
5593
diff
changeset
|
95 |
by (cut_inst_tac [("m","m")] int_Suc_int_1 1); |
5593 | 96 |
by (full_simp_tac (simpset() addsimps zadd_ac) 1); |
97 |
by (asm_full_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1); |
|
98 |
by (auto_tac (claset(), |
|
99 |
simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, zadd_int])); |
|
100 |
qed "zless_add_int_Suc_eq"; |
|
101 |
||
102 |
||
103 |
Goal "(w + int (Suc m) <= z) = (w + int m < z)"; |
|
104 |
by (simp_tac (simpset() addsimps [zle_def, zless_add_int_Suc_eq]) 1); |
|
105 |
by (auto_tac (claset() addIs [zle_anti_sym] addEs [zless_asym], |
|
106 |
simpset() addsimps [zless_imp_zle, symmetric zle_def])); |
|
107 |
qed "add_int_Suc_zle_eq"; |
|
108 |
||
109 |
||
110 |
(* (w < int (Suc m)) = (w < int m | w = int m) *) |
|
111 |
bind_thm ("less_int_Suc_eq", |
|
112 |
simplify (simpset()) |
|
113 |
(read_instantiate [("z", "int 0")] zless_add_int_Suc_eq)); |
|
114 |
||
115 |
Goal "(w <= int (Suc m)) = (w <= int m | w = int (Suc m))"; |
|
116 |
by (simp_tac (simpset() addsimps [less_int_Suc_eq, order_le_less]) 1); |
|
117 |
qed "le_int_Suc_eq"; |
|
118 |
||
119 |
||
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
120 |
(*** Monotonicity results ***) |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
121 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
122 |
Goal "(v+z < w+z) = (v < (w::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
123 |
by (Simp_tac 1); |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
124 |
qed "zadd_right_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 "(z+v < z+w) = (v < (w::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
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_left_cancel_zless"; |
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 |
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
|
131 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
132 |
Goal "(v+z <= w+z) = (v <= (w::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
133 |
by (Simp_tac 1); |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
134 |
qed "zadd_right_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 |
Goal "(z+v <= z+w) = (v <= (w::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
137 |
by (Simp_tac 1); |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
138 |
qed "zadd_left_cancel_zle"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
139 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
140 |
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
|
141 |
|
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_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
|
144 |
|
6998 | 145 |
(*"v<=w ==> z+v <= z+w"*) |
146 |
bind_thm ("zadd_zless_mono2", zadd_left_cancel_zless RS iffD2); |
|
147 |
||
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
148 |
(*"v<=w ==> v+z <= w+z"*) |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
149 |
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
|
150 |
|
6998 | 151 |
(*"v<=w ==> z+v <= z+w"*) |
152 |
bind_thm ("zadd_zle_mono2", zadd_left_cancel_zle RS iffD2); |
|
153 |
||
7081 | 154 |
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
|
155 |
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
|
156 |
by (Simp_tac 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
157 |
qed "zadd_zle_mono"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
158 |
|
7081 | 159 |
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
|
160 |
by (etac (zadd_zless_mono1 RS zless_zle_trans) 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
161 |
by (Simp_tac 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
162 |
qed "zadd_zless_mono"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
163 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
164 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
165 |
(*** Comparison laws ***) |
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 "(- x < - y) = (y < (x::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
168 |
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
|
169 |
qed "zminus_zless_zminus"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
170 |
Addsimps [zminus_zless_zminus]; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
171 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
172 |
Goal "(- x <= - y) = (y <= (x::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
173 |
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
|
174 |
qed "zminus_zle_zminus"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
175 |
Addsimps [zminus_zle_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 |
(** 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
|
178 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
179 |
Goal "(x < - y) = (y < - (x::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
180 |
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
|
181 |
qed "zless_zminus"; |
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 |
Goal "(- x < y) = (- y < (x::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
184 |
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
|
185 |
qed "zminus_zless"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
186 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
187 |
Goal "(x <= - y) = (y <= - (x::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
188 |
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
|
189 |
qed "zle_zminus"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
190 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
191 |
Goal "(- x <= y) = (- y <= (x::int))"; |
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
192 |
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
|
193 |
qed "zminus_zle"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
194 |
|
6917 | 195 |
Goal "(x = - y) = (y = - (x::int))"; |
196 |
by Auto_tac; |
|
197 |
qed "equation_zminus"; |
|
198 |
||
199 |
Goal "(- x = y) = (- (y::int) = x)"; |
|
200 |
by Auto_tac; |
|
201 |
qed "zminus_equation"; |
|
202 |
||
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
203 |
Goal "- (int (Suc n)) < int 0"; |
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
204 |
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
|
205 |
qed "negative_zless_0"; |
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:
5562
diff
changeset
|
207 |
Goal "- (int (Suc n)) < int m"; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
208 |
by (rtac (negative_zless_0 RS zless_zle_trans) 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
209 |
by (Simp_tac 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
210 |
qed "negative_zless"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
211 |
AddIffs [negative_zless]; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
212 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
213 |
Goal "- int n <= int 0"; |
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
214 |
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
|
215 |
qed "negative_zle_0"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
216 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
217 |
Goal "- int n <= int m"; |
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
218 |
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
|
219 |
qed "negative_zle"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
220 |
AddIffs [negative_zle]; |
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:
5562
diff
changeset
|
222 |
Goal "~(int 0 <= - (int (Suc n)))"; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
223 |
by (stac zle_zminus 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
224 |
by (Simp_tac 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
225 |
qed "not_zle_0_negative"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
226 |
Addsimps [not_zle_0_negative]; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
227 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
228 |
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
|
229 |
by Safe_tac; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
230 |
by (Simp_tac 3); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
231 |
by (dtac (zle_zminus RS iffD1) 2); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
232 |
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
|
233 |
by (ALLGOALS Asm_full_simp_tac); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
234 |
qed "int_zle_neg"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
235 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
236 |
Goal "~(int n < - int m)"; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
237 |
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
|
238 |
qed "not_int_zless_negative"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
239 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
240 |
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
|
241 |
by (rtac iffI 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
242 |
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
|
243 |
by (dtac sym 1); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
244 |
by (ALLGOALS Asm_simp_tac); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
245 |
qed "negative_eq_positive"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
246 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
247 |
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
|
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:
5562
diff
changeset
|
250 |
Goal "(w <= z) = (EX n. z = w + int n)"; |
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
251 |
by (auto_tac (claset() addSIs [not_sym RS not0_implies_Suc], |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
252 |
simpset() addsimps [zless_iff_Suc_zadd, integ_le_less])); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
253 |
qed "zle_iff_zadd"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
254 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
255 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
256 |
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
|
257 |
by Auto_tac; |
6917 | 258 |
qed "neg_eq_less_int0"; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
259 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
260 |
Goalw [zle_def] "(~neg x) = (int 0 <= x)"; |
6917 | 261 |
by (simp_tac (simpset() addsimps [neg_eq_less_int0]) 1); |
262 |
qed "not_neg_eq_ge_int0"; |
|
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 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
265 |
(**** 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
|
266 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
267 |
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
|
268 |
by Auto_tac; |
7009
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
paulson
parents:
6998
diff
changeset
|
269 |
qed "nat_int"; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
270 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
271 |
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
|
272 |
by (auto_tac (claset(), |
6917 | 273 |
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:
6998
diff
changeset
|
274 |
qed "nat_zminus_int"; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
275 |
|
7009
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
paulson
parents:
6998
diff
changeset
|
276 |
Addsimps [nat_int, nat_zminus_int]; |
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
277 |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
278 |
Goal "~ neg z ==> int (nat z) = z"; |
6917 | 279 |
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
|
280 |
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
|
281 |
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
|
282 |
qed "not_neg_nat"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
283 |
|
7081 | 284 |
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
|
285 |
by (auto_tac (claset(), |
6917 | 286 |
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
|
287 |
zdiff_eq_eq RS sym, zdiff_def])); |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
288 |
qed "negD"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
289 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
290 |
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
|
291 |
by Auto_tac; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
292 |
qed "neg_nat"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
293 |
|
7518
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
294 |
Goal "(m < nat z) = (int m < z)"; |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
295 |
by (case_tac "neg z" 1); |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
296 |
by (etac (not_neg_nat RS subst) 2); |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
297 |
by (auto_tac (claset(), simpset() addsimps [neg_nat])); |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
298 |
by (auto_tac (claset() addDs [order_less_trans], |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
299 |
simpset() addsimps [neg_eq_less_int0])); |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
300 |
qed "zless_nat_eq_int_zless"; |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
301 |
|
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
302 |
Goal "z <= int 0 ==> nat z = 0"; |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
303 |
by (auto_tac (claset(), |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
304 |
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:
7081
diff
changeset
|
305 |
zle_def, neg_nat])); |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
306 |
qed "nat_le_int0"; |
67bde103ec0c
generalized the theorem zless_zero_nat to zless_nat_eq_int_zless, and
paulson
parents:
7081
diff
changeset
|
307 |
|
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
308 |
(*An alternative condition is int 0 <= w *) |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
309 |
Goal "int 0 < z ==> (nat w < nat z) = (w < z)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
310 |
by (stac (zless_int RS sym) 1); |
6917 | 311 |
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:
6717
diff
changeset
|
312 |
order_le_less]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
313 |
by (case_tac "neg w" 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
314 |
by (asm_simp_tac (simpset() addsimps [not_neg_nat]) 2); |
6917 | 315 |
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:
6717
diff
changeset
|
316 |
by (blast_tac (claset() addIs [order_less_trans]) 1); |
6917 | 317 |
val lemma = result(); |
318 |
||
319 |
Goal "(nat w < nat z) = (int 0 < z & w < z)"; |
|
320 |
by (case_tac "int 0 < z" 1); |
|
321 |
by (auto_tac (claset(), |
|
7009
d6a721e7125d
more renaming of theorems from _nat to _int (corresponding to a function that
paulson
parents:
6998
diff
changeset
|
322 |
simpset() addsimps [lemma, nat_le_int0, linorder_not_less])); |
6917 | 323 |
qed "zless_nat_conj"; |
324 |
||
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
325 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
326 |
(* 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
|
327 |
|
6942 | 328 |
val prems = Goal |
329 |
"[|!! 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
|
330 |
by (case_tac "neg z" 1); |
6942 | 331 |
by (fast_tac (claset() addSDs [negD] addSEs prems) 1); |
332 |
by (dtac (not_neg_nat RS sym) 1); |
|
333 |
by (eresolve_tac prems 1); |
|
5562
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
334 |
qed "int_cases"; |
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
335 |
|
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff
changeset
|
336 |
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
|
337 |
|
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
338 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
339 |
(*** Monotonicity of Multiplication ***) |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
340 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
341 |
Goal "i <= (j::int) ==> i * int k <= j * int k"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
342 |
by (induct_tac "k" 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
343 |
by (stac int_Suc_int_1 2); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
344 |
by (ALLGOALS |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
345 |
(asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2, zadd_zle_mono]))); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
346 |
val lemma = result(); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
347 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
348 |
Goal "[| i <= j; int 0 <= k |] ==> i*k <= j*k"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
349 |
by (res_inst_tac [("t", "k")] (not_neg_nat RS subst) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
350 |
by (etac lemma 2); |
6917 | 351 |
by (full_simp_tac (simpset() addsimps [not_neg_eq_ge_int0]) 1); |
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
352 |
qed "zmult_zle_mono1"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
353 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
354 |
Goal "[| i <= j; k <= int 0 |] ==> j*k <= i*k"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
355 |
by (rtac (zminus_zle_zminus RS iffD1) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
356 |
by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
357 |
zmult_zle_mono1, zle_zminus]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
358 |
qed "zmult_zle_mono1_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
359 |
|
6942 | 360 |
Goal "[| i <= j; int 0 <= k |] ==> k*i <= k*j"; |
361 |
by (dtac zmult_zle_mono1 1); |
|
362 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); |
|
363 |
qed "zmult_zle_mono2"; |
|
364 |
||
365 |
Goal "[| i <= j; k <= int 0 |] ==> k*j <= k*i"; |
|
366 |
by (dtac zmult_zle_mono1_neg 1); |
|
367 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); |
|
368 |
qed "zmult_zle_mono2_neg"; |
|
369 |
||
6990 | 370 |
(* <= monotonicity, BOTH arguments*) |
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
371 |
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:
6717
diff
changeset
|
372 |
by (etac (zmult_zle_mono1 RS order_trans) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
373 |
by (assume_tac 1); |
6990 | 374 |
by (etac zmult_zle_mono2 1); |
375 |
by (assume_tac 1); |
|
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
376 |
qed "zmult_zle_mono"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
377 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
378 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
379 |
(** strict, in 1st argument; proof is by induction on k>0 **) |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
380 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
381 |
Goal "i<j ==> 0<k --> int k * i < int k * j"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
382 |
by (induct_tac "k" 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
383 |
by (stac int_Suc_int_1 2); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
384 |
by (case_tac "n=0" 2); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
385 |
by (ALLGOALS (asm_full_simp_tac |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
386 |
(simpset() addsimps [zadd_zmult_distrib, zadd_zless_mono, |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
387 |
order_le_less]))); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
388 |
val lemma = result() RS mp; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
389 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
390 |
Goal "[| i<j; int 0 < k |] ==> k*i < k*j"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
391 |
by (res_inst_tac [("t", "k")] (not_neg_nat RS subst) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
392 |
by (etac lemma 2); |
6917 | 393 |
by (asm_simp_tac (simpset() addsimps [not_neg_eq_ge_int0, |
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
394 |
order_le_less]) 1); |
6917 | 395 |
by (forward_tac [conjI RS (zless_nat_conj RS iffD2)] 1); |
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
396 |
by Auto_tac; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
397 |
qed "zmult_zless_mono2"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
398 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
399 |
Goal "[| i<j; int 0 < k |] ==> i*k < j*k"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
400 |
by (dtac zmult_zless_mono2 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
401 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
402 |
qed "zmult_zless_mono1"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
403 |
|
6990 | 404 |
(* < monotonicity, BOTH arguments*) |
405 |
Goal "[| i < j; k < l; int 0 < j; int 0 < k |] ==> i*k < j*l"; |
|
406 |
by (etac (zmult_zless_mono1 RS order_less_trans) 1); |
|
407 |
by (assume_tac 1); |
|
408 |
by (etac zmult_zless_mono2 1); |
|
409 |
by (assume_tac 1); |
|
410 |
qed "zmult_zless_mono"; |
|
411 |
||
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
412 |
Goal "[| i<j; k < int 0 |] ==> j*k < i*k"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
413 |
by (rtac (zminus_zless_zminus RS iffD1) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
414 |
by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
415 |
zmult_zless_mono1, zless_zminus]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
416 |
qed "zmult_zless_mono1_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
417 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
418 |
Goal "[| i<j; k < int 0 |] ==> k*j < k*i"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
419 |
by (rtac (zminus_zless_zminus RS iffD1) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
420 |
by (asm_simp_tac (simpset() addsimps [zmult_zminus RS sym, |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
421 |
zmult_zless_mono2, zless_zminus]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
422 |
qed "zmult_zless_mono2_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
423 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
424 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
425 |
Goal "(m*n = int 0) = (m = int 0 | n = int 0)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
426 |
by (case_tac "m < int 0" 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
427 |
by (auto_tac (claset(), |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
428 |
simpset() addsimps [linorder_not_less, order_le_less, |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
429 |
linorder_neq_iff])); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
430 |
by (REPEAT |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
431 |
(force_tac (claset() addDs [zmult_zless_mono1_neg, zmult_zless_mono1], |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
432 |
simpset()) 1)); |
6990 | 433 |
qed "zmult_eq_int0_iff"; |
6866
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
434 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
435 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
436 |
Goal "int 0 < k ==> (m*k < n*k) = (m<n)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
437 |
by (safe_tac (claset() addSIs [zmult_zless_mono1])); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
438 |
by (cut_facts_tac [linorder_less_linear] 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
439 |
by (blast_tac (claset() addIs [zmult_zless_mono1] addEs [order_less_asym]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
440 |
qed "zmult_zless_cancel2"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
441 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
442 |
Goal "int 0 < k ==> (k*m < k*n) = (m<n)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
443 |
by (dtac zmult_zless_cancel2 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
444 |
by (asm_full_simp_tac (simpset() addsimps [zmult_commute]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
445 |
qed "zmult_zless_cancel1"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
446 |
Addsimps [zmult_zless_cancel1, zmult_zless_cancel2]; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
447 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
448 |
Goal "k < int 0 ==> (m*k < n*k) = (n<m)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
449 |
by (safe_tac (claset() addSIs [zmult_zless_mono1_neg])); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
450 |
by (cut_facts_tac [linorder_less_linear] 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
451 |
by (blast_tac (claset() addIs [zmult_zless_mono1_neg] |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
452 |
addEs [order_less_asym]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
453 |
qed "zmult_zless_cancel2_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
454 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
455 |
Goal "k < int 0 ==> (k*m < k*n) = (n<m)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
456 |
by (dtac zmult_zless_cancel2_neg 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
457 |
by (asm_full_simp_tac (simpset() addsimps [zmult_commute]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
458 |
qed "zmult_zless_cancel1_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
459 |
Addsimps [zmult_zless_cancel1_neg, zmult_zless_cancel2_neg]; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
460 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
461 |
Goal "int 0 < k ==> (m*k <= n*k) = (m<=n)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
462 |
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
463 |
qed "zmult_zle_cancel2"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
464 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
465 |
Goal "int 0 < k ==> (k*m <= k*n) = (m<=n)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
466 |
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
467 |
qed "zmult_zle_cancel1"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
468 |
Addsimps [zmult_zle_cancel1, zmult_zle_cancel2]; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
469 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
470 |
Goal "k < int 0 ==> (m*k <= n*k) = (n<=m)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
471 |
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
472 |
qed "zmult_zle_cancel2_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
473 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
474 |
Goal "k < int 0 ==> (k*m <= k*n) = (n<=m)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
475 |
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
476 |
qed "zmult_zle_cancel1_neg"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
477 |
Addsimps [zmult_zle_cancel1_neg, zmult_zle_cancel2_neg]; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
478 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
479 |
Goal "k ~= int 0 ==> (m*k = n*k) = (m=n)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
480 |
by (cut_facts_tac [linorder_less_linear] 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
481 |
by Safe_tac; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
482 |
by (assume_tac 2); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
483 |
by (REPEAT |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
484 |
(force_tac (claset() addD2 ("mono_neg", zmult_zless_mono1_neg) |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
485 |
addD2 ("mono_pos", zmult_zless_mono1), |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
486 |
simpset() addsimps [linorder_neq_iff]) 1)); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
487 |
qed "zmult_cancel2"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
488 |
|
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
489 |
Goal "k ~= int 0 ==> (k*m = k*n) = (m=n)"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
490 |
by (dtac zmult_cancel2 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
491 |
by (asm_full_simp_tac (simpset() addsimps [zmult_commute]) 1); |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
492 |
qed "zmult_cancel1"; |
f795b63139ec
many new theorems concerning multiplication and (in)equations
paulson
parents:
6717
diff
changeset
|
493 |
Addsimps [zmult_cancel1, zmult_cancel2]; |