47108
|
1 |
(* Title: HOL/Num.thy
|
|
2 |
Author: Florian Haftmann
|
|
3 |
Author: Brian Huffman
|
|
4 |
*)
|
|
5 |
|
|
6 |
header {* Binary Numerals *}
|
|
7 |
|
|
8 |
theory Num
|
|
9 |
imports Datatype Power
|
|
10 |
begin
|
|
11 |
|
|
12 |
subsection {* The @{text num} type *}
|
|
13 |
|
|
14 |
datatype num = One | Bit0 num | Bit1 num
|
|
15 |
|
|
16 |
text {* Increment function for type @{typ num} *}
|
|
17 |
|
|
18 |
primrec inc :: "num \<Rightarrow> num" where
|
|
19 |
"inc One = Bit0 One" |
|
|
20 |
"inc (Bit0 x) = Bit1 x" |
|
|
21 |
"inc (Bit1 x) = Bit0 (inc x)"
|
|
22 |
|
|
23 |
text {* Converting between type @{typ num} and type @{typ nat} *}
|
|
24 |
|
|
25 |
primrec nat_of_num :: "num \<Rightarrow> nat" where
|
|
26 |
"nat_of_num One = Suc 0" |
|
|
27 |
"nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" |
|
|
28 |
"nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)"
|
|
29 |
|
|
30 |
primrec num_of_nat :: "nat \<Rightarrow> num" where
|
|
31 |
"num_of_nat 0 = One" |
|
|
32 |
"num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)"
|
|
33 |
|
|
34 |
lemma nat_of_num_pos: "0 < nat_of_num x"
|
|
35 |
by (induct x) simp_all
|
|
36 |
|
|
37 |
lemma nat_of_num_neq_0: " nat_of_num x \<noteq> 0"
|
|
38 |
by (induct x) simp_all
|
|
39 |
|
|
40 |
lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)"
|
|
41 |
by (induct x) simp_all
|
|
42 |
|
|
43 |
lemma num_of_nat_double:
|
|
44 |
"0 < n \<Longrightarrow> num_of_nat (n + n) = Bit0 (num_of_nat n)"
|
|
45 |
by (induct n) simp_all
|
|
46 |
|
|
47 |
text {*
|
|
48 |
Type @{typ num} is isomorphic to the strictly positive
|
|
49 |
natural numbers.
|
|
50 |
*}
|
|
51 |
|
|
52 |
lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x"
|
|
53 |
by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos)
|
|
54 |
|
|
55 |
lemma num_of_nat_inverse: "0 < n \<Longrightarrow> nat_of_num (num_of_nat n) = n"
|
|
56 |
by (induct n) (simp_all add: nat_of_num_inc)
|
|
57 |
|
|
58 |
lemma num_eq_iff: "x = y \<longleftrightarrow> nat_of_num x = nat_of_num y"
|
|
59 |
apply safe
|
|
60 |
apply (drule arg_cong [where f=num_of_nat])
|
|
61 |
apply (simp add: nat_of_num_inverse)
|
|
62 |
done
|
|
63 |
|
|
64 |
lemma num_induct [case_names One inc]:
|
|
65 |
fixes P :: "num \<Rightarrow> bool"
|
|
66 |
assumes One: "P One"
|
|
67 |
and inc: "\<And>x. P x \<Longrightarrow> P (inc x)"
|
|
68 |
shows "P x"
|
|
69 |
proof -
|
|
70 |
obtain n where n: "Suc n = nat_of_num x"
|
|
71 |
by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0)
|
|
72 |
have "P (num_of_nat (Suc n))"
|
|
73 |
proof (induct n)
|
|
74 |
case 0 show ?case using One by simp
|
|
75 |
next
|
|
76 |
case (Suc n)
|
|
77 |
then have "P (inc (num_of_nat (Suc n)))" by (rule inc)
|
|
78 |
then show "P (num_of_nat (Suc (Suc n)))" by simp
|
|
79 |
qed
|
|
80 |
with n show "P x"
|
|
81 |
by (simp add: nat_of_num_inverse)
|
|
82 |
qed
|
|
83 |
|
|
84 |
text {*
|
|
85 |
From now on, there are two possible models for @{typ num}:
|
|
86 |
as positive naturals (rule @{text "num_induct"})
|
|
87 |
and as digit representation (rules @{text "num.induct"}, @{text "num.cases"}).
|
|
88 |
*}
|
|
89 |
|
|
90 |
|
|
91 |
subsection {* Numeral operations *}
|
|
92 |
|
|
93 |
instantiation num :: "{plus,times,linorder}"
|
|
94 |
begin
|
|
95 |
|
|
96 |
definition [code del]:
|
|
97 |
"m + n = num_of_nat (nat_of_num m + nat_of_num n)"
|
|
98 |
|
|
99 |
definition [code del]:
|
|
100 |
"m * n = num_of_nat (nat_of_num m * nat_of_num n)"
|
|
101 |
|
|
102 |
definition [code del]:
|
|
103 |
"m \<le> n \<longleftrightarrow> nat_of_num m \<le> nat_of_num n"
|
|
104 |
|
|
105 |
definition [code del]:
|
|
106 |
"m < n \<longleftrightarrow> nat_of_num m < nat_of_num n"
|
|
107 |
|
|
108 |
instance
|
|
109 |
by (default, auto simp add: less_num_def less_eq_num_def num_eq_iff)
|
|
110 |
|
|
111 |
end
|
|
112 |
|
|
113 |
lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y"
|
|
114 |
unfolding plus_num_def
|
|
115 |
by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos)
|
|
116 |
|
|
117 |
lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y"
|
|
118 |
unfolding times_num_def
|
|
119 |
by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos)
|
|
120 |
|
|
121 |
lemma add_num_simps [simp, code]:
|
|
122 |
"One + One = Bit0 One"
|
|
123 |
"One + Bit0 n = Bit1 n"
|
|
124 |
"One + Bit1 n = Bit0 (n + One)"
|
|
125 |
"Bit0 m + One = Bit1 m"
|
|
126 |
"Bit0 m + Bit0 n = Bit0 (m + n)"
|
|
127 |
"Bit0 m + Bit1 n = Bit1 (m + n)"
|
|
128 |
"Bit1 m + One = Bit0 (m + One)"
|
|
129 |
"Bit1 m + Bit0 n = Bit1 (m + n)"
|
|
130 |
"Bit1 m + Bit1 n = Bit0 (m + n + One)"
|
|
131 |
by (simp_all add: num_eq_iff nat_of_num_add)
|
|
132 |
|
|
133 |
lemma mult_num_simps [simp, code]:
|
|
134 |
"m * One = m"
|
|
135 |
"One * n = n"
|
|
136 |
"Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))"
|
|
137 |
"Bit0 m * Bit1 n = Bit0 (m * Bit1 n)"
|
|
138 |
"Bit1 m * Bit0 n = Bit0 (Bit1 m * n)"
|
|
139 |
"Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))"
|
|
140 |
by (simp_all add: num_eq_iff nat_of_num_add
|
|
141 |
nat_of_num_mult left_distrib right_distrib)
|
|
142 |
|
|
143 |
lemma eq_num_simps:
|
|
144 |
"One = One \<longleftrightarrow> True"
|
|
145 |
"One = Bit0 n \<longleftrightarrow> False"
|
|
146 |
"One = Bit1 n \<longleftrightarrow> False"
|
|
147 |
"Bit0 m = One \<longleftrightarrow> False"
|
|
148 |
"Bit1 m = One \<longleftrightarrow> False"
|
|
149 |
"Bit0 m = Bit0 n \<longleftrightarrow> m = n"
|
|
150 |
"Bit0 m = Bit1 n \<longleftrightarrow> False"
|
|
151 |
"Bit1 m = Bit0 n \<longleftrightarrow> False"
|
|
152 |
"Bit1 m = Bit1 n \<longleftrightarrow> m = n"
|
|
153 |
by simp_all
|
|
154 |
|
|
155 |
lemma le_num_simps [simp, code]:
|
|
156 |
"One \<le> n \<longleftrightarrow> True"
|
|
157 |
"Bit0 m \<le> One \<longleftrightarrow> False"
|
|
158 |
"Bit1 m \<le> One \<longleftrightarrow> False"
|
|
159 |
"Bit0 m \<le> Bit0 n \<longleftrightarrow> m \<le> n"
|
|
160 |
"Bit0 m \<le> Bit1 n \<longleftrightarrow> m \<le> n"
|
|
161 |
"Bit1 m \<le> Bit1 n \<longleftrightarrow> m \<le> n"
|
|
162 |
"Bit1 m \<le> Bit0 n \<longleftrightarrow> m < n"
|
|
163 |
using nat_of_num_pos [of n] nat_of_num_pos [of m]
|
|
164 |
by (auto simp add: less_eq_num_def less_num_def)
|
|
165 |
|
|
166 |
lemma less_num_simps [simp, code]:
|
|
167 |
"m < One \<longleftrightarrow> False"
|
|
168 |
"One < Bit0 n \<longleftrightarrow> True"
|
|
169 |
"One < Bit1 n \<longleftrightarrow> True"
|
|
170 |
"Bit0 m < Bit0 n \<longleftrightarrow> m < n"
|
|
171 |
"Bit0 m < Bit1 n \<longleftrightarrow> m \<le> n"
|
|
172 |
"Bit1 m < Bit1 n \<longleftrightarrow> m < n"
|
|
173 |
"Bit1 m < Bit0 n \<longleftrightarrow> m < n"
|
|
174 |
using nat_of_num_pos [of n] nat_of_num_pos [of m]
|
|
175 |
by (auto simp add: less_eq_num_def less_num_def)
|
|
176 |
|
|
177 |
text {* Rules using @{text One} and @{text inc} as constructors *}
|
|
178 |
|
|
179 |
lemma add_One: "x + One = inc x"
|
|
180 |
by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
|
|
181 |
|
|
182 |
lemma add_One_commute: "One + n = n + One"
|
|
183 |
by (induct n) simp_all
|
|
184 |
|
|
185 |
lemma add_inc: "x + inc y = inc (x + y)"
|
|
186 |
by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
|
|
187 |
|
|
188 |
lemma mult_inc: "x * inc y = x * y + x"
|
|
189 |
by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc)
|
|
190 |
|
|
191 |
text {* The @{const num_of_nat} conversion *}
|
|
192 |
|
|
193 |
lemma num_of_nat_One:
|
|
194 |
"n \<le> 1 \<Longrightarrow> num_of_nat n = Num.One"
|
|
195 |
by (cases n) simp_all
|
|
196 |
|
|
197 |
lemma num_of_nat_plus_distrib:
|
|
198 |
"0 < m \<Longrightarrow> 0 < n \<Longrightarrow> num_of_nat (m + n) = num_of_nat m + num_of_nat n"
|
|
199 |
by (induct n) (auto simp add: add_One add_One_commute add_inc)
|
|
200 |
|
|
201 |
text {* A double-and-decrement function *}
|
|
202 |
|
|
203 |
primrec BitM :: "num \<Rightarrow> num" where
|
|
204 |
"BitM One = One" |
|
|
205 |
"BitM (Bit0 n) = Bit1 (BitM n)" |
|
|
206 |
"BitM (Bit1 n) = Bit1 (Bit0 n)"
|
|
207 |
|
|
208 |
lemma BitM_plus_one: "BitM n + One = Bit0 n"
|
|
209 |
by (induct n) simp_all
|
|
210 |
|
|
211 |
lemma one_plus_BitM: "One + BitM n = Bit0 n"
|
|
212 |
unfolding add_One_commute BitM_plus_one ..
|
|
213 |
|
|
214 |
text {* Squaring and exponentiation *}
|
|
215 |
|
|
216 |
primrec sqr :: "num \<Rightarrow> num" where
|
|
217 |
"sqr One = One" |
|
|
218 |
"sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" |
|
|
219 |
"sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))"
|
|
220 |
|
|
221 |
primrec pow :: "num \<Rightarrow> num \<Rightarrow> num" where
|
|
222 |
"pow x One = x" |
|
|
223 |
"pow x (Bit0 y) = sqr (pow x y)" |
|
|
224 |
"pow x (Bit1 y) = x * sqr (pow x y)"
|
|
225 |
|
|
226 |
lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x"
|
|
227 |
by (induct x, simp_all add: algebra_simps nat_of_num_add)
|
|
228 |
|
|
229 |
lemma sqr_conv_mult: "sqr x = x * x"
|
|
230 |
by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult)
|
|
231 |
|
|
232 |
|
|
233 |
subsection {* Numary numerals *}
|
|
234 |
|
|
235 |
text {*
|
|
236 |
We embed numary representations into a generic algebraic
|
|
237 |
structure using @{text numeral}.
|
|
238 |
*}
|
|
239 |
|
|
240 |
class numeral = one + semigroup_add
|
|
241 |
begin
|
|
242 |
|
|
243 |
primrec numeral :: "num \<Rightarrow> 'a" where
|
|
244 |
numeral_One: "numeral One = 1" |
|
|
245 |
numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" |
|
|
246 |
numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1"
|
|
247 |
|
|
248 |
lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1"
|
|
249 |
apply (induct x)
|
|
250 |
apply simp
|
|
251 |
apply (simp add: add_assoc [symmetric], simp add: add_assoc)
|
|
252 |
apply (simp add: add_assoc [symmetric], simp add: add_assoc)
|
|
253 |
done
|
|
254 |
|
|
255 |
lemma numeral_inc: "numeral (inc x) = numeral x + 1"
|
|
256 |
proof (induct x)
|
|
257 |
case (Bit1 x)
|
|
258 |
have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1"
|
|
259 |
by (simp only: one_plus_numeral_commute)
|
|
260 |
with Bit1 show ?case
|
|
261 |
by (simp add: add_assoc)
|
|
262 |
qed simp_all
|
|
263 |
|
|
264 |
declare numeral.simps [simp del]
|
|
265 |
|
|
266 |
abbreviation "Numeral1 \<equiv> numeral One"
|
|
267 |
|
|
268 |
declare numeral_One [code_post]
|
|
269 |
|
|
270 |
end
|
|
271 |
|
|
272 |
text {* Negative numerals. *}
|
|
273 |
|
|
274 |
class neg_numeral = numeral + group_add
|
|
275 |
begin
|
|
276 |
|
|
277 |
definition neg_numeral :: "num \<Rightarrow> 'a" where
|
|
278 |
"neg_numeral k = - numeral k"
|
|
279 |
|
|
280 |
end
|
|
281 |
|
|
282 |
text {* Numeral syntax. *}
|
|
283 |
|
|
284 |
syntax
|
|
285 |
"_Numeral" :: "num_const \<Rightarrow> 'a" ("_")
|
|
286 |
|
|
287 |
parse_translation {*
|
|
288 |
let
|
|
289 |
fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2)
|
|
290 |
of (0, 1) => Syntax.const @{const_name One}
|
|
291 |
| (n, 0) => Syntax.const @{const_name Bit0} $ num_of_int n
|
|
292 |
| (n, 1) => Syntax.const @{const_name Bit1} $ num_of_int n
|
|
293 |
else raise Match;
|
|
294 |
val pos = Syntax.const @{const_name numeral}
|
|
295 |
val neg = Syntax.const @{const_name neg_numeral}
|
|
296 |
val one = Syntax.const @{const_name Groups.one}
|
|
297 |
val zero = Syntax.const @{const_name Groups.zero}
|
|
298 |
fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
|
|
299 |
c $ numeral_tr [t] $ u
|
|
300 |
| numeral_tr [Const (num, _)] =
|
|
301 |
let
|
|
302 |
val {value, ...} = Lexicon.read_xnum num;
|
|
303 |
in
|
|
304 |
if value = 0 then zero else
|
|
305 |
if value > 0
|
|
306 |
then pos $ num_of_int value
|
|
307 |
else neg $ num_of_int (~value)
|
|
308 |
end
|
|
309 |
| numeral_tr ts = raise TERM ("numeral_tr", ts);
|
|
310 |
in [("_Numeral", numeral_tr)] end
|
|
311 |
*}
|
|
312 |
|
|
313 |
typed_print_translation (advanced) {*
|
|
314 |
let
|
|
315 |
fun dest_num (Const (@{const_syntax Bit0}, _) $ n) = 2 * dest_num n
|
|
316 |
| dest_num (Const (@{const_syntax Bit1}, _) $ n) = 2 * dest_num n + 1
|
|
317 |
| dest_num (Const (@{const_syntax One}, _)) = 1;
|
|
318 |
fun num_tr' sign ctxt T [n] =
|
|
319 |
let
|
|
320 |
val k = dest_num n;
|
|
321 |
val t' = Syntax.const @{syntax_const "_Numeral"} $
|
|
322 |
Syntax.free (sign ^ string_of_int k);
|
|
323 |
in
|
|
324 |
case T of
|
|
325 |
Type (@{type_name fun}, [_, T']) =>
|
|
326 |
if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t'
|
|
327 |
else Syntax.const @{syntax_const "_constrain"} $ t' $ Syntax_Phases.term_of_typ ctxt T'
|
|
328 |
| T' => if T' = dummyT then t' else raise Match
|
|
329 |
end;
|
|
330 |
in [(@{const_syntax numeral}, num_tr' ""),
|
|
331 |
(@{const_syntax neg_numeral}, num_tr' "-")] end
|
|
332 |
*}
|
|
333 |
|
|
334 |
subsection {* Class-specific numeral rules *}
|
|
335 |
|
|
336 |
text {*
|
|
337 |
@{const numeral} is a morphism.
|
|
338 |
*}
|
|
339 |
|
|
340 |
subsubsection {* Structures with addition: class @{text numeral} *}
|
|
341 |
|
|
342 |
context numeral
|
|
343 |
begin
|
|
344 |
|
|
345 |
lemma numeral_add: "numeral (m + n) = numeral m + numeral n"
|
|
346 |
by (induct n rule: num_induct)
|
|
347 |
(simp_all only: numeral_One add_One add_inc numeral_inc add_assoc)
|
|
348 |
|
|
349 |
lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)"
|
|
350 |
by (rule numeral_add [symmetric])
|
|
351 |
|
|
352 |
lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)"
|
|
353 |
using numeral_add [of n One] by (simp add: numeral_One)
|
|
354 |
|
|
355 |
lemma one_plus_numeral: "1 + numeral n = numeral (One + n)"
|
|
356 |
using numeral_add [of One n] by (simp add: numeral_One)
|
|
357 |
|
|
358 |
lemma one_add_one: "1 + 1 = 2"
|
|
359 |
using numeral_add [of One One] by (simp add: numeral_One)
|
|
360 |
|
|
361 |
lemmas add_numeral_special =
|
|
362 |
numeral_plus_one one_plus_numeral one_add_one
|
|
363 |
|
|
364 |
end
|
|
365 |
|
|
366 |
subsubsection {*
|
|
367 |
Structures with negation: class @{text neg_numeral}
|
|
368 |
*}
|
|
369 |
|
|
370 |
context neg_numeral
|
|
371 |
begin
|
|
372 |
|
|
373 |
text {* Numerals form an abelian subgroup. *}
|
|
374 |
|
|
375 |
inductive is_num :: "'a \<Rightarrow> bool" where
|
|
376 |
"is_num 1" |
|
|
377 |
"is_num x \<Longrightarrow> is_num (- x)" |
|
|
378 |
"\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> is_num (x + y)"
|
|
379 |
|
|
380 |
lemma is_num_numeral: "is_num (numeral k)"
|
|
381 |
by (induct k, simp_all add: numeral.simps is_num.intros)
|
|
382 |
|
|
383 |
lemma is_num_add_commute:
|
|
384 |
"\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> x + y = y + x"
|
|
385 |
apply (induct x rule: is_num.induct)
|
|
386 |
apply (induct y rule: is_num.induct)
|
|
387 |
apply simp
|
|
388 |
apply (rule_tac a=x in add_left_imp_eq)
|
|
389 |
apply (rule_tac a=x in add_right_imp_eq)
|
|
390 |
apply (simp add: add_assoc minus_add_cancel)
|
|
391 |
apply (simp add: add_assoc [symmetric], simp add: add_assoc)
|
|
392 |
apply (rule_tac a=x in add_left_imp_eq)
|
|
393 |
apply (rule_tac a=x in add_right_imp_eq)
|
|
394 |
apply (simp add: add_assoc minus_add_cancel add_minus_cancel)
|
|
395 |
apply (simp add: add_assoc, simp add: add_assoc [symmetric])
|
|
396 |
done
|
|
397 |
|
|
398 |
lemma is_num_add_left_commute:
|
|
399 |
"\<lbrakk>is_num x; is_num y\<rbrakk> \<Longrightarrow> x + (y + z) = y + (x + z)"
|
|
400 |
by (simp only: add_assoc [symmetric] is_num_add_commute)
|
|
401 |
|
|
402 |
lemmas is_num_normalize =
|
|
403 |
add_assoc is_num_add_commute is_num_add_left_commute
|
|
404 |
is_num.intros is_num_numeral
|
|
405 |
diff_minus minus_add add_minus_cancel minus_add_cancel
|
|
406 |
|
|
407 |
definition dbl :: "'a \<Rightarrow> 'a" where "dbl x = x + x"
|
|
408 |
definition dbl_inc :: "'a \<Rightarrow> 'a" where "dbl_inc x = x + x + 1"
|
|
409 |
definition dbl_dec :: "'a \<Rightarrow> 'a" where "dbl_dec x = x + x - 1"
|
|
410 |
|
|
411 |
definition sub :: "num \<Rightarrow> num \<Rightarrow> 'a" where
|
|
412 |
"sub k l = numeral k - numeral l"
|
|
413 |
|
|
414 |
lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1"
|
|
415 |
by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq)
|
|
416 |
|
|
417 |
lemma dbl_simps [simp]:
|
|
418 |
"dbl (neg_numeral k) = neg_numeral (Bit0 k)"
|
|
419 |
"dbl 0 = 0"
|
|
420 |
"dbl 1 = 2"
|
|
421 |
"dbl (numeral k) = numeral (Bit0 k)"
|
|
422 |
unfolding dbl_def neg_numeral_def numeral.simps
|
|
423 |
by (simp_all add: minus_add)
|
|
424 |
|
|
425 |
lemma dbl_inc_simps [simp]:
|
|
426 |
"dbl_inc (neg_numeral k) = neg_numeral (BitM k)"
|
|
427 |
"dbl_inc 0 = 1"
|
|
428 |
"dbl_inc 1 = 3"
|
|
429 |
"dbl_inc (numeral k) = numeral (Bit1 k)"
|
|
430 |
unfolding dbl_inc_def neg_numeral_def numeral.simps numeral_BitM
|
|
431 |
by (simp_all add: is_num_normalize)
|
|
432 |
|
|
433 |
lemma dbl_dec_simps [simp]:
|
|
434 |
"dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)"
|
|
435 |
"dbl_dec 0 = -1"
|
|
436 |
"dbl_dec 1 = 1"
|
|
437 |
"dbl_dec (numeral k) = numeral (BitM k)"
|
|
438 |
unfolding dbl_dec_def neg_numeral_def numeral.simps numeral_BitM
|
|
439 |
by (simp_all add: is_num_normalize)
|
|
440 |
|
|
441 |
lemma sub_num_simps [simp]:
|
|
442 |
"sub One One = 0"
|
|
443 |
"sub One (Bit0 l) = neg_numeral (BitM l)"
|
|
444 |
"sub One (Bit1 l) = neg_numeral (Bit0 l)"
|
|
445 |
"sub (Bit0 k) One = numeral (BitM k)"
|
|
446 |
"sub (Bit1 k) One = numeral (Bit0 k)"
|
|
447 |
"sub (Bit0 k) (Bit0 l) = dbl (sub k l)"
|
|
448 |
"sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)"
|
|
449 |
"sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)"
|
|
450 |
"sub (Bit1 k) (Bit1 l) = dbl (sub k l)"
|
|
451 |
unfolding dbl_def dbl_dec_def dbl_inc_def sub_def
|
|
452 |
unfolding neg_numeral_def numeral.simps numeral_BitM
|
|
453 |
by (simp_all add: is_num_normalize)
|
|
454 |
|
|
455 |
lemma add_neg_numeral_simps:
|
|
456 |
"numeral m + neg_numeral n = sub m n"
|
|
457 |
"neg_numeral m + numeral n = sub n m"
|
|
458 |
"neg_numeral m + neg_numeral n = neg_numeral (m + n)"
|
|
459 |
unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
|
|
460 |
by (simp_all add: is_num_normalize)
|
|
461 |
|
|
462 |
lemma add_neg_numeral_special:
|
|
463 |
"1 + neg_numeral m = sub One m"
|
|
464 |
"neg_numeral m + 1 = sub One m"
|
|
465 |
unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
|
|
466 |
by (simp_all add: is_num_normalize)
|
|
467 |
|
|
468 |
lemma diff_numeral_simps:
|
|
469 |
"numeral m - numeral n = sub m n"
|
|
470 |
"numeral m - neg_numeral n = numeral (m + n)"
|
|
471 |
"neg_numeral m - numeral n = neg_numeral (m + n)"
|
|
472 |
"neg_numeral m - neg_numeral n = sub n m"
|
|
473 |
unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
|
|
474 |
by (simp_all add: is_num_normalize)
|
|
475 |
|
|
476 |
lemma diff_numeral_special:
|
|
477 |
"1 - numeral n = sub One n"
|
|
478 |
"1 - neg_numeral n = numeral (One + n)"
|
|
479 |
"numeral m - 1 = sub m One"
|
|
480 |
"neg_numeral m - 1 = neg_numeral (m + One)"
|
|
481 |
unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
|
|
482 |
by (simp_all add: is_num_normalize)
|
|
483 |
|
|
484 |
lemma minus_one: "- 1 = -1"
|
|
485 |
unfolding neg_numeral_def numeral.simps ..
|
|
486 |
|
|
487 |
lemma minus_numeral: "- numeral n = neg_numeral n"
|
|
488 |
unfolding neg_numeral_def ..
|
|
489 |
|
|
490 |
lemma minus_neg_numeral: "- neg_numeral n = numeral n"
|
|
491 |
unfolding neg_numeral_def by simp
|
|
492 |
|
|
493 |
lemmas minus_numeral_simps [simp] =
|
|
494 |
minus_one minus_numeral minus_neg_numeral
|
|
495 |
|
|
496 |
end
|
|
497 |
|
|
498 |
subsubsection {*
|
|
499 |
Structures with multiplication: class @{text semiring_numeral}
|
|
500 |
*}
|
|
501 |
|
|
502 |
class semiring_numeral = semiring + monoid_mult
|
|
503 |
begin
|
|
504 |
|
|
505 |
subclass numeral ..
|
|
506 |
|
|
507 |
lemma numeral_mult: "numeral (m * n) = numeral m * numeral n"
|
|
508 |
apply (induct n rule: num_induct)
|
|
509 |
apply (simp add: numeral_One)
|
|
510 |
apply (simp add: mult_inc numeral_inc numeral_add numeral_inc right_distrib)
|
|
511 |
done
|
|
512 |
|
|
513 |
lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)"
|
|
514 |
by (rule numeral_mult [symmetric])
|
|
515 |
|
|
516 |
end
|
|
517 |
|
|
518 |
subsubsection {*
|
|
519 |
Structures with a zero: class @{text semiring_1}
|
|
520 |
*}
|
|
521 |
|
|
522 |
context semiring_1
|
|
523 |
begin
|
|
524 |
|
|
525 |
subclass semiring_numeral ..
|
|
526 |
|
|
527 |
lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n"
|
|
528 |
by (induct n,
|
|
529 |
simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1)
|
|
530 |
|
|
531 |
end
|
|
532 |
|
|
533 |
lemma nat_of_num_numeral: "nat_of_num = numeral"
|
|
534 |
proof
|
|
535 |
fix n
|
|
536 |
have "numeral n = nat_of_num n"
|
|
537 |
by (induct n) (simp_all add: numeral.simps)
|
|
538 |
then show "nat_of_num n = numeral n" by simp
|
|
539 |
qed
|
|
540 |
|
|
541 |
subsubsection {*
|
|
542 |
Equality: class @{text semiring_char_0}
|
|
543 |
*}
|
|
544 |
|
|
545 |
context semiring_char_0
|
|
546 |
begin
|
|
547 |
|
|
548 |
lemma numeral_eq_iff: "numeral m = numeral n \<longleftrightarrow> m = n"
|
|
549 |
unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric]
|
|
550 |
of_nat_eq_iff num_eq_iff ..
|
|
551 |
|
|
552 |
lemma numeral_eq_one_iff: "numeral n = 1 \<longleftrightarrow> n = One"
|
|
553 |
by (rule numeral_eq_iff [of n One, unfolded numeral_One])
|
|
554 |
|
|
555 |
lemma one_eq_numeral_iff: "1 = numeral n \<longleftrightarrow> One = n"
|
|
556 |
by (rule numeral_eq_iff [of One n, unfolded numeral_One])
|
|
557 |
|
|
558 |
lemma numeral_neq_zero: "numeral n \<noteq> 0"
|
|
559 |
unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric]
|
|
560 |
by (simp add: nat_of_num_pos)
|
|
561 |
|
|
562 |
lemma zero_neq_numeral: "0 \<noteq> numeral n"
|
|
563 |
unfolding eq_commute [of 0] by (rule numeral_neq_zero)
|
|
564 |
|
|
565 |
lemmas eq_numeral_simps [simp] =
|
|
566 |
numeral_eq_iff
|
|
567 |
numeral_eq_one_iff
|
|
568 |
one_eq_numeral_iff
|
|
569 |
numeral_neq_zero
|
|
570 |
zero_neq_numeral
|
|
571 |
|
|
572 |
end
|
|
573 |
|
|
574 |
subsubsection {*
|
|
575 |
Comparisons: class @{text linordered_semidom}
|
|
576 |
*}
|
|
577 |
|
|
578 |
text {* Could be perhaps more general than here. *}
|
|
579 |
|
|
580 |
context linordered_semidom
|
|
581 |
begin
|
|
582 |
|
|
583 |
lemma numeral_le_iff: "numeral m \<le> numeral n \<longleftrightarrow> m \<le> n"
|
|
584 |
proof -
|
|
585 |
have "of_nat (numeral m) \<le> of_nat (numeral n) \<longleftrightarrow> m \<le> n"
|
|
586 |
unfolding less_eq_num_def nat_of_num_numeral of_nat_le_iff ..
|
|
587 |
then show ?thesis by simp
|
|
588 |
qed
|
|
589 |
|
|
590 |
lemma one_le_numeral: "1 \<le> numeral n"
|
|
591 |
using numeral_le_iff [of One n] by (simp add: numeral_One)
|
|
592 |
|
|
593 |
lemma numeral_le_one_iff: "numeral n \<le> 1 \<longleftrightarrow> n \<le> One"
|
|
594 |
using numeral_le_iff [of n One] by (simp add: numeral_One)
|
|
595 |
|
|
596 |
lemma numeral_less_iff: "numeral m < numeral n \<longleftrightarrow> m < n"
|
|
597 |
proof -
|
|
598 |
have "of_nat (numeral m) < of_nat (numeral n) \<longleftrightarrow> m < n"
|
|
599 |
unfolding less_num_def nat_of_num_numeral of_nat_less_iff ..
|
|
600 |
then show ?thesis by simp
|
|
601 |
qed
|
|
602 |
|
|
603 |
lemma not_numeral_less_one: "\<not> numeral n < 1"
|
|
604 |
using numeral_less_iff [of n One] by (simp add: numeral_One)
|
|
605 |
|
|
606 |
lemma one_less_numeral_iff: "1 < numeral n \<longleftrightarrow> One < n"
|
|
607 |
using numeral_less_iff [of One n] by (simp add: numeral_One)
|
|
608 |
|
|
609 |
lemma zero_le_numeral: "0 \<le> numeral n"
|
|
610 |
by (induct n) (simp_all add: numeral.simps)
|
|
611 |
|
|
612 |
lemma zero_less_numeral: "0 < numeral n"
|
|
613 |
by (induct n) (simp_all add: numeral.simps add_pos_pos)
|
|
614 |
|
|
615 |
lemma not_numeral_le_zero: "\<not> numeral n \<le> 0"
|
|
616 |
by (simp add: not_le zero_less_numeral)
|
|
617 |
|
|
618 |
lemma not_numeral_less_zero: "\<not> numeral n < 0"
|
|
619 |
by (simp add: not_less zero_le_numeral)
|
|
620 |
|
|
621 |
lemmas le_numeral_extra =
|
|
622 |
zero_le_one not_one_le_zero
|
|
623 |
order_refl [of 0] order_refl [of 1]
|
|
624 |
|
|
625 |
lemmas less_numeral_extra =
|
|
626 |
zero_less_one not_one_less_zero
|
|
627 |
less_irrefl [of 0] less_irrefl [of 1]
|
|
628 |
|
|
629 |
lemmas le_numeral_simps [simp] =
|
|
630 |
numeral_le_iff
|
|
631 |
one_le_numeral
|
|
632 |
numeral_le_one_iff
|
|
633 |
zero_le_numeral
|
|
634 |
not_numeral_le_zero
|
|
635 |
|
|
636 |
lemmas less_numeral_simps [simp] =
|
|
637 |
numeral_less_iff
|
|
638 |
one_less_numeral_iff
|
|
639 |
not_numeral_less_one
|
|
640 |
zero_less_numeral
|
|
641 |
not_numeral_less_zero
|
|
642 |
|
|
643 |
end
|
|
644 |
|
|
645 |
subsubsection {*
|
|
646 |
Multiplication and negation: class @{text ring_1}
|
|
647 |
*}
|
|
648 |
|
|
649 |
context ring_1
|
|
650 |
begin
|
|
651 |
|
|
652 |
subclass neg_numeral ..
|
|
653 |
|
|
654 |
lemma mult_neg_numeral_simps:
|
|
655 |
"neg_numeral m * neg_numeral n = numeral (m * n)"
|
|
656 |
"neg_numeral m * numeral n = neg_numeral (m * n)"
|
|
657 |
"numeral m * neg_numeral n = neg_numeral (m * n)"
|
|
658 |
unfolding neg_numeral_def mult_minus_left mult_minus_right
|
|
659 |
by (simp_all only: minus_minus numeral_mult)
|
|
660 |
|
|
661 |
lemma mult_minus1 [simp]: "-1 * z = - z"
|
|
662 |
unfolding neg_numeral_def numeral.simps mult_minus_left by simp
|
|
663 |
|
|
664 |
lemma mult_minus1_right [simp]: "z * -1 = - z"
|
|
665 |
unfolding neg_numeral_def numeral.simps mult_minus_right by simp
|
|
666 |
|
|
667 |
end
|
|
668 |
|
|
669 |
subsubsection {*
|
|
670 |
Equality using @{text iszero} for rings with non-zero characteristic
|
|
671 |
*}
|
|
672 |
|
|
673 |
context ring_1
|
|
674 |
begin
|
|
675 |
|
|
676 |
definition iszero :: "'a \<Rightarrow> bool"
|
|
677 |
where "iszero z \<longleftrightarrow> z = 0"
|
|
678 |
|
|
679 |
lemma iszero_0 [simp]: "iszero 0"
|
|
680 |
by (simp add: iszero_def)
|
|
681 |
|
|
682 |
lemma not_iszero_1 [simp]: "\<not> iszero 1"
|
|
683 |
by (simp add: iszero_def)
|
|
684 |
|
|
685 |
lemma not_iszero_Numeral1: "\<not> iszero Numeral1"
|
|
686 |
by (simp add: numeral_One)
|
|
687 |
|
|
688 |
lemma iszero_neg_numeral [simp]:
|
|
689 |
"iszero (neg_numeral w) \<longleftrightarrow> iszero (numeral w)"
|
|
690 |
unfolding iszero_def neg_numeral_def
|
|
691 |
by (rule neg_equal_0_iff_equal)
|
|
692 |
|
|
693 |
lemma eq_iff_iszero_diff: "x = y \<longleftrightarrow> iszero (x - y)"
|
|
694 |
unfolding iszero_def by (rule eq_iff_diff_eq_0)
|
|
695 |
|
|
696 |
text {* The @{text "eq_numeral_iff_iszero"} lemmas are not declared
|
|
697 |
@{text "[simp]"} by default, because for rings of characteristic zero,
|
|
698 |
better simp rules are possible. For a type like integers mod @{text
|
|
699 |
"n"}, type-instantiated versions of these rules should be added to the
|
|
700 |
simplifier, along with a type-specific rule for deciding propositions
|
|
701 |
of the form @{text "iszero (numeral w)"}.
|
|
702 |
|
|
703 |
bh: Maybe it would not be so bad to just declare these as simp
|
|
704 |
rules anyway? I should test whether these rules take precedence over
|
|
705 |
the @{text "ring_char_0"} rules in the simplifier.
|
|
706 |
*}
|
|
707 |
|
|
708 |
lemma eq_numeral_iff_iszero:
|
|
709 |
"numeral x = numeral y \<longleftrightarrow> iszero (sub x y)"
|
|
710 |
"numeral x = neg_numeral y \<longleftrightarrow> iszero (numeral (x + y))"
|
|
711 |
"neg_numeral x = numeral y \<longleftrightarrow> iszero (numeral (x + y))"
|
|
712 |
"neg_numeral x = neg_numeral y \<longleftrightarrow> iszero (sub y x)"
|
|
713 |
"numeral x = 1 \<longleftrightarrow> iszero (sub x One)"
|
|
714 |
"1 = numeral y \<longleftrightarrow> iszero (sub One y)"
|
|
715 |
"neg_numeral x = 1 \<longleftrightarrow> iszero (numeral (x + One))"
|
|
716 |
"1 = neg_numeral y \<longleftrightarrow> iszero (numeral (One + y))"
|
|
717 |
"numeral x = 0 \<longleftrightarrow> iszero (numeral x)"
|
|
718 |
"0 = numeral y \<longleftrightarrow> iszero (numeral y)"
|
|
719 |
"neg_numeral x = 0 \<longleftrightarrow> iszero (numeral x)"
|
|
720 |
"0 = neg_numeral y \<longleftrightarrow> iszero (numeral y)"
|
|
721 |
unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special
|
|
722 |
by simp_all
|
|
723 |
|
|
724 |
end
|
|
725 |
|
|
726 |
subsubsection {*
|
|
727 |
Equality and negation: class @{text ring_char_0}
|
|
728 |
*}
|
|
729 |
|
|
730 |
class ring_char_0 = ring_1 + semiring_char_0
|
|
731 |
begin
|
|
732 |
|
|
733 |
lemma not_iszero_numeral [simp]: "\<not> iszero (numeral w)"
|
|
734 |
by (simp add: iszero_def)
|
|
735 |
|
|
736 |
lemma neg_numeral_eq_iff: "neg_numeral m = neg_numeral n \<longleftrightarrow> m = n"
|
|
737 |
by (simp only: neg_numeral_def neg_equal_iff_equal numeral_eq_iff)
|
|
738 |
|
|
739 |
lemma numeral_neq_neg_numeral: "numeral m \<noteq> neg_numeral n"
|
|
740 |
unfolding neg_numeral_def eq_neg_iff_add_eq_0
|
|
741 |
by (simp add: numeral_plus_numeral)
|
|
742 |
|
|
743 |
lemma neg_numeral_neq_numeral: "neg_numeral m \<noteq> numeral n"
|
|
744 |
by (rule numeral_neq_neg_numeral [symmetric])
|
|
745 |
|
|
746 |
lemma zero_neq_neg_numeral: "0 \<noteq> neg_numeral n"
|
|
747 |
unfolding neg_numeral_def neg_0_equal_iff_equal by simp
|
|
748 |
|
|
749 |
lemma neg_numeral_neq_zero: "neg_numeral n \<noteq> 0"
|
|
750 |
unfolding neg_numeral_def neg_equal_0_iff_equal by simp
|
|
751 |
|
|
752 |
lemma one_neq_neg_numeral: "1 \<noteq> neg_numeral n"
|
|
753 |
using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One)
|
|
754 |
|
|
755 |
lemma neg_numeral_neq_one: "neg_numeral n \<noteq> 1"
|
|
756 |
using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One)
|
|
757 |
|
|
758 |
lemmas eq_neg_numeral_simps [simp] =
|
|
759 |
neg_numeral_eq_iff
|
|
760 |
numeral_neq_neg_numeral neg_numeral_neq_numeral
|
|
761 |
one_neq_neg_numeral neg_numeral_neq_one
|
|
762 |
zero_neq_neg_numeral neg_numeral_neq_zero
|
|
763 |
|
|
764 |
end
|
|
765 |
|
|
766 |
subsubsection {*
|
|
767 |
Structures with negation and order: class @{text linordered_idom}
|
|
768 |
*}
|
|
769 |
|
|
770 |
context linordered_idom
|
|
771 |
begin
|
|
772 |
|
|
773 |
subclass ring_char_0 ..
|
|
774 |
|
|
775 |
lemma neg_numeral_le_iff: "neg_numeral m \<le> neg_numeral n \<longleftrightarrow> n \<le> m"
|
|
776 |
by (simp only: neg_numeral_def neg_le_iff_le numeral_le_iff)
|
|
777 |
|
|
778 |
lemma neg_numeral_less_iff: "neg_numeral m < neg_numeral n \<longleftrightarrow> n < m"
|
|
779 |
by (simp only: neg_numeral_def neg_less_iff_less numeral_less_iff)
|
|
780 |
|
|
781 |
lemma neg_numeral_less_zero: "neg_numeral n < 0"
|
|
782 |
by (simp only: neg_numeral_def neg_less_0_iff_less zero_less_numeral)
|
|
783 |
|
|
784 |
lemma neg_numeral_le_zero: "neg_numeral n \<le> 0"
|
|
785 |
by (simp only: neg_numeral_def neg_le_0_iff_le zero_le_numeral)
|
|
786 |
|
|
787 |
lemma not_zero_less_neg_numeral: "\<not> 0 < neg_numeral n"
|
|
788 |
by (simp only: not_less neg_numeral_le_zero)
|
|
789 |
|
|
790 |
lemma not_zero_le_neg_numeral: "\<not> 0 \<le> neg_numeral n"
|
|
791 |
by (simp only: not_le neg_numeral_less_zero)
|
|
792 |
|
|
793 |
lemma neg_numeral_less_numeral: "neg_numeral m < numeral n"
|
|
794 |
using neg_numeral_less_zero zero_less_numeral by (rule less_trans)
|
|
795 |
|
|
796 |
lemma neg_numeral_le_numeral: "neg_numeral m \<le> numeral n"
|
|
797 |
by (simp only: less_imp_le neg_numeral_less_numeral)
|
|
798 |
|
|
799 |
lemma not_numeral_less_neg_numeral: "\<not> numeral m < neg_numeral n"
|
|
800 |
by (simp only: not_less neg_numeral_le_numeral)
|
|
801 |
|
|
802 |
lemma not_numeral_le_neg_numeral: "\<not> numeral m \<le> neg_numeral n"
|
|
803 |
by (simp only: not_le neg_numeral_less_numeral)
|
|
804 |
|
|
805 |
lemma neg_numeral_less_one: "neg_numeral m < 1"
|
|
806 |
by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One])
|
|
807 |
|
|
808 |
lemma neg_numeral_le_one: "neg_numeral m \<le> 1"
|
|
809 |
by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One])
|
|
810 |
|
|
811 |
lemma not_one_less_neg_numeral: "\<not> 1 < neg_numeral m"
|
|
812 |
by (simp only: not_less neg_numeral_le_one)
|
|
813 |
|
|
814 |
lemma not_one_le_neg_numeral: "\<not> 1 \<le> neg_numeral m"
|
|
815 |
by (simp only: not_le neg_numeral_less_one)
|
|
816 |
|
|
817 |
lemma sub_non_negative:
|
|
818 |
"sub n m \<ge> 0 \<longleftrightarrow> n \<ge> m"
|
|
819 |
by (simp only: sub_def le_diff_eq) simp
|
|
820 |
|
|
821 |
lemma sub_positive:
|
|
822 |
"sub n m > 0 \<longleftrightarrow> n > m"
|
|
823 |
by (simp only: sub_def less_diff_eq) simp
|
|
824 |
|
|
825 |
lemma sub_non_positive:
|
|
826 |
"sub n m \<le> 0 \<longleftrightarrow> n \<le> m"
|
|
827 |
by (simp only: sub_def diff_le_eq) simp
|
|
828 |
|
|
829 |
lemma sub_negative:
|
|
830 |
"sub n m < 0 \<longleftrightarrow> n < m"
|
|
831 |
by (simp only: sub_def diff_less_eq) simp
|
|
832 |
|
|
833 |
lemmas le_neg_numeral_simps [simp] =
|
|
834 |
neg_numeral_le_iff
|
|
835 |
neg_numeral_le_numeral not_numeral_le_neg_numeral
|
|
836 |
neg_numeral_le_zero not_zero_le_neg_numeral
|
|
837 |
neg_numeral_le_one not_one_le_neg_numeral
|
|
838 |
|
|
839 |
lemmas less_neg_numeral_simps [simp] =
|
|
840 |
neg_numeral_less_iff
|
|
841 |
neg_numeral_less_numeral not_numeral_less_neg_numeral
|
|
842 |
neg_numeral_less_zero not_zero_less_neg_numeral
|
|
843 |
neg_numeral_less_one not_one_less_neg_numeral
|
|
844 |
|
|
845 |
lemma abs_numeral [simp]: "abs (numeral n) = numeral n"
|
|
846 |
by simp
|
|
847 |
|
|
848 |
lemma abs_neg_numeral [simp]: "abs (neg_numeral n) = numeral n"
|
|
849 |
by (simp only: neg_numeral_def abs_minus_cancel abs_numeral)
|
|
850 |
|
|
851 |
end
|
|
852 |
|
|
853 |
subsubsection {*
|
|
854 |
Natural numbers
|
|
855 |
*}
|
|
856 |
|
|
857 |
lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)"
|
|
858 |
unfolding numeral_plus_one [symmetric] by simp
|
|
859 |
|
|
860 |
lemma nat_number:
|
|
861 |
"1 = Suc 0"
|
|
862 |
"numeral One = Suc 0"
|
|
863 |
"numeral (Bit0 n) = Suc (numeral (BitM n))"
|
|
864 |
"numeral (Bit1 n) = Suc (numeral (Bit0 n))"
|
|
865 |
by (simp_all add: numeral.simps BitM_plus_one)
|
|
866 |
|
|
867 |
subsubsection {*
|
|
868 |
Structures with exponentiation
|
|
869 |
*}
|
|
870 |
|
|
871 |
context semiring_numeral
|
|
872 |
begin
|
|
873 |
|
|
874 |
lemma numeral_sqr: "numeral (sqr n) = numeral n * numeral n"
|
|
875 |
by (simp add: sqr_conv_mult numeral_mult)
|
|
876 |
|
|
877 |
lemma numeral_pow: "numeral (pow m n) = numeral m ^ numeral n"
|
|
878 |
by (induct n, simp_all add: numeral_class.numeral.simps
|
|
879 |
power_add numeral_sqr numeral_mult)
|
|
880 |
|
|
881 |
lemma power_numeral [simp]: "numeral m ^ numeral n = numeral (pow m n)"
|
|
882 |
by (rule numeral_pow [symmetric])
|
|
883 |
|
|
884 |
end
|
|
885 |
|
|
886 |
context semiring_1
|
|
887 |
begin
|
|
888 |
|
|
889 |
lemma power_zero_numeral [simp]: "(0::'a) ^ numeral n = 0"
|
|
890 |
by (induct n, simp_all add: numeral_class.numeral.simps power_add)
|
|
891 |
|
|
892 |
end
|
|
893 |
|
|
894 |
context ring_1
|
|
895 |
begin
|
|
896 |
|
|
897 |
lemma power_minus_Bit0: "(- x) ^ numeral (Bit0 n) = x ^ numeral (Bit0 n)"
|
|
898 |
by (induct n, simp_all add: numeral_class.numeral.simps power_add)
|
|
899 |
|
|
900 |
lemma power_minus_Bit1: "(- x) ^ numeral (Bit1 n) = - (x ^ numeral (Bit1 n))"
|
|
901 |
by (simp only: nat_number(4) power_Suc power_minus_Bit0 mult_minus_left)
|
|
902 |
|
|
903 |
lemma power_neg_numeral_Bit0 [simp]:
|
|
904 |
"neg_numeral m ^ numeral (Bit0 n) = numeral (pow m (Bit0 n))"
|
|
905 |
by (simp only: neg_numeral_def power_minus_Bit0 power_numeral)
|
|
906 |
|
|
907 |
lemma power_neg_numeral_Bit1 [simp]:
|
|
908 |
"neg_numeral m ^ numeral (Bit1 n) = neg_numeral (pow m (Bit1 n))"
|
|
909 |
by (simp only: neg_numeral_def power_minus_Bit1 power_numeral pow.simps)
|
|
910 |
|
|
911 |
end
|
|
912 |
|
|
913 |
subsection {* Numeral equations as default simplification rules *}
|
|
914 |
|
|
915 |
declare (in numeral) numeral_One [simp]
|
|
916 |
declare (in numeral) numeral_plus_numeral [simp]
|
|
917 |
declare (in numeral) add_numeral_special [simp]
|
|
918 |
declare (in neg_numeral) add_neg_numeral_simps [simp]
|
|
919 |
declare (in neg_numeral) add_neg_numeral_special [simp]
|
|
920 |
declare (in neg_numeral) diff_numeral_simps [simp]
|
|
921 |
declare (in neg_numeral) diff_numeral_special [simp]
|
|
922 |
declare (in semiring_numeral) numeral_times_numeral [simp]
|
|
923 |
declare (in ring_1) mult_neg_numeral_simps [simp]
|
|
924 |
|
|
925 |
subsection {* Setting up simprocs *}
|
|
926 |
|
|
927 |
lemma numeral_reorient:
|
|
928 |
"(numeral w = x) = (x = numeral w)"
|
|
929 |
by auto
|
|
930 |
|
|
931 |
lemma mult_numeral_1: "Numeral1 * a = (a::'a::semiring_numeral)"
|
|
932 |
by simp
|
|
933 |
|
|
934 |
lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::semiring_numeral)"
|
|
935 |
by simp
|
|
936 |
|
|
937 |
lemma divide_numeral_1: "a / Numeral1 = (a::'a::field)"
|
|
938 |
by simp
|
|
939 |
|
|
940 |
lemma inverse_numeral_1:
|
|
941 |
"inverse Numeral1 = (Numeral1::'a::division_ring)"
|
|
942 |
by simp
|
|
943 |
|
|
944 |
text{*Theorem lists for the cancellation simprocs. The use of a numary
|
|
945 |
numeral for 1 reduces the number of special cases.*}
|
|
946 |
|
|
947 |
lemmas mult_1s =
|
|
948 |
mult_numeral_1 mult_numeral_1_right
|
|
949 |
mult_minus1 mult_minus1_right
|
|
950 |
|
|
951 |
|
|
952 |
subsubsection {* Simplification of arithmetic operations on integer constants. *}
|
|
953 |
|
|
954 |
lemmas arith_special = (* already declared simp above *)
|
|
955 |
add_numeral_special add_neg_numeral_special
|
|
956 |
diff_numeral_special minus_one
|
|
957 |
|
|
958 |
(* rules already in simpset *)
|
|
959 |
lemmas arith_extra_simps =
|
|
960 |
numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right
|
|
961 |
minus_numeral minus_neg_numeral minus_zero minus_one
|
|
962 |
diff_numeral_simps diff_0 diff_0_right
|
|
963 |
numeral_times_numeral mult_neg_numeral_simps
|
|
964 |
mult_zero_left mult_zero_right
|
|
965 |
abs_numeral abs_neg_numeral
|
|
966 |
|
|
967 |
text {*
|
|
968 |
For making a minimal simpset, one must include these default simprules.
|
|
969 |
Also include @{text simp_thms}.
|
|
970 |
*}
|
|
971 |
|
|
972 |
lemmas arith_simps =
|
|
973 |
add_num_simps mult_num_simps sub_num_simps
|
|
974 |
BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps
|
|
975 |
abs_zero abs_one arith_extra_simps
|
|
976 |
|
|
977 |
text {* Simplification of relational operations *}
|
|
978 |
|
|
979 |
lemmas eq_numeral_extra =
|
|
980 |
zero_neq_one one_neq_zero
|
|
981 |
|
|
982 |
lemmas rel_simps =
|
|
983 |
le_num_simps less_num_simps eq_num_simps
|
|
984 |
le_numeral_simps le_neg_numeral_simps le_numeral_extra
|
|
985 |
less_numeral_simps less_neg_numeral_simps less_numeral_extra
|
|
986 |
eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra
|
|
987 |
|
|
988 |
|
|
989 |
subsubsection {* Simplification of arithmetic when nested to the right. *}
|
|
990 |
|
|
991 |
lemma add_numeral_left [simp]:
|
|
992 |
"numeral v + (numeral w + z) = (numeral(v + w) + z)"
|
|
993 |
by (simp_all add: add_assoc [symmetric])
|
|
994 |
|
|
995 |
lemma add_neg_numeral_left [simp]:
|
|
996 |
"numeral v + (neg_numeral w + y) = (sub v w + y)"
|
|
997 |
"neg_numeral v + (numeral w + y) = (sub w v + y)"
|
|
998 |
"neg_numeral v + (neg_numeral w + y) = (neg_numeral(v + w) + y)"
|
|
999 |
by (simp_all add: add_assoc [symmetric])
|
|
1000 |
|
|
1001 |
lemma mult_numeral_left [simp]:
|
|
1002 |
"numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)"
|
|
1003 |
"neg_numeral v * (numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)"
|
|
1004 |
"numeral v * (neg_numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)"
|
|
1005 |
"neg_numeral v * (neg_numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)"
|
|
1006 |
by (simp_all add: mult_assoc [symmetric])
|
|
1007 |
|
|
1008 |
hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec
|
|
1009 |
|
|
1010 |
subsection {* code module namespace *}
|
|
1011 |
|
|
1012 |
code_modulename SML
|
47126
|
1013 |
Num Arith
|
47108
|
1014 |
|
|
1015 |
code_modulename OCaml
|
47126
|
1016 |
Num Arith
|
47108
|
1017 |
|
|
1018 |
code_modulename Haskell
|
47126
|
1019 |
Num Arith
|
47108
|
1020 |
|
|
1021 |
end
|