separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
1 (* Title: HOL/Code_Numeral.thy
2 Author: Florian Haftmann, TU Muenchen
5 section {* Numeric types for code generation onto target language numerals only *}
8 imports Nat_Transfer Divides Lifting
11 subsection {* Type of target language integers *}
13 typedef integer = "UNIV \<Colon> int set"
14 morphisms int_of_integer integer_of_int ..
16 setup_lifting type_definition_integer
19 "k = l \<longleftrightarrow> int_of_integer k = int_of_integer l"
23 "int_of_integer k = int_of_integer l \<Longrightarrow> k = l"
24 using integer_eq_iff [of k l] by simp
26 lemma int_of_integer_integer_of_int [simp]:
27 "int_of_integer (integer_of_int k) = k"
30 lemma integer_of_int_int_of_integer [simp]:
31 "integer_of_int (int_of_integer k) = k"
34 instantiation integer :: ring_1
37 lift_definition zero_integer :: integer
41 declare zero_integer.rep_eq [simp]
43 lift_definition one_integer :: integer
47 declare one_integer.rep_eq [simp]
49 lift_definition plus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
50 is "plus :: int \<Rightarrow> int \<Rightarrow> int"
53 declare plus_integer.rep_eq [simp]
55 lift_definition uminus_integer :: "integer \<Rightarrow> integer"
56 is "uminus :: int \<Rightarrow> int"
59 declare uminus_integer.rep_eq [simp]
61 lift_definition minus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
62 is "minus :: int \<Rightarrow> int \<Rightarrow> int"
65 declare minus_integer.rep_eq [simp]
67 lift_definition times_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
68 is "times :: int \<Rightarrow> int \<Rightarrow> int"
71 declare times_integer.rep_eq [simp]
74 qed (transfer, simp add: algebra_simps)+
78 lemma [transfer_rule]:
79 "rel_fun HOL.eq pcr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)"
80 by (unfold of_nat_def [abs_def]) transfer_prover
82 lemma [transfer_rule]:
83 "rel_fun HOL.eq pcr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)"
85 have "rel_fun HOL.eq pcr_integer (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> integer)"
86 by (unfold of_int_of_nat [abs_def]) transfer_prover
87 then show ?thesis by (simp add: id_def)
90 lemma [transfer_rule]:
91 "rel_fun HOL.eq pcr_integer (numeral :: num \<Rightarrow> int) (numeral :: num \<Rightarrow> integer)"
93 have "rel_fun HOL.eq pcr_integer (numeral :: num \<Rightarrow> int) (\<lambda>n. of_int (numeral n))"
95 then show ?thesis by simp
98 lemma [transfer_rule]:
99 "rel_fun HOL.eq (rel_fun HOL.eq pcr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)"
100 by (unfold Num.sub_def [abs_def]) transfer_prover
102 lemma int_of_integer_of_nat [simp]:
103 "int_of_integer (of_nat n) = of_nat n"
106 lift_definition integer_of_nat :: "nat \<Rightarrow> integer"
107 is "of_nat :: nat \<Rightarrow> int"
110 lemma integer_of_nat_eq_of_nat [code]:
111 "integer_of_nat = of_nat"
114 lemma int_of_integer_integer_of_nat [simp]:
115 "int_of_integer (integer_of_nat n) = of_nat n"
118 lift_definition nat_of_integer :: "integer \<Rightarrow> nat"
122 lemma nat_of_integer_of_nat [simp]:
123 "nat_of_integer (of_nat n) = n"
126 lemma int_of_integer_of_int [simp]:
127 "int_of_integer (of_int k) = k"
130 lemma nat_of_integer_integer_of_nat [simp]:
131 "nat_of_integer (integer_of_nat n) = n"
134 lemma integer_of_int_eq_of_int [simp, code_abbrev]:
135 "integer_of_int = of_int"
136 by transfer (simp add: fun_eq_iff)
138 lemma of_int_integer_of [simp]:
139 "of_int (int_of_integer k) = (k :: integer)"
142 lemma int_of_integer_numeral [simp]:
143 "int_of_integer (numeral k) = numeral k"
146 lemma int_of_integer_sub [simp]:
147 "int_of_integer (Num.sub k l) = Num.sub k l"
150 instantiation integer :: "{ring_div, equal, linordered_idom}"
153 lift_definition divide_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
154 is "divide :: int \<Rightarrow> int \<Rightarrow> int"
157 declare divide_integer.rep_eq [simp]
159 lift_definition mod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
160 is "Divides.mod :: int \<Rightarrow> int \<Rightarrow> int"
163 declare mod_integer.rep_eq [simp]
165 lift_definition abs_integer :: "integer \<Rightarrow> integer"
166 is "abs :: int \<Rightarrow> int"
169 declare abs_integer.rep_eq [simp]
171 lift_definition sgn_integer :: "integer \<Rightarrow> integer"
172 is "sgn :: int \<Rightarrow> int"
175 declare sgn_integer.rep_eq [simp]
177 lift_definition less_eq_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
178 is "less_eq :: int \<Rightarrow> int \<Rightarrow> bool"
181 lift_definition less_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
182 is "less :: int \<Rightarrow> int \<Rightarrow> bool"
185 lift_definition equal_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
186 is "HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool"
190 qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+
194 lemma [transfer_rule]:
195 "rel_fun pcr_integer (rel_fun pcr_integer pcr_integer) (min :: _ \<Rightarrow> _ \<Rightarrow> int) (min :: _ \<Rightarrow> _ \<Rightarrow> integer)"
196 by (unfold min_def [abs_def]) transfer_prover
198 lemma [transfer_rule]:
199 "rel_fun pcr_integer (rel_fun pcr_integer pcr_integer) (max :: _ \<Rightarrow> _ \<Rightarrow> int) (max :: _ \<Rightarrow> _ \<Rightarrow> integer)"
200 by (unfold max_def [abs_def]) transfer_prover
202 lemma int_of_integer_min [simp]:
203 "int_of_integer (min k l) = min (int_of_integer k) (int_of_integer l)"
206 lemma int_of_integer_max [simp]:
207 "int_of_integer (max k l) = max (int_of_integer k) (int_of_integer l)"
210 lemma nat_of_integer_non_positive [simp]:
211 "k \<le> 0 \<Longrightarrow> nat_of_integer k = 0"
214 lemma of_nat_of_integer [simp]:
215 "of_nat (nat_of_integer k) = max 0 k"
218 instance integer :: semiring_numeral_div
219 by intro_classes (transfer,
220 fact semiring_numeral_div_class.le_add_diff_inverse2
221 semiring_numeral_div_class.div_less
222 semiring_numeral_div_class.mod_less
223 semiring_numeral_div_class.div_positive
224 semiring_numeral_div_class.mod_less_eq_dividend
225 semiring_numeral_div_class.pos_mod_bound
226 semiring_numeral_div_class.pos_mod_sign
227 semiring_numeral_div_class.mod_mult2_eq
228 semiring_numeral_div_class.div_mult2_eq
229 semiring_numeral_div_class.discrete)+
231 lemma integer_of_nat_0: "integer_of_nat 0 = 0"
234 lemma integer_of_nat_1: "integer_of_nat 1 = 1"
237 lemma integer_of_nat_numeral:
238 "integer_of_nat (numeral n) = numeral n"
241 subsection {* Code theorems for target language integers *}
243 text {* Constructors *}
245 definition Pos :: "num \<Rightarrow> integer"
247 [simp, code_abbrev]: "Pos = numeral"
249 lemma [transfer_rule]:
250 "rel_fun HOL.eq pcr_integer numeral Pos"
251 by simp transfer_prover
253 definition Neg :: "num \<Rightarrow> integer"
255 [simp, code_abbrev]: "Neg n = - Pos n"
257 lemma [transfer_rule]:
258 "rel_fun HOL.eq pcr_integer (\<lambda>n. - numeral n) Neg"
259 by (simp add: Neg_def [abs_def]) transfer_prover
261 code_datatype "0::integer" Pos Neg
264 text {* Auxiliary operations *}
266 lift_definition dup :: "integer \<Rightarrow> integer"
267 is "\<lambda>k::int. k + k"
270 lemma dup_code [code]:
272 "dup (Pos n) = Pos (Num.Bit0 n)"
273 "dup (Neg n) = Neg (Num.Bit0 n)"
274 by (transfer, simp only: numeral_Bit0 minus_add_distrib)+
276 lift_definition sub :: "num \<Rightarrow> num \<Rightarrow> integer"
277 is "\<lambda>m n. numeral m - numeral n :: int"
280 lemma sub_code [code]:
281 "sub Num.One Num.One = 0"
282 "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)"
283 "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)"
284 "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)"
285 "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)"
286 "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)"
287 "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
288 "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
289 "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
290 by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+
293 text {* Implementations *}
295 lemma one_integer_code [code, code_unfold]:
299 lemma plus_integer_code [code]:
300 "k + 0 = (k::integer)"
301 "0 + l = (l::integer)"
302 "Pos m + Pos n = Pos (m + n)"
303 "Pos m + Neg n = sub m n"
304 "Neg m + Pos n = sub n m"
305 "Neg m + Neg n = Neg (m + n)"
308 lemma uminus_integer_code [code]:
309 "uminus 0 = (0::integer)"
310 "uminus (Pos m) = Neg m"
311 "uminus (Neg m) = Pos m"
314 lemma minus_integer_code [code]:
315 "k - 0 = (k::integer)"
316 "0 - l = uminus (l::integer)"
317 "Pos m - Pos n = sub m n"
318 "Pos m - Neg n = Pos (m + n)"
319 "Neg m - Pos n = Neg (m + n)"
320 "Neg m - Neg n = sub n m"
323 lemma abs_integer_code [code]:
324 "\<bar>k\<bar> = (if (k::integer) < 0 then - k else k)"
327 lemma sgn_integer_code [code]:
328 "sgn k = (if k = 0 then 0 else if (k::integer) < 0 then - 1 else 1)"
331 lemma times_integer_code [code]:
332 "k * 0 = (0::integer)"
333 "0 * l = (0::integer)"
334 "Pos m * Pos n = Pos (m * n)"
335 "Pos m * Neg n = Neg (m * n)"
336 "Neg m * Pos n = Neg (m * n)"
337 "Neg m * Neg n = Pos (m * n)"
340 definition divmod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
342 "divmod_integer k l = (k div l, k mod l)"
344 lemma fst_divmod [simp]:
345 "fst (divmod_integer k l) = k div l"
346 by (simp add: divmod_integer_def)
348 lemma snd_divmod [simp]:
349 "snd (divmod_integer k l) = k mod l"
350 by (simp add: divmod_integer_def)
352 definition divmod_abs :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
354 "divmod_abs k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"
356 lemma fst_divmod_abs [simp]:
357 "fst (divmod_abs k l) = \<bar>k\<bar> div \<bar>l\<bar>"
358 by (simp add: divmod_abs_def)
360 lemma snd_divmod_abs [simp]:
361 "snd (divmod_abs k l) = \<bar>k\<bar> mod \<bar>l\<bar>"
362 by (simp add: divmod_abs_def)
364 lemma divmod_abs_code [code]:
365 "divmod_abs (Pos k) (Pos l) = divmod k l"
366 "divmod_abs (Neg k) (Neg l) = divmod k l"
367 "divmod_abs (Neg k) (Pos l) = divmod k l"
368 "divmod_abs (Pos k) (Neg l) = divmod k l"
369 "divmod_abs j 0 = (0, \<bar>j\<bar>)"
370 "divmod_abs 0 j = (0, 0)"
371 by (simp_all add: prod_eq_iff)
373 lemma divmod_integer_code [code]:
374 "divmod_integer k l =
375 (if k = 0 then (0, 0) else if l = 0 then (0, k) else
376 (apsnd \<circ> times \<circ> sgn) l (if sgn k = sgn l
378 else (let (r, s) = divmod_abs k l in
379 if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
381 have aux1: "\<And>k l::int. sgn k = sgn l \<longleftrightarrow> k = 0 \<and> l = 0 \<or> 0 < l \<and> 0 < k \<or> l < 0 \<and> k < 0"
382 by (auto simp add: sgn_if)
383 have aux2: "\<And>q::int. - int_of_integer k = int_of_integer l * q \<longleftrightarrow> int_of_integer k = int_of_integer l * - q" by auto
385 by (simp add: prod_eq_iff integer_eq_iff case_prod_beta aux1)
386 (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2)
389 lemma div_integer_code [code]:
390 "k div l = fst (divmod_integer k l)"
393 lemma mod_integer_code [code]:
394 "k mod l = snd (divmod_integer k l)"
397 lemma equal_integer_code [code]:
398 "HOL.equal 0 (0::integer) \<longleftrightarrow> True"
399 "HOL.equal 0 (Pos l) \<longleftrightarrow> False"
400 "HOL.equal 0 (Neg l) \<longleftrightarrow> False"
401 "HOL.equal (Pos k) 0 \<longleftrightarrow> False"
402 "HOL.equal (Pos k) (Pos l) \<longleftrightarrow> HOL.equal k l"
403 "HOL.equal (Pos k) (Neg l) \<longleftrightarrow> False"
404 "HOL.equal (Neg k) 0 \<longleftrightarrow> False"
405 "HOL.equal (Neg k) (Pos l) \<longleftrightarrow> False"
406 "HOL.equal (Neg k) (Neg l) \<longleftrightarrow> HOL.equal k l"
407 by (simp_all add: equal)
409 lemma equal_integer_refl [code nbe]:
410 "HOL.equal (k::integer) k \<longleftrightarrow> True"
413 lemma less_eq_integer_code [code]:
414 "0 \<le> (0::integer) \<longleftrightarrow> True"
415 "0 \<le> Pos l \<longleftrightarrow> True"
416 "0 \<le> Neg l \<longleftrightarrow> False"
417 "Pos k \<le> 0 \<longleftrightarrow> False"
418 "Pos k \<le> Pos l \<longleftrightarrow> k \<le> l"
419 "Pos k \<le> Neg l \<longleftrightarrow> False"
420 "Neg k \<le> 0 \<longleftrightarrow> True"
421 "Neg k \<le> Pos l \<longleftrightarrow> True"
422 "Neg k \<le> Neg l \<longleftrightarrow> l \<le> k"
425 lemma less_integer_code [code]:
426 "0 < (0::integer) \<longleftrightarrow> False"
427 "0 < Pos l \<longleftrightarrow> True"
428 "0 < Neg l \<longleftrightarrow> False"
429 "Pos k < 0 \<longleftrightarrow> False"
430 "Pos k < Pos l \<longleftrightarrow> k < l"
431 "Pos k < Neg l \<longleftrightarrow> False"
432 "Neg k < 0 \<longleftrightarrow> True"
433 "Neg k < Pos l \<longleftrightarrow> True"
434 "Neg k < Neg l \<longleftrightarrow> l < k"
437 lift_definition integer_of_num :: "num \<Rightarrow> integer"
438 is "numeral :: num \<Rightarrow> int"
441 lemma integer_of_num [code]:
442 "integer_of_num num.One = 1"
443 "integer_of_num (num.Bit0 n) = (let k = integer_of_num n in k + k)"
444 "integer_of_num (num.Bit1 n) = (let k = integer_of_num n in k + k + 1)"
445 by (transfer, simp only: numeral.simps Let_def)+
447 lift_definition num_of_integer :: "integer \<Rightarrow> num"
448 is "num_of_nat \<circ> nat"
451 lemma num_of_integer_code [code]:
452 "num_of_integer k = (if k \<le> 1 then Num.One
454 (l, j) = divmod_integer k 2;
455 l' = num_of_integer l;
457 in if j = 0 then l'' else l'' + Num.One)"
460 assume "int_of_integer k mod 2 = 1"
461 then have "nat (int_of_integer k mod 2) = nat 1" by simp
462 moreover assume *: "1 < int_of_integer k"
463 ultimately have **: "nat (int_of_integer k) mod 2 = 1" by (simp add: nat_mod_distrib)
464 have "num_of_nat (nat (int_of_integer k)) =
465 num_of_nat (2 * (nat (int_of_integer k) div 2) + nat (int_of_integer k) mod 2)"
467 then have "num_of_nat (nat (int_of_integer k)) =
468 num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + nat (int_of_integer k) mod 2)"
469 by (simp add: mult_2)
470 with ** have "num_of_nat (nat (int_of_integer k)) =
471 num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + 1)"
476 by (auto simp add: num_of_integer_def nat_of_integer_def Let_def case_prod_beta
477 not_le integer_eq_iff less_eq_integer_def
478 nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib
479 mult_2 [where 'a=nat] aux add_One)
482 lemma nat_of_integer_code [code]:
483 "nat_of_integer k = (if k \<le> 0 then 0
485 (l, j) = divmod_integer k 2;
486 l' = nat_of_integer l;
488 in if j = 0 then l'' else l'' + 1)"
490 obtain j where "k = integer_of_int j"
492 show "k = integer_of_int (int_of_integer k)" by simp
494 moreover have "2 * (j div 2) = j - j mod 2"
495 by (simp add: zmult_div_cancel mult.commute)
496 ultimately show ?thesis
497 by (auto simp add: split_def Let_def mod_integer_def nat_of_integer_def not_le
498 nat_add_distrib [symmetric] Suc_nat_eq_nat_zadd1)
499 (auto simp add: mult_2 [symmetric])
502 lemma int_of_integer_code [code]:
503 "int_of_integer k = (if k < 0 then - (int_of_integer (- k))
506 (l, j) = divmod_integer k 2;
507 l' = 2 * int_of_integer l
508 in if j = 0 then l' else l' + 1)"
509 by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
511 lemma integer_of_int_code [code]:
512 "integer_of_int k = (if k < 0 then - (integer_of_int (- k))
515 (l, j) = divmod_int k 2;
516 l' = 2 * integer_of_int l
517 in if j = 0 then l' else l' + 1)"
518 by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
520 hide_const (open) Pos Neg sub dup divmod_abs
523 subsection {* Serializer setup for target language integers *}
525 code_reserved Eval int Integer abs
528 type_constructor integer \<rightharpoonup>
530 and (OCaml) "Big'_int.big'_int"
531 and (Haskell) "Integer"
534 | class_instance integer :: equal \<rightharpoonup>
538 constant "0::integer" \<rightharpoonup>
539 (SML) "!(0/ :/ IntInf.int)"
540 and (OCaml) "Big'_int.zero'_big'_int"
541 and (Haskell) "!(0/ ::/ Integer)"
542 and (Scala) "BigInt(0)"
546 Numeral.add_code @{const_name Code_Numeral.Pos} I Code_Printer.literal_numeral target
547 #> Numeral.add_code @{const_name Code_Numeral.Neg} (op ~) Code_Printer.literal_numeral target)
548 ["SML", "OCaml", "Haskell", "Scala"]
552 constant "plus :: integer \<Rightarrow> _ \<Rightarrow> _" \<rightharpoonup>
553 (SML) "IntInf.+ ((_), (_))"
554 and (OCaml) "Big'_int.add'_big'_int"
555 and (Haskell) infixl 6 "+"
556 and (Scala) infixl 7 "+"
557 and (Eval) infixl 8 "+"
558 | constant "uminus :: integer \<Rightarrow> _" \<rightharpoonup>
560 and (OCaml) "Big'_int.minus'_big'_int"
561 and (Haskell) "negate"
564 | constant "minus :: integer \<Rightarrow> _" \<rightharpoonup>
565 (SML) "IntInf.- ((_), (_))"
566 and (OCaml) "Big'_int.sub'_big'_int"
567 and (Haskell) infixl 6 "-"
568 and (Scala) infixl 7 "-"
569 and (Eval) infixl 8 "-"
570 | constant Code_Numeral.dup \<rightharpoonup>
571 (SML) "IntInf.*/ (2,/ (_))"
572 and (OCaml) "Big'_int.mult'_big'_int/ (Big'_int.big'_int'_of'_int/ 2)"
573 and (Haskell) "!(2 * _)"
574 and (Scala) "!(2 * _)"
575 and (Eval) "!(2 * _)"
576 | constant Code_Numeral.sub \<rightharpoonup>
577 (SML) "!(raise/ Fail/ \"sub\")"
578 and (OCaml) "failwith/ \"sub\""
579 and (Haskell) "error/ \"sub\""
580 and (Scala) "!sys.error(\"sub\")"
581 | constant "times :: integer \<Rightarrow> _ \<Rightarrow> _" \<rightharpoonup>
582 (SML) "IntInf.* ((_), (_))"
583 and (OCaml) "Big'_int.mult'_big'_int"
584 and (Haskell) infixl 7 "*"
585 and (Scala) infixl 8 "*"
586 and (Eval) infixl 9 "*"
587 | constant Code_Numeral.divmod_abs \<rightharpoonup>
588 (SML) "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)"
589 and (OCaml) "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)"
590 and (Haskell) "divMod/ (abs _)/ (abs _)"
591 and (Scala) "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))"
592 and (Eval) "Integer.div'_mod/ (abs _)/ (abs _)"
593 | constant "HOL.equal :: integer \<Rightarrow> _ \<Rightarrow> bool" \<rightharpoonup>
594 (SML) "!((_ : IntInf.int) = _)"
595 and (OCaml) "Big'_int.eq'_big'_int"
596 and (Haskell) infix 4 "=="
597 and (Scala) infixl 5 "=="
598 and (Eval) infixl 6 "="
599 | constant "less_eq :: integer \<Rightarrow> _ \<Rightarrow> bool" \<rightharpoonup>
600 (SML) "IntInf.<= ((_), (_))"
601 and (OCaml) "Big'_int.le'_big'_int"
602 and (Haskell) infix 4 "<="
603 and (Scala) infixl 4 "<="
604 and (Eval) infixl 6 "<="
605 | constant "less :: integer \<Rightarrow> _ \<Rightarrow> bool" \<rightharpoonup>
606 (SML) "IntInf.< ((_), (_))"
607 and (OCaml) "Big'_int.lt'_big'_int"
608 and (Haskell) infix 4 "<"
609 and (Scala) infixl 4 "<"
610 and (Eval) infixl 6 "<"
613 code_module Code_Numeral \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
616 subsection {* Type of target language naturals *}
618 typedef natural = "UNIV \<Colon> nat set"
619 morphisms nat_of_natural natural_of_nat ..
621 setup_lifting type_definition_natural
623 lemma natural_eq_iff [termination_simp]:
624 "m = n \<longleftrightarrow> nat_of_natural m = nat_of_natural n"
628 "nat_of_natural m = nat_of_natural n \<Longrightarrow> m = n"
629 using natural_eq_iff [of m n] by simp
631 lemma nat_of_natural_of_nat_inverse [simp]:
632 "nat_of_natural (natural_of_nat n) = n"
635 lemma natural_of_nat_of_natural_inverse [simp]:
636 "natural_of_nat (nat_of_natural n) = n"
639 instantiation natural :: "{comm_monoid_diff, semiring_1}"
642 lift_definition zero_natural :: natural
646 declare zero_natural.rep_eq [simp]
648 lift_definition one_natural :: natural
652 declare one_natural.rep_eq [simp]
654 lift_definition plus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
655 is "plus :: nat \<Rightarrow> nat \<Rightarrow> nat"
658 declare plus_natural.rep_eq [simp]
660 lift_definition minus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
661 is "minus :: nat \<Rightarrow> nat \<Rightarrow> nat"
664 declare minus_natural.rep_eq [simp]
666 lift_definition times_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
667 is "times :: nat \<Rightarrow> nat \<Rightarrow> nat"
670 declare times_natural.rep_eq [simp]
673 qed (transfer, simp add: algebra_simps)+
677 lemma [transfer_rule]:
678 "rel_fun HOL.eq pcr_natural (\<lambda>n::nat. n) (of_nat :: nat \<Rightarrow> natural)"
680 have "rel_fun HOL.eq pcr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)"
681 by (unfold of_nat_def [abs_def]) transfer_prover
682 then show ?thesis by (simp add: id_def)
685 lemma [transfer_rule]:
686 "rel_fun HOL.eq pcr_natural (numeral :: num \<Rightarrow> nat) (numeral :: num \<Rightarrow> natural)"
688 have "rel_fun HOL.eq pcr_natural (numeral :: num \<Rightarrow> nat) (\<lambda>n. of_nat (numeral n))"
690 then show ?thesis by simp
693 lemma nat_of_natural_of_nat [simp]:
694 "nat_of_natural (of_nat n) = n"
697 lemma natural_of_nat_of_nat [simp, code_abbrev]:
698 "natural_of_nat = of_nat"
701 lemma of_nat_of_natural [simp]:
702 "of_nat (nat_of_natural n) = n"
705 lemma nat_of_natural_numeral [simp]:
706 "nat_of_natural (numeral k) = numeral k"
709 instantiation natural :: "{semiring_div, equal, linordered_semiring}"
712 lift_definition divide_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
713 is "divide :: nat \<Rightarrow> nat \<Rightarrow> nat"
716 declare divide_natural.rep_eq [simp]
718 lift_definition mod_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
719 is "Divides.mod :: nat \<Rightarrow> nat \<Rightarrow> nat"
722 declare mod_natural.rep_eq [simp]
724 lift_definition less_eq_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
725 is "less_eq :: nat \<Rightarrow> nat \<Rightarrow> bool"
728 declare less_eq_natural.rep_eq [termination_simp]
730 lift_definition less_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
731 is "less :: nat \<Rightarrow> nat \<Rightarrow> bool"
734 declare less_natural.rep_eq [termination_simp]
736 lift_definition equal_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
737 is "HOL.equal :: nat \<Rightarrow> nat \<Rightarrow> bool"
741 qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] linear)+
745 lemma [transfer_rule]:
746 "rel_fun pcr_natural (rel_fun pcr_natural pcr_natural) (min :: _ \<Rightarrow> _ \<Rightarrow> nat) (min :: _ \<Rightarrow> _ \<Rightarrow> natural)"
747 by (unfold min_def [abs_def]) transfer_prover
749 lemma [transfer_rule]:
750 "rel_fun pcr_natural (rel_fun pcr_natural pcr_natural) (max :: _ \<Rightarrow> _ \<Rightarrow> nat) (max :: _ \<Rightarrow> _ \<Rightarrow> natural)"
751 by (unfold max_def [abs_def]) transfer_prover
753 lemma nat_of_natural_min [simp]:
754 "nat_of_natural (min k l) = min (nat_of_natural k) (nat_of_natural l)"
757 lemma nat_of_natural_max [simp]:
758 "nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)"
761 lift_definition natural_of_integer :: "integer \<Rightarrow> natural"
762 is "nat :: int \<Rightarrow> nat"
765 lift_definition integer_of_natural :: "natural \<Rightarrow> integer"
766 is "of_nat :: nat \<Rightarrow> int"
769 lemma natural_of_integer_of_natural [simp]:
770 "natural_of_integer (integer_of_natural n) = n"
773 lemma integer_of_natural_of_integer [simp]:
774 "integer_of_natural (natural_of_integer k) = max 0 k"
777 lemma int_of_integer_of_natural [simp]:
778 "int_of_integer (integer_of_natural n) = of_nat (nat_of_natural n)"
781 lemma integer_of_natural_of_nat [simp]:
782 "integer_of_natural (of_nat n) = of_nat n"
785 lemma [measure_function]:
786 "is_measure nat_of_natural"
787 by (rule is_measure_trivial)
790 subsection {* Inductive representation of target language naturals *}
792 lift_definition Suc :: "natural \<Rightarrow> natural"
796 declare Suc.rep_eq [simp]
798 old_rep_datatype "0::natural" Suc
799 by (transfer, fact nat.induct nat.inject nat.distinct)+
801 lemma natural_cases [case_names nat, cases type: natural]:
803 assumes "\<And>n. m = of_nat n \<Longrightarrow> P"
805 using assms by transfer blast
807 lemma [simp, code]: "size_natural = nat_of_natural"
810 show "size_natural n = nat_of_natural n"
811 by (induct n) simp_all
814 lemma [simp, code]: "size = nat_of_natural"
817 show "size n = nat_of_natural n"
818 by (induct n) simp_all
821 lemma natural_decr [termination_simp]:
822 "n \<noteq> 0 \<Longrightarrow> nat_of_natural n - Nat.Suc 0 < nat_of_natural n"
825 lemma natural_zero_minus_one: "(0::natural) - 1 = 0"
828 lemma Suc_natural_minus_one: "Suc n - 1 = n"
831 hide_const (open) Suc
834 subsection {* Code refinement for target language naturals *}
836 lift_definition Nat :: "integer \<Rightarrow> natural"
843 "Nat (numeral k) = numeral k"
846 lemma [code abstype]:
847 "Nat (integer_of_natural n) = n"
850 lemma [code abstract]:
851 "integer_of_natural (natural_of_nat n) = of_nat n"
854 lemma [code abstract]:
855 "integer_of_natural (natural_of_integer k) = max 0 k"
859 "natural_of_integer (Code_Numeral.Pos k) = numeral k"
862 lemma [code abstract]:
863 "integer_of_natural 0 = 0"
866 lemma [code abstract]:
867 "integer_of_natural 1 = 1"
870 lemma [code abstract]:
871 "integer_of_natural (Code_Numeral.Suc n) = integer_of_natural n + 1"
875 "nat_of_natural = nat_of_integer \<circ> integer_of_natural"
876 by transfer (simp add: fun_eq_iff)
878 lemma [code, code_unfold]:
879 "case_natural f g n = (if n = 0 then f else g (n - 1))"
880 by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def)
882 declare natural.rec [code del]
884 lemma [code abstract]:
885 "integer_of_natural (m + n) = integer_of_natural m + integer_of_natural n"
888 lemma [code abstract]:
889 "integer_of_natural (m - n) = max 0 (integer_of_natural m - integer_of_natural n)"
892 lemma [code abstract]:
893 "integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n"
894 by transfer (simp add: of_nat_mult)
896 lemma [code abstract]:
897 "integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n"
898 by transfer (simp add: zdiv_int)
900 lemma [code abstract]:
901 "integer_of_natural (m mod n) = integer_of_natural m mod integer_of_natural n"
902 by transfer (simp add: zmod_int)
905 "HOL.equal m n \<longleftrightarrow> HOL.equal (integer_of_natural m) (integer_of_natural n)"
906 by transfer (simp add: equal)
908 lemma [code nbe]: "HOL.equal n (n::natural) \<longleftrightarrow> True"
909 by (rule equal_class.equal_refl)
911 lemma [code]: "m \<le> n \<longleftrightarrow> integer_of_natural m \<le> integer_of_natural n"
914 lemma [code]: "m < n \<longleftrightarrow> integer_of_natural m < integer_of_natural n"
917 hide_const (open) Nat
919 lifting_update integer.lifting
920 lifting_forget integer.lifting
922 lifting_update natural.lifting
923 lifting_forget natural.lifting
925 code_reflect Code_Numeral
926 datatypes natural = _
927 functions integer_of_natural natural_of_integer