| author | paulson | 
| Mon, 12 Jan 2004 16:45:35 +0100 | |
| changeset 14352 | a8b1a44d8264 | 
| parent 14295 | 7f115e5c5de4 | 
| child 14353 | 79f9fbef9106 | 
| permissions | -rw-r--r-- | 
| 14259 | 1 | (* Title: HOL/Integ/IntArith.thy | 
| 2 | ID: $Id$ | |
| 3 | Authors: Larry Paulson and Tobias Nipkow | |
| 4 | *) | |
| 12023 | 5 | |
| 6 | header {* Integer arithmetic *}
 | |
| 7 | ||
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
9214diff
changeset | 8 | theory IntArith = Bin | 
| 14259 | 9 | files ("int_arith1.ML"):
 | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
9214diff
changeset | 10 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 11 | subsection{*Inequality Reasoning for the Arithmetic Simproc*}
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 12 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 13 | lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 14 | proof (auto simp add: zle_def zless_iff_Suc_zadd) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 15 | fix m n | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 16 | assume "w + 1 = w + (1 + int m) + (1 + int n)" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 17 | hence "(w + 1) + (1 + int (m + n)) = (w + 1) + 0" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 18 | by (simp add: add_ac zadd_int [symmetric]) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 19 | hence "int (Suc(m+n)) = 0" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 20 | by (simp only: Ring_and_Field.add_left_cancel int_Suc) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 21 | thus False by (simp only: int_eq_0_conv) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 22 | qed | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 23 | |
| 12023 | 24 | use "int_arith1.ML" | 
| 25 | setup int_arith_setup | |
| 14259 | 26 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 27 | subsection{*More inequality reasoning*}
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 28 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 29 | lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)" | 
| 14259 | 30 | by arith | 
| 31 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 32 | lemma add1_zle_eq: "(w + (1::int) \<le> z) = (w<z)" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 33 | by arith | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 34 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 35 | lemma zle_diff1_eq [simp]: "(w \<le> z - (1::int)) = (w<(z::int))" | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 36 | by arith | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 37 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 38 | lemma zle_add1_eq_le [simp]: "(w < z + 1) = (w\<le>(z::int))" | 
| 14259 | 39 | by arith | 
| 40 | ||
| 41 | lemma zadd_left_cancel0 [simp]: "(z = z + w) = (w = (0::int))" | |
| 42 | by arith | |
| 43 | ||
| 44 | subsection{*Results about @{term nat}*}
 | |
| 45 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 46 | lemma nonneg_eq_int: "[| 0 \<le> z; !!m. z = int m ==> P |] ==> P" | 
| 14259 | 47 | by (blast dest: nat_0_le sym) | 
| 48 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 49 | lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)" | 
| 14259 | 50 | by auto | 
| 51 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 52 | lemma nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)" | 
| 14259 | 53 | by auto | 
| 54 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 55 | lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)" | 
| 14259 | 56 | apply (rule iffI) | 
| 57 | apply (erule nat_0_le [THEN subst]) | |
| 58 | apply (simp_all del: zless_int add: zless_int [symmetric]) | |
| 59 | done | |
| 60 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 61 | lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)" | 
| 14259 | 62 | by (auto simp add: nat_eq_iff2) | 
| 63 | ||
| 64 | ||
| 65 | (*Users don't want to see (int 0), int(Suc 0) or w + - z*) | |
| 66 | declare Zero_int_def [symmetric, simp] | |
| 67 | declare One_int_def [symmetric, simp] | |
| 68 | ||
| 69 | text{*cooper.ML refers to this theorem*}
 | |
| 70 | lemmas zdiff_def_symmetric = zdiff_def [symmetric, simp] | |
| 71 | ||
| 72 | lemma nat_0: "nat 0 = 0" | |
| 73 | by (simp add: nat_eq_iff) | |
| 74 | ||
| 75 | lemma nat_1: "nat 1 = Suc 0" | |
| 76 | by (subst nat_eq_iff, simp) | |
| 77 | ||
| 78 | lemma nat_2: "nat 2 = Suc (Suc 0)" | |
| 79 | by (subst nat_eq_iff, simp) | |
| 80 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 81 | lemma nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)" | 
| 14259 | 82 | apply (case_tac "neg z") | 
| 83 | apply (auto simp add: nat_less_iff) | |
| 84 | apply (auto intro: zless_trans simp add: neg_eq_less_0 zle_def) | |
| 85 | done | |
| 86 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 87 | lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)" | 
| 14259 | 88 | by (auto simp add: linorder_not_less [symmetric] zless_nat_conj) | 
| 89 | ||
| 90 | subsection{*@{term abs}: Absolute Value, as an Integer*}
 | |
| 91 | ||
| 92 | (* Simpler: use zabs_def as rewrite rule | |
| 93 | but arith_tac is not parameterized by such simp rules | |
| 94 | *) | |
| 95 | ||
| 14295 | 96 | lemma zabs_split [arith_split]: | 
| 97 | "P(abs(i::int)) = ((0 \<le> i --> P i) & (i < 0 --> P(-i)))" | |
| 14259 | 98 | by (simp add: zabs_def) | 
| 99 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 100 | lemma zero_le_zabs [iff]: "0 \<le> abs (z::int)" | 
| 14259 | 101 | by (simp add: zabs_def) | 
| 102 | ||
| 103 | ||
| 104 | text{*This simplifies expressions of the form @{term "int n = z"} where
 | |
| 105 | z is an integer literal.*} | |
| 106 | declare int_eq_iff [of _ "number_of v", standard, simp] | |
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13685diff
changeset | 107 | |
| 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13685diff
changeset | 108 | lemma zabs_eq_iff: | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 109 | "(abs (z::int) = w) = (z = w \<and> 0 \<le> z \<or> z = -w \<and> z < 0)" | 
| 13837 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13685diff
changeset | 110 | by (auto simp add: zabs_def) | 
| 
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
 paulson parents: 
13685diff
changeset | 111 | |
| 13849 | 112 | lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)" | 
| 113 | by simp | |
| 114 | ||
| 14295 | 115 | lemma split_nat [arith_split]: | 
| 14259 | 116 | "P(nat(i::int)) = ((\<forall>n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))" | 
| 13575 | 117 | (is "?P = (?L & ?R)") | 
| 118 | proof (cases "i < 0") | |
| 119 | case True thus ?thesis by simp | |
| 120 | next | |
| 121 | case False | |
| 122 | have "?P = ?L" | |
| 123 | proof | |
| 124 | assume ?P thus ?L using False by clarsimp | |
| 125 | next | |
| 126 | assume ?L thus ?P using False by simp | |
| 127 | qed | |
| 128 | with False show ?thesis by simp | |
| 129 | qed | |
| 130 | ||
| 13685 | 131 | subsubsection "Induction principles for int" | 
| 132 | ||
| 133 | (* `set:int': dummy construction *) | |
| 134 | theorem int_ge_induct[case_names base step,induct set:int]: | |
| 135 | assumes ge: "k \<le> (i::int)" and | |
| 136 | base: "P(k)" and | |
| 137 | step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)" | |
| 138 | shows "P i" | |
| 139 | proof - | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 140 |   { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i"
 | 
| 13685 | 141 | proof (induct n) | 
| 142 | case 0 | |
| 143 | hence "i = k" by arith | |
| 144 | thus "P i" using base by simp | |
| 145 | next | |
| 146 | case (Suc n) | |
| 147 | hence "n = nat((i - 1) - k)" by arith | |
| 148 | moreover | |
| 149 | have ki1: "k \<le> i - 1" using Suc.prems by arith | |
| 150 | ultimately | |
| 151 | have "P(i - 1)" by(rule Suc.hyps) | |
| 152 | from step[OF ki1 this] show ?case by simp | |
| 153 | qed | |
| 154 | } | |
| 155 | from this ge show ?thesis by fast | |
| 156 | qed | |
| 157 | ||
| 158 | (* `set:int': dummy construction *) | |
| 159 | theorem int_gr_induct[case_names base step,induct set:int]: | |
| 160 | assumes gr: "k < (i::int)" and | |
| 161 | base: "P(k+1)" and | |
| 162 | step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)" | |
| 163 | shows "P i" | |
| 164 | apply(rule int_ge_induct[of "k + 1"]) | |
| 165 | using gr apply arith | |
| 166 | apply(rule base) | |
| 14259 | 167 | apply (rule step, simp+) | 
| 13685 | 168 | done | 
| 169 | ||
| 170 | theorem int_le_induct[consumes 1,case_names base step]: | |
| 171 | assumes le: "i \<le> (k::int)" and | |
| 172 | base: "P(k)" and | |
| 173 | step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)" | |
| 174 | shows "P i" | |
| 175 | proof - | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 176 |   { fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i"
 | 
| 13685 | 177 | proof (induct n) | 
| 178 | case 0 | |
| 179 | hence "i = k" by arith | |
| 180 | thus "P i" using base by simp | |
| 181 | next | |
| 182 | case (Suc n) | |
| 183 | hence "n = nat(k - (i+1))" by arith | |
| 184 | moreover | |
| 185 | have ki1: "i + 1 \<le> k" using Suc.prems by arith | |
| 186 | ultimately | |
| 187 | have "P(i+1)" by(rule Suc.hyps) | |
| 188 | from step[OF ki1 this] show ?case by simp | |
| 189 | qed | |
| 190 | } | |
| 191 | from this le show ?thesis by fast | |
| 192 | qed | |
| 193 | ||
| 194 | theorem int_less_induct[consumes 1,case_names base step]: | |
| 195 | assumes less: "(i::int) < k" and | |
| 196 | base: "P(k - 1)" and | |
| 197 | step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)" | |
| 198 | shows "P i" | |
| 199 | apply(rule int_le_induct[of _ "k - 1"]) | |
| 200 | using less apply arith | |
| 201 | apply(rule base) | |
| 14259 | 202 | apply (rule step, simp+) | 
| 203 | done | |
| 204 | ||
| 205 | subsection{*Simple Equations*}
 | |
| 206 | ||
| 207 | lemma int_diff_minus_eq [simp]: "x - - y = x + (y::int)" | |
| 208 | by simp | |
| 209 | ||
| 210 | lemma abs_abs [simp]: "abs(abs(x::int)) = abs(x)" | |
| 211 | by arith | |
| 212 | ||
| 213 | lemma abs_minus [simp]: "abs(-(x::int)) = abs(x)" | |
| 214 | by arith | |
| 215 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 216 | lemma triangle_ineq: "abs(x+y) \<le> abs(x) + abs(y::int)" | 
| 14259 | 217 | by arith | 
| 218 | ||
| 219 | ||
| 220 | subsection{*Intermediate value theorems*}
 | |
| 221 | ||
| 222 | lemma int_val_lemma: | |
| 223 | "(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) --> | |
| 224 | f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))" | |
| 14271 | 225 | apply (induct_tac "n", simp) | 
| 14259 | 226 | apply (intro strip) | 
| 227 | apply (erule impE, simp) | |
| 228 | apply (erule_tac x = n in allE, simp) | |
| 229 | apply (case_tac "k = f (n+1) ") | |
| 230 | apply force | |
| 231 | apply (erule impE) | |
| 232 | apply (simp add: zabs_def split add: split_if_asm) | |
| 233 | apply (blast intro: le_SucI) | |
| 234 | done | |
| 235 | ||
| 236 | lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)] | |
| 237 | ||
| 238 | lemma nat_intermed_int_val: | |
| 239 | "[| \<forall>i. m \<le> i & i < n --> abs(f(i + 1::nat) - f i) \<le> 1; m < n; | |
| 240 | f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)" | |
| 241 | apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k | |
| 242 | in int_val_lemma) | |
| 243 | apply simp | |
| 244 | apply (erule impE) | |
| 245 | apply (intro strip) | |
| 246 | apply (erule_tac x = "i+m" in allE, arith) | |
| 247 | apply (erule exE) | |
| 248 | apply (rule_tac x = "i+m" in exI, arith) | |
| 249 | done | |
| 250 | ||
| 251 | ||
| 252 | subsection{*Some convenient biconditionals for products of signs*}
 | |
| 253 | ||
| 254 | lemma zmult_pos: "[| (0::int) < i; 0 < j |] ==> 0 < i*j" | |
| 14266 | 255 | by (rule Ring_and_Field.mult_pos) | 
| 14259 | 256 | |
| 257 | lemma zmult_neg: "[| i < (0::int); j < 0 |] ==> 0 < i*j" | |
| 14266 | 258 | by (rule Ring_and_Field.mult_neg) | 
| 14259 | 259 | |
| 260 | lemma zmult_pos_neg: "[| (0::int) < i; j < 0 |] ==> i*j < 0" | |
| 14266 | 261 | by (rule Ring_and_Field.mult_pos_neg) | 
| 14259 | 262 | |
| 263 | lemma int_0_less_mult_iff: "((0::int) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)" | |
| 14266 | 264 | by (rule Ring_and_Field.zero_less_mult_iff) | 
| 14259 | 265 | |
| 266 | lemma int_0_le_mult_iff: "((0::int) \<le> x*y) = (0 \<le> x & 0 \<le> y | x \<le> 0 & y \<le> 0)" | |
| 14266 | 267 | by (rule Ring_and_Field.zero_le_mult_iff) | 
| 14259 | 268 | |
| 269 | lemma zmult_less_0_iff: "(x*y < (0::int)) = (0 < x & y < 0 | x < 0 & 0 < y)" | |
| 14266 | 270 | by (rule Ring_and_Field.mult_less_0_iff) | 
| 14259 | 271 | |
| 272 | lemma zmult_le_0_iff: "(x*y \<le> (0::int)) = (0 \<le> x & y \<le> 0 | x \<le> 0 & 0 \<le> y)" | |
| 14266 | 273 | by (rule Ring_and_Field.mult_le_0_iff) | 
| 14259 | 274 | |
| 275 | lemma abs_eq_0 [iff]: "(abs x = 0) = (x = (0::int))" | |
| 14266 | 276 | by (rule Ring_and_Field.abs_eq_0) | 
| 14259 | 277 | |
| 278 | lemma zero_less_abs_iff [iff]: "(0 < abs x) = (x ~= (0::int))" | |
| 14266 | 279 | by (rule Ring_and_Field.zero_less_abs_iff) | 
| 14259 | 280 | |
| 281 | lemma square_nonzero [simp]: "0 \<le> x * (x::int)" | |
| 14266 | 282 | by (rule Ring_and_Field.zero_le_square) | 
| 14259 | 283 | |
| 284 | ||
| 285 | subsection{*Products and 1, by T. M. Rasmussen*}
 | |
| 286 | ||
| 287 | lemma zmult_eq_self_iff: "(m = m*(n::int)) = (n = 1 | m = 0)" | |
| 288 | apply auto | |
| 289 | apply (subgoal_tac "m*1 = m*n") | |
| 290 | apply (drule zmult_cancel1 [THEN iffD1], auto) | |
| 13685 | 291 | done | 
| 292 | ||
| 14259 | 293 | lemma zless_1_zmult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::int)" | 
| 294 | apply (rule_tac y = "1*n" in order_less_trans) | |
| 295 | apply (rule_tac [2] zmult_zless_mono1) | |
| 296 | apply (simp_all (no_asm_simp)) | |
| 297 | done | |
| 298 | ||
| 299 | lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)" | |
| 300 | apply auto | |
| 301 | apply (case_tac "m=1") | |
| 302 | apply (case_tac [2] "n=1") | |
| 303 | apply (case_tac [4] "m=1") | |
| 304 | apply (case_tac [5] "n=1", auto) | |
| 305 | apply (tactic"distinct_subgoals_tac") | |
| 306 | apply (subgoal_tac "1<m*n", simp) | |
| 307 | apply (rule zless_1_zmult, arith) | |
| 308 | apply (subgoal_tac "0<n", arith) | |
| 309 | apply (subgoal_tac "0<m*n") | |
| 310 | apply (drule int_0_less_mult_iff [THEN iffD1], auto) | |
| 311 | done | |
| 312 | ||
| 313 | lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))" | |
| 314 | apply (case_tac "0<m") | |
| 14271 | 315 | apply (simp add: pos_zmult_eq_1_iff) | 
| 14259 | 316 | apply (case_tac "m=0") | 
| 14271 | 317 | apply (simp del: number_of_reorient) | 
| 14259 | 318 | apply (subgoal_tac "0 < -m") | 
| 319 | apply (drule_tac n = "-n" in pos_zmult_eq_1_iff, auto) | |
| 320 | done | |
| 321 | ||
| 322 | ||
| 323 | subsection{*More about nat*}
 | |
| 324 | ||
| 14271 | 325 | (*Analogous to zadd_int*) | 
| 326 | lemma zdiff_int: "n \<le> m ==> int m - int n = int (m-n)" | |
| 327 | by (induct m n rule: diff_induct, simp_all) | |
| 328 | ||
| 14259 | 329 | lemma nat_add_distrib: | 
| 330 | "[| (0::int) \<le> z; 0 \<le> z' |] ==> nat (z+z') = nat z + nat z'" | |
| 331 | apply (rule inj_int [THEN injD]) | |
| 14271 | 332 | apply (simp add: zadd_int [symmetric]) | 
| 14259 | 333 | done | 
| 334 | ||
| 335 | lemma nat_diff_distrib: | |
| 336 | "[| (0::int) \<le> z'; z' \<le> z |] ==> nat (z-z') = nat z - nat z'" | |
| 337 | apply (rule inj_int [THEN injD]) | |
| 14271 | 338 | apply (simp add: zdiff_int [symmetric] nat_le_eq_zle) | 
| 14259 | 339 | done | 
| 340 | ||
| 341 | lemma nat_mult_distrib: "(0::int) \<le> z ==> nat (z*z') = nat z * nat z'" | |
| 342 | apply (case_tac "0 \<le> z'") | |
| 343 | apply (rule inj_int [THEN injD]) | |
| 14271 | 344 | apply (simp add: zmult_int [symmetric] int_0_le_mult_iff) | 
| 14259 | 345 | apply (simp add: zmult_le_0_iff) | 
| 346 | done | |
| 347 | ||
| 348 | lemma nat_mult_distrib_neg: "z \<le> (0::int) ==> nat(z*z') = nat(-z) * nat(-z')" | |
| 349 | apply (rule trans) | |
| 350 | apply (rule_tac [2] nat_mult_distrib, auto) | |
| 351 | done | |
| 352 | ||
| 353 | lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)" | |
| 354 | apply (case_tac "z=0 | w=0") | |
| 355 | apply (auto simp add: zabs_def nat_mult_distrib [symmetric] | |
| 356 | nat_mult_distrib_neg [symmetric] zmult_less_0_iff) | |
| 357 | done | |
| 358 | ||
| 359 | ML | |
| 360 | {*
 | |
| 361 | val zle_diff1_eq = thm "zle_diff1_eq"; | |
| 362 | val zle_add1_eq_le = thm "zle_add1_eq_le"; | |
| 363 | val nonneg_eq_int = thm "nonneg_eq_int"; | |
| 364 | val nat_eq_iff = thm "nat_eq_iff"; | |
| 365 | val nat_eq_iff2 = thm "nat_eq_iff2"; | |
| 366 | val nat_less_iff = thm "nat_less_iff"; | |
| 367 | val int_eq_iff = thm "int_eq_iff"; | |
| 368 | val nat_0 = thm "nat_0"; | |
| 369 | val nat_1 = thm "nat_1"; | |
| 370 | val nat_2 = thm "nat_2"; | |
| 371 | val nat_less_eq_zless = thm "nat_less_eq_zless"; | |
| 372 | val nat_le_eq_zle = thm "nat_le_eq_zle"; | |
| 373 | val zabs_split = thm "zabs_split"; | |
| 374 | val zero_le_zabs = thm "zero_le_zabs"; | |
| 375 | ||
| 376 | val int_diff_minus_eq = thm "int_diff_minus_eq"; | |
| 377 | val abs_abs = thm "abs_abs"; | |
| 378 | val abs_minus = thm "abs_minus"; | |
| 379 | val triangle_ineq = thm "triangle_ineq"; | |
| 380 | val nat_intermed_int_val = thm "nat_intermed_int_val"; | |
| 381 | val zmult_pos = thm "zmult_pos"; | |
| 382 | val zmult_neg = thm "zmult_neg"; | |
| 383 | val zmult_pos_neg = thm "zmult_pos_neg"; | |
| 384 | val int_0_less_mult_iff = thm "int_0_less_mult_iff"; | |
| 385 | val int_0_le_mult_iff = thm "int_0_le_mult_iff"; | |
| 386 | val zmult_less_0_iff = thm "zmult_less_0_iff"; | |
| 387 | val zmult_le_0_iff = thm "zmult_le_0_iff"; | |
| 388 | val abs_mult = thm "abs_mult"; | |
| 389 | val abs_eq_0 = thm "abs_eq_0"; | |
| 390 | val zero_less_abs_iff = thm "zero_less_abs_iff"; | |
| 391 | val square_nonzero = thm "square_nonzero"; | |
| 392 | val zmult_eq_self_iff = thm "zmult_eq_self_iff"; | |
| 393 | val zless_1_zmult = thm "zless_1_zmult"; | |
| 394 | val pos_zmult_eq_1_iff = thm "pos_zmult_eq_1_iff"; | |
| 395 | val zmult_eq_1_iff = thm "zmult_eq_1_iff"; | |
| 396 | val nat_add_distrib = thm "nat_add_distrib"; | |
| 397 | val nat_diff_distrib = thm "nat_diff_distrib"; | |
| 398 | val nat_mult_distrib = thm "nat_mult_distrib"; | |
| 399 | val nat_mult_distrib_neg = thm "nat_mult_distrib_neg"; | |
| 400 | val nat_abs_mult_distrib = thm "nat_abs_mult_distrib"; | |
| 401 | *} | |
| 402 | ||
| 7707 | 403 | end |