author | kleing |
Tue, 13 May 2003 08:59:21 +0200 | |
changeset 14024 | 213dcc39358f |
parent 13849 | 2584233cf3ef |
child 14259 | 79f7d3451b1e |
permissions | -rw-r--r-- |
12023 | 1 |
|
2 |
header {* Integer arithmetic *} |
|
3 |
||
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
9214
diff
changeset
|
4 |
theory IntArith = Bin |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
9214
diff
changeset
|
5 |
files ("int_arith1.ML") ("int_arith2.ML"): |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
9214
diff
changeset
|
6 |
|
12023 | 7 |
use "int_arith1.ML" |
8 |
setup int_arith_setup |
|
9 |
use "int_arith2.ML" |
|
13837
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13685
diff
changeset
|
10 |
|
12023 | 11 |
declare zabs_split [arith_split] |
12 |
||
13837
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13685
diff
changeset
|
13 |
lemma zabs_eq_iff: |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13685
diff
changeset
|
14 |
"(abs (z::int) = w) = (z = w \<and> 0 <= z \<or> z = -w \<and> z < 0)" |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13685
diff
changeset
|
15 |
by (auto simp add: zabs_def) |
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13685
diff
changeset
|
16 |
|
13849 | 17 |
lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)" |
18 |
by simp |
|
19 |
||
13575 | 20 |
lemma split_nat[arith_split]: |
21 |
"P(nat(i::int)) = ((ALL n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))" |
|
22 |
(is "?P = (?L & ?R)") |
|
23 |
proof (cases "i < 0") |
|
24 |
case True thus ?thesis by simp |
|
25 |
next |
|
26 |
case False |
|
27 |
have "?P = ?L" |
|
28 |
proof |
|
29 |
assume ?P thus ?L using False by clarsimp |
|
30 |
next |
|
31 |
assume ?L thus ?P using False by simp |
|
32 |
qed |
|
33 |
with False show ?thesis by simp |
|
34 |
qed |
|
35 |
||
13685 | 36 |
subsubsection "Induction principles for int" |
37 |
||
38 |
(* `set:int': dummy construction *) |
|
39 |
theorem int_ge_induct[case_names base step,induct set:int]: |
|
40 |
assumes ge: "k \<le> (i::int)" and |
|
41 |
base: "P(k)" and |
|
42 |
step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)" |
|
43 |
shows "P i" |
|
44 |
proof - |
|
45 |
{ fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k <= i \<Longrightarrow> P i" |
|
46 |
proof (induct n) |
|
47 |
case 0 |
|
48 |
hence "i = k" by arith |
|
49 |
thus "P i" using base by simp |
|
50 |
next |
|
51 |
case (Suc n) |
|
52 |
hence "n = nat((i - 1) - k)" by arith |
|
53 |
moreover |
|
54 |
have ki1: "k \<le> i - 1" using Suc.prems by arith |
|
55 |
ultimately |
|
56 |
have "P(i - 1)" by(rule Suc.hyps) |
|
57 |
from step[OF ki1 this] show ?case by simp |
|
58 |
qed |
|
59 |
} |
|
60 |
from this ge show ?thesis by fast |
|
61 |
qed |
|
62 |
||
63 |
(* `set:int': dummy construction *) |
|
64 |
theorem int_gr_induct[case_names base step,induct set:int]: |
|
65 |
assumes gr: "k < (i::int)" and |
|
66 |
base: "P(k+1)" and |
|
67 |
step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)" |
|
68 |
shows "P i" |
|
69 |
apply(rule int_ge_induct[of "k + 1"]) |
|
70 |
using gr apply arith |
|
71 |
apply(rule base) |
|
72 |
apply(rule step) |
|
73 |
apply simp+ |
|
74 |
done |
|
75 |
||
76 |
theorem int_le_induct[consumes 1,case_names base step]: |
|
77 |
assumes le: "i \<le> (k::int)" and |
|
78 |
base: "P(k)" and |
|
79 |
step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)" |
|
80 |
shows "P i" |
|
81 |
proof - |
|
82 |
{ fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i <= k \<Longrightarrow> P i" |
|
83 |
proof (induct n) |
|
84 |
case 0 |
|
85 |
hence "i = k" by arith |
|
86 |
thus "P i" using base by simp |
|
87 |
next |
|
88 |
case (Suc n) |
|
89 |
hence "n = nat(k - (i+1))" by arith |
|
90 |
moreover |
|
91 |
have ki1: "i + 1 \<le> k" using Suc.prems by arith |
|
92 |
ultimately |
|
93 |
have "P(i+1)" by(rule Suc.hyps) |
|
94 |
from step[OF ki1 this] show ?case by simp |
|
95 |
qed |
|
96 |
} |
|
97 |
from this le show ?thesis by fast |
|
98 |
qed |
|
99 |
||
100 |
theorem int_less_induct[consumes 1,case_names base step]: |
|
101 |
assumes less: "(i::int) < k" and |
|
102 |
base: "P(k - 1)" and |
|
103 |
step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)" |
|
104 |
shows "P i" |
|
105 |
apply(rule int_le_induct[of _ "k - 1"]) |
|
106 |
using less apply arith |
|
107 |
apply(rule base) |
|
108 |
apply(rule step) |
|
109 |
apply simp+ |
|
110 |
done |
|
111 |
||
7707 | 112 |
end |
13849 | 113 |