| author | aspinall |
| Sat, 17 Feb 2007 18:01:22 +0100 | |
| changeset 22338 | c7feeba2249e |
| parent 22242 | 020f65c2cdab |
| child 22801 | caffcb450ef4 |
| 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 |
||
| 15131 | 8 |
theory IntArith |
| 22059 | 9 |
imports Numeral "../Wellfounded_Relations" |
| 16417 | 10 |
uses ("int_arith1.ML")
|
| 15131 | 11 |
begin |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
9214
diff
changeset
|
12 |
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
13 |
text{*Duplicate: can't understand why it's necessary*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
14 |
declare numeral_0_eq_0 [simp] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
15 |
|
| 16413 | 16 |
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
17 |
subsection{*Instantiating Binary Arithmetic for the Integers*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
18 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
19 |
instance int :: number_ring |
|
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21191
diff
changeset
|
20 |
int_number_of_def: "number_of w \<equiv> of_int w" |
|
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21191
diff
changeset
|
21 |
by intro_classes (simp only: int_number_of_def) |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
22 |
|
|
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14295
diff
changeset
|
23 |
|
|
14272
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
24 |
subsection{*Inequality Reasoning for the Arithmetic Simproc*}
|
|
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
25 |
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
26 |
lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
27 |
by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
28 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
29 |
lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
30 |
by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
31 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
32 |
lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
33 |
by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
34 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
35 |
lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
36 |
by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
37 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
38 |
text{*Theorem lists for the cancellation simprocs. The use of binary numerals
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
39 |
for 0 and 1 reduces the number of special cases.*} |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
40 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
41 |
lemmas add_0s = add_numeral_0 add_numeral_0_right |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
42 |
lemmas mult_1s = mult_numeral_1 mult_numeral_1_right |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
43 |
mult_minus1 mult_minus1_right |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
44 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
45 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
46 |
subsection{*Special Arithmetic Rules for Abstract 0 and 1*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
47 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
48 |
text{*Arithmetic computations are defined for binary literals, which leaves 0
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
49 |
and 1 as special cases. Addition already has rules for 0, but not 1. |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
50 |
Multiplication and unary minus already have rules for both 0 and 1.*} |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
51 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
52 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
53 |
lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
54 |
by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
55 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
56 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
57 |
lemmas add_number_of_eq = number_of_add [symmetric] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
58 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
59 |
text{*Allow 1 on either or both sides*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
60 |
lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
61 |
by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq) |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
62 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
63 |
lemmas add_special = |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
64 |
one_add_one_is_two |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
65 |
binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
66 |
binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
67 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
68 |
text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
69 |
lemmas diff_special = |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
70 |
binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
71 |
binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
72 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
73 |
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
74 |
lemmas eq_special = |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
75 |
binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
76 |
binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
77 |
binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
78 |
binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
79 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
80 |
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
81 |
lemmas less_special = |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
82 |
binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
83 |
binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
84 |
binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
85 |
binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
86 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
87 |
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
88 |
lemmas le_special = |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
89 |
binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
90 |
binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
91 |
binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
92 |
binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard] |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
93 |
|
| 22192 | 94 |
lemmas arith_special[simp] = |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
95 |
add_special diff_special eq_special less_special le_special |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
96 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
97 |
|
| 22192 | 98 |
lemma min_max_01: "min (0::int) 1 = 0 & min (1::int) 0 = 0 & |
99 |
max (0::int) 1 = 1 & max (1::int) 0 = 1" |
|
100 |
by(simp add:min_def max_def) |
|
101 |
||
102 |
lemmas min_max_special[simp] = |
|
103 |
min_max_01 |
|
104 |
max_def[of "0::int" "number_of v", standard, simp] |
|
105 |
min_def[of "0::int" "number_of v", standard, simp] |
|
106 |
max_def[of "number_of u" "0::int", standard, simp] |
|
107 |
min_def[of "number_of u" "0::int", standard, simp] |
|
108 |
max_def[of "1::int" "number_of v", standard, simp] |
|
109 |
min_def[of "1::int" "number_of v", standard, simp] |
|
110 |
max_def[of "number_of u" "1::int", standard, simp] |
|
111 |
min_def[of "number_of u" "1::int", standard, simp] |
|
112 |
||
113 |
||
| 12023 | 114 |
use "int_arith1.ML" |
115 |
setup int_arith_setup |
|
| 14259 | 116 |
|
|
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14295
diff
changeset
|
117 |
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
118 |
subsection{*Lemmas About Small Numerals*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
119 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
120 |
lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
121 |
proof - |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
122 |
have "(of_int -1 :: 'a) = of_int (- 1)" by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
123 |
also have "... = - of_int 1" by (simp only: of_int_minus) |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
124 |
also have "... = -1" by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
125 |
finally show ?thesis . |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
126 |
qed |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
127 |
|
| 14738 | 128 |
lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_idom,number_ring})"
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
129 |
by (simp add: abs_if) |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
130 |
|
| 14436 | 131 |
lemma abs_power_minus_one [simp]: |
| 15003 | 132 |
"abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,recpower})"
|
| 14436 | 133 |
by (simp add: power_abs) |
134 |
||
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
135 |
lemma of_int_number_of_eq: |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
136 |
"of_int (number_of v) = (number_of v :: 'a :: number_ring)" |
| 15013 | 137 |
by (simp add: number_of_eq) |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
138 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
139 |
text{*Lemmas for specialist use, NOT as default simprules*}
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
140 |
lemma mult_2: "2 * z = (z+z::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
141 |
proof - |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
142 |
have "2*z = (1 + 1)*z" by simp |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
143 |
also have "... = z+z" by (simp add: left_distrib) |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
144 |
finally show ?thesis . |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
145 |
qed |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
146 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
147 |
lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)" |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
148 |
by (subst mult_commute, rule mult_2) |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
149 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
150 |
|
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
151 |
subsection{*More Inequality Reasoning*}
|
|
14272
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
152 |
|
|
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
153 |
lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)" |
| 14259 | 154 |
by arith |
155 |
||
|
14272
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
156 |
lemma add1_zle_eq: "(w + (1::int) \<le> z) = (w<z)" |
|
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
157 |
by arith |
|
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
158 |
|
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
159 |
lemma zle_diff1_eq [simp]: "(w \<le> z - (1::int)) = (w<z)" |
|
14272
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
160 |
by arith |
|
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
161 |
|
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
162 |
lemma zle_add1_eq_le [simp]: "(w < z + (1::int)) = (w\<le>z)" |
| 14259 | 163 |
by arith |
164 |
||
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14353
diff
changeset
|
165 |
lemma int_one_le_iff_zero_less: "((1::int) \<le> z) = (0 < z)" |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14353
diff
changeset
|
166 |
by arith |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14353
diff
changeset
|
167 |
|
|
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14295
diff
changeset
|
168 |
|
|
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14295
diff
changeset
|
169 |
subsection{*The Functions @{term nat} and @{term int}*}
|
| 14259 | 170 |
|
|
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14295
diff
changeset
|
171 |
text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and
|
|
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14295
diff
changeset
|
172 |
@{term "w + - z"}*}
|
| 14259 | 173 |
declare Zero_int_def [symmetric, simp] |
174 |
declare One_int_def [symmetric, simp] |
|
175 |
||
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
176 |
lemmas diff_int_def_symmetric = diff_int_def [symmetric, simp] |
| 14259 | 177 |
|
178 |
lemma nat_0: "nat 0 = 0" |
|
179 |
by (simp add: nat_eq_iff) |
|
180 |
||
181 |
lemma nat_1: "nat 1 = Suc 0" |
|
182 |
by (subst nat_eq_iff, simp) |
|
183 |
||
184 |
lemma nat_2: "nat 2 = Suc (Suc 0)" |
|
185 |
by (subst nat_eq_iff, simp) |
|
186 |
||
| 16413 | 187 |
lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)" |
188 |
apply (insert zless_nat_conj [of 1 z]) |
|
189 |
apply (auto simp add: nat_1) |
|
190 |
done |
|
191 |
||
| 14259 | 192 |
text{*This simplifies expressions of the form @{term "int n = z"} where
|
193 |
z is an integer literal.*} |
|
| 17085 | 194 |
lemmas int_eq_iff_number_of = int_eq_iff [of _ "number_of v", standard] |
195 |
declare int_eq_iff_number_of [simp] |
|
196 |
||
|
13837
8dd150d36c65
Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents:
13685
diff
changeset
|
197 |
|
| 14295 | 198 |
lemma split_nat [arith_split]: |
| 14259 | 199 |
"P(nat(i::int)) = ((\<forall>n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))" |
| 13575 | 200 |
(is "?P = (?L & ?R)") |
201 |
proof (cases "i < 0") |
|
202 |
case True thus ?thesis by simp |
|
203 |
next |
|
204 |
case False |
|
205 |
have "?P = ?L" |
|
206 |
proof |
|
207 |
assume ?P thus ?L using False by clarsimp |
|
208 |
next |
|
209 |
assume ?L thus ?P using False by simp |
|
210 |
qed |
|
211 |
with False show ?thesis by simp |
|
212 |
qed |
|
213 |
||
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
214 |
|
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
215 |
(*Analogous to zadd_int*) |
| 15013 | 216 |
lemma zdiff_int: "n \<le> m ==> int m - int n = int (m-n)" |
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
217 |
by (induct m n rule: diff_induct, simp_all) |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
218 |
|
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
219 |
lemma nat_mult_distrib: "(0::int) \<le> z ==> nat (z*z') = nat z * nat z'" |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
220 |
apply (case_tac "0 \<le> z'") |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
221 |
apply (rule inj_int [THEN injD]) |
| 16413 | 222 |
apply (simp add: int_mult zero_le_mult_iff) |
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
223 |
apply (simp add: mult_le_0_iff) |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
224 |
done |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
225 |
|
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
226 |
lemma nat_mult_distrib_neg: "z \<le> (0::int) ==> nat(z*z') = nat(-z) * nat(-z')" |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
227 |
apply (rule trans) |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
228 |
apply (rule_tac [2] nat_mult_distrib, auto) |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
229 |
done |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
230 |
|
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
231 |
lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)" |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
232 |
apply (case_tac "z=0 | w=0") |
| 15003 | 233 |
apply (auto simp add: abs_if nat_mult_distrib [symmetric] |
|
14479
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
234 |
nat_mult_distrib_neg [symmetric] mult_less_0_iff) |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
235 |
done |
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
236 |
|
|
0eca4aabf371
streamlined treatment of quotients for the integers
paulson
parents:
14473
diff
changeset
|
237 |
|
| 17472 | 238 |
subsection "Induction principles for int" |
| 13685 | 239 |
|
| 22059 | 240 |
text{*Well-founded segments of the integers*}
|
241 |
||
242 |
definition |
|
243 |
int_ge_less_than :: "int => (int * int) set" |
|
244 |
where |
|
245 |
"int_ge_less_than d = {(z',z). d \<le> z' & z' < z}"
|
|
246 |
||
247 |
theorem wf_int_ge_less_than: "wf (int_ge_less_than d)" |
|
248 |
proof - |
|
249 |
have "int_ge_less_than d \<subseteq> measure (%z. nat (z-d))" |
|
250 |
by (auto simp add: int_ge_less_than_def) |
|
251 |
thus ?thesis |
|
252 |
by (rule wf_subset [OF wf_measure]) |
|
253 |
qed |
|
254 |
||
255 |
text{*This variant looks odd, but is typical of the relations suggested
|
|
256 |
by RankFinder.*} |
|
257 |
||
258 |
definition |
|
259 |
int_ge_less_than2 :: "int => (int * int) set" |
|
260 |
where |
|
261 |
"int_ge_less_than2 d = {(z',z). d \<le> z & z' < z}"
|
|
262 |
||
263 |
theorem wf_int_ge_less_than2: "wf (int_ge_less_than2 d)" |
|
264 |
proof - |
|
265 |
have "int_ge_less_than2 d \<subseteq> measure (%z. nat (1+z-d))" |
|
266 |
by (auto simp add: int_ge_less_than2_def) |
|
267 |
thus ?thesis |
|
268 |
by (rule wf_subset [OF wf_measure]) |
|
269 |
qed |
|
270 |
||
| 13685 | 271 |
(* `set:int': dummy construction *) |
272 |
theorem int_ge_induct[case_names base step,induct set:int]: |
|
273 |
assumes ge: "k \<le> (i::int)" and |
|
274 |
base: "P(k)" and |
|
275 |
step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)" |
|
276 |
shows "P i" |
|
277 |
proof - |
|
|
14272
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
278 |
{ fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i"
|
| 13685 | 279 |
proof (induct n) |
280 |
case 0 |
|
281 |
hence "i = k" by arith |
|
282 |
thus "P i" using base by simp |
|
283 |
next |
|
284 |
case (Suc n) |
|
285 |
hence "n = nat((i - 1) - k)" by arith |
|
286 |
moreover |
|
287 |
have ki1: "k \<le> i - 1" using Suc.prems by arith |
|
288 |
ultimately |
|
289 |
have "P(i - 1)" by(rule Suc.hyps) |
|
290 |
from step[OF ki1 this] show ?case by simp |
|
291 |
qed |
|
292 |
} |
|
| 14473 | 293 |
with ge show ?thesis by fast |
| 13685 | 294 |
qed |
295 |
||
296 |
(* `set:int': dummy construction *) |
|
297 |
theorem int_gr_induct[case_names base step,induct set:int]: |
|
298 |
assumes gr: "k < (i::int)" and |
|
299 |
base: "P(k+1)" and |
|
300 |
step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)" |
|
301 |
shows "P i" |
|
302 |
apply(rule int_ge_induct[of "k + 1"]) |
|
303 |
using gr apply arith |
|
304 |
apply(rule base) |
|
| 14259 | 305 |
apply (rule step, simp+) |
| 13685 | 306 |
done |
307 |
||
308 |
theorem int_le_induct[consumes 1,case_names base step]: |
|
309 |
assumes le: "i \<le> (k::int)" and |
|
310 |
base: "P(k)" and |
|
311 |
step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)" |
|
312 |
shows "P i" |
|
313 |
proof - |
|
|
14272
5efbb548107d
Tidying of the integer development; towards removing the
paulson
parents:
14271
diff
changeset
|
314 |
{ fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i"
|
| 13685 | 315 |
proof (induct n) |
316 |
case 0 |
|
317 |
hence "i = k" by arith |
|
318 |
thus "P i" using base by simp |
|
319 |
next |
|
320 |
case (Suc n) |
|
321 |
hence "n = nat(k - (i+1))" by arith |
|
322 |
moreover |
|
323 |
have ki1: "i + 1 \<le> k" using Suc.prems by arith |
|
324 |
ultimately |
|
325 |
have "P(i+1)" by(rule Suc.hyps) |
|
326 |
from step[OF ki1 this] show ?case by simp |
|
327 |
qed |
|
328 |
} |
|
| 14473 | 329 |
with le show ?thesis by fast |
| 13685 | 330 |
qed |
331 |
||
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
332 |
theorem int_less_induct [consumes 1,case_names base step]: |
| 13685 | 333 |
assumes less: "(i::int) < k" and |
334 |
base: "P(k - 1)" and |
|
335 |
step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)" |
|
336 |
shows "P i" |
|
337 |
apply(rule int_le_induct[of _ "k - 1"]) |
|
338 |
using less apply arith |
|
339 |
apply(rule base) |
|
| 14259 | 340 |
apply (rule step, simp+) |
341 |
done |
|
342 |
||
343 |
subsection{*Intermediate value theorems*}
|
|
344 |
||
345 |
lemma int_val_lemma: |
|
346 |
"(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) --> |
|
347 |
f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))" |
|
| 14271 | 348 |
apply (induct_tac "n", simp) |
| 14259 | 349 |
apply (intro strip) |
350 |
apply (erule impE, simp) |
|
351 |
apply (erule_tac x = n in allE, simp) |
|
352 |
apply (case_tac "k = f (n+1) ") |
|
353 |
apply force |
|
354 |
apply (erule impE) |
|
| 15003 | 355 |
apply (simp add: abs_if split add: split_if_asm) |
| 14259 | 356 |
apply (blast intro: le_SucI) |
357 |
done |
|
358 |
||
359 |
lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)] |
|
360 |
||
361 |
lemma nat_intermed_int_val: |
|
362 |
"[| \<forall>i. m \<le> i & i < n --> abs(f(i + 1::nat) - f i) \<le> 1; m < n; |
|
363 |
f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)" |
|
364 |
apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k |
|
365 |
in int_val_lemma) |
|
366 |
apply simp |
|
367 |
apply (erule exE) |
|
368 |
apply (rule_tac x = "i+m" in exI, arith) |
|
369 |
done |
|
370 |
||
371 |
||
372 |
subsection{*Products and 1, by T. M. Rasmussen*}
|
|
373 |
||
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
374 |
lemma zabs_less_one_iff [simp]: "(\<bar>z\<bar> < 1) = (z = (0::int))" |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
375 |
by arith |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
376 |
|
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
377 |
lemma abs_zmult_eq_1: "(\<bar>m * n\<bar> = 1) ==> \<bar>m\<bar> = (1::int)" |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
378 |
apply (case_tac "\<bar>n\<bar>=1") |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
379 |
apply (simp add: abs_mult) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
380 |
apply (rule ccontr) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
381 |
apply (auto simp add: linorder_neq_iff abs_mult) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
382 |
apply (subgoal_tac "2 \<le> \<bar>m\<bar> & 2 \<le> \<bar>n\<bar>") |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
383 |
prefer 2 apply arith |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
384 |
apply (subgoal_tac "2*2 \<le> \<bar>m\<bar> * \<bar>n\<bar>", simp) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
385 |
apply (rule mult_mono, auto) |
| 13685 | 386 |
done |
387 |
||
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
388 |
lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1" |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
389 |
by (insert abs_zmult_eq_1 [of m n], arith) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
390 |
|
| 14259 | 391 |
lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)" |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
392 |
apply (auto dest: pos_zmult_eq_1_iff_lemma) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
393 |
apply (simp add: mult_commute [of m]) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
394 |
apply (frule pos_zmult_eq_1_iff_lemma, auto) |
| 14259 | 395 |
done |
396 |
||
397 |
lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))" |
|
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
398 |
apply (rule iffI) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
399 |
apply (frule pos_zmult_eq_1_iff_lemma) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
400 |
apply (simp add: mult_commute [of m]) |
|
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15140
diff
changeset
|
401 |
apply (frule pos_zmult_eq_1_iff_lemma, auto) |
| 14259 | 402 |
done |
403 |
||
| 20355 | 404 |
|
405 |
subsection {* code generator setup *}
|
|
406 |
||
| 21191 | 407 |
code_modulename SML |
408 |
Numeral Integer |
|
| 20355 | 409 |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
410 |
code_modulename OCaml |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
411 |
Numeral Integer |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
412 |
|
| 20595 | 413 |
lemma Numeral_Pls_refl [code func]: |
414 |
"Numeral.Pls = Numeral.Pls" .. |
|
|
20402
e2c6d096af01
more concise preprocessing of numerals for code generation
haftmann
parents:
20380
diff
changeset
|
415 |
|
| 20595 | 416 |
lemma Numeral_Min_refl [code func]: |
417 |
"Numeral.Min = Numeral.Min" .. |
|
|
20402
e2c6d096af01
more concise preprocessing of numerals for code generation
haftmann
parents:
20380
diff
changeset
|
418 |
|
| 20595 | 419 |
lemma zero_int_refl [code func]: |
420 |
"(0\<Colon>int) = 0" .. |
|
| 20355 | 421 |
|
| 20595 | 422 |
lemma one_int_refl [code func]: |
423 |
"(1\<Colon>int) = 1" .. |
|
| 20355 | 424 |
|
| 20595 | 425 |
lemma number_of_int_refl [code func]: |
426 |
"(number_of \<Colon> int \<Rightarrow> int) = number_of" .. |
|
427 |
||
428 |
lemma number_of_is_id: |
|
| 20485 | 429 |
"number_of (k::int) = k" |
430 |
unfolding int_number_of_def by simp |
|
| 20355 | 431 |
|
| 21060 | 432 |
lemma zero_is_num_zero [code inline, symmetric, normal post]: |
| 20595 | 433 |
"(0::int) = number_of Numeral.Pls" |
434 |
by simp |
|
435 |
||
| 21060 | 436 |
lemma one_is_num_one [code inline, symmetric, normal post]: |
| 20595 | 437 |
"(1::int) = number_of (Numeral.Pls BIT bit.B1)" |
438 |
by simp |
|
439 |
||
440 |
lemmas int_code_rewrites = |
|
441 |
arith_simps(5-27) |
|
| 20900 | 442 |
arith_extra_simps(1-5) [where 'a = int] |
| 20595 | 443 |
|
444 |
declare int_code_rewrites [code func] |
|
| 20355 | 445 |
|
| 20699 | 446 |
code_type bit |
| 21113 | 447 |
(SML "bool") |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
448 |
(OCaml "bool") |
| 21113 | 449 |
(Haskell "Bool") |
| 20699 | 450 |
code_const "Numeral.bit.B0" and "Numeral.bit.B1" |
| 21113 | 451 |
(SML "false" and "true") |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
452 |
(OCaml "false" and "true") |
| 21113 | 453 |
(Haskell "False" and "True") |
| 20699 | 454 |
|
455 |
code_const "number_of \<Colon> int \<Rightarrow> int" |
|
456 |
and "Numeral.Pls" and "Numeral.Min" and "Numeral.Bit" |
|
457 |
and "Numeral.succ" and "Numeral.pred" |
|
| 20595 | 458 |
(SML "_" |
| 21113 | 459 |
and "0/ :/ IntInf.int" |
460 |
and "~1/ :/ IntInf.int" |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
461 |
and "!(_; _; raise Fail \"BIT\")" |
| 21113 | 462 |
and "IntInf.+/ (_,/ 1)" |
463 |
and "IntInf.-/ (_,/ 1))") |
|
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
464 |
(OCaml "_" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
465 |
and "Big'_int.big'_int'_of'_int/ 0" |
| 22176 | 466 |
and "Big'_int.big'_int'_of'_int/ (-1)" |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
467 |
and "!(_; _; failwith \"BIT\")" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
468 |
and "Big'_int.succ'_big'_int" |
|
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
469 |
and "Big'_int.pred'_big'_int") |
| 20595 | 470 |
(Haskell "_" |
| 21113 | 471 |
and "0" |
|
21911
e29bcab0c81c
added OCaml code generation (without dictionaries)
haftmann
parents:
21872
diff
changeset
|
472 |
and "!(-1)" |
| 21113 | 473 |
and "error/ \"BIT\"" |
474 |
and "(+)/ 1" |
|
475 |
and "(-)/ _/ 1") |
|
| 20485 | 476 |
|
| 20355 | 477 |
setup {*
|
|
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21191
diff
changeset
|
478 |
CodegenPackage.add_appconst ("Numeral.Bit", CodegenPackage.appgen_numeral (try HOLogic.dest_numeral))
|
| 20355 | 479 |
*} |
480 |
||
481 |
||
| 19601 | 482 |
subsection {* legacy ML bindings *}
|
483 |
||
| 14259 | 484 |
ML |
485 |
{*
|
|
486 |
val zle_diff1_eq = thm "zle_diff1_eq"; |
|
487 |
val zle_add1_eq_le = thm "zle_add1_eq_le"; |
|
488 |
val nonneg_eq_int = thm "nonneg_eq_int"; |
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
489 |
val abs_minus_one = thm "abs_minus_one"; |
| 14390 | 490 |
val of_int_number_of_eq = thm"of_int_number_of_eq"; |
| 14259 | 491 |
val nat_eq_iff = thm "nat_eq_iff"; |
492 |
val nat_eq_iff2 = thm "nat_eq_iff2"; |
|
493 |
val nat_less_iff = thm "nat_less_iff"; |
|
494 |
val int_eq_iff = thm "int_eq_iff"; |
|
495 |
val nat_0 = thm "nat_0"; |
|
496 |
val nat_1 = thm "nat_1"; |
|
497 |
val nat_2 = thm "nat_2"; |
|
498 |
val nat_less_eq_zless = thm "nat_less_eq_zless"; |
|
499 |
val nat_le_eq_zle = thm "nat_le_eq_zle"; |
|
500 |
||
501 |
val nat_intermed_int_val = thm "nat_intermed_int_val"; |
|
502 |
val pos_zmult_eq_1_iff = thm "pos_zmult_eq_1_iff"; |
|
503 |
val zmult_eq_1_iff = thm "zmult_eq_1_iff"; |
|
504 |
val nat_add_distrib = thm "nat_add_distrib"; |
|
505 |
val nat_diff_distrib = thm "nat_diff_distrib"; |
|
506 |
val nat_mult_distrib = thm "nat_mult_distrib"; |
|
507 |
val nat_mult_distrib_neg = thm "nat_mult_distrib_neg"; |
|
508 |
val nat_abs_mult_distrib = thm "nat_abs_mult_distrib"; |
|
509 |
*} |
|
510 |
||
| 7707 | 511 |
end |