src/HOL/ex/Numeral.thy
author haftmann
Fri, 10 Sep 2010 10:21:25 +0200
changeset 39499 0b61951d2682
parent 39393 917b4b6ba3d2
child 41494 e1fce873b814
permissions -rw-r--r--
Haskell == is infix, not infixl
     1 (*  Title:      HOL/ex/Numeral.thy
     2     Author:     Florian Haftmann
     3 *)
     4 
     5 header {* An experimental alternative numeral representation. *}
     6 
     7 theory Numeral
     8 imports Main
     9 begin
    10 
    11 subsection {* The @{text num} type *}
    12 
    13 datatype num = One | Dig0 num | Dig1 num
    14 
    15 text {* Increment function for type @{typ num} *}
    16 
    17 primrec inc :: "num \<Rightarrow> num" where
    18   "inc One = Dig0 One"
    19 | "inc (Dig0 x) = Dig1 x"
    20 | "inc (Dig1 x) = Dig0 (inc x)"
    21 
    22 text {* Converting between type @{typ num} and type @{typ nat} *}
    23 
    24 primrec nat_of_num :: "num \<Rightarrow> nat" where
    25   "nat_of_num One = Suc 0"
    26 | "nat_of_num (Dig0 x) = nat_of_num x + nat_of_num x"
    27 | "nat_of_num (Dig1 x) = Suc (nat_of_num x + nat_of_num x)"
    28 
    29 primrec num_of_nat :: "nat \<Rightarrow> num" where
    30   "num_of_nat 0 = One"
    31 | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)"
    32 
    33 lemma nat_of_num_pos: "0 < nat_of_num x"
    34   by (induct x) simp_all
    35 
    36 lemma nat_of_num_neq_0: "nat_of_num x \<noteq> 0"
    37   by (induct x) simp_all
    38 
    39 lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)"
    40   by (induct x) simp_all
    41 
    42 lemma num_of_nat_double:
    43   "0 < n \<Longrightarrow> num_of_nat (n + n) = Dig0 (num_of_nat n)"
    44   by (induct n) simp_all
    45 
    46 text {*
    47   Type @{typ num} is isomorphic to the strictly positive
    48   natural numbers.
    49 *}
    50 
    51 lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x"
    52   by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos)
    53 
    54 lemma num_of_nat_inverse: "0 < n \<Longrightarrow> nat_of_num (num_of_nat n) = n"
    55   by (induct n) (simp_all add: nat_of_num_inc)
    56 
    57 lemma num_eq_iff: "x = y \<longleftrightarrow> nat_of_num x = nat_of_num y"
    58 proof
    59   assume "nat_of_num x = nat_of_num y"
    60   then have "num_of_nat (nat_of_num x) = num_of_nat (nat_of_num y)" by simp
    61   then show "x = y" by (simp add: nat_of_num_inverse)
    62 qed simp
    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}: as
    86   positive naturals (rule @{text "num_induct"}) and as digit
    87   representation (rules @{text "num.induct"}, @{text "num.cases"}).
    88 
    89   It is not entirely clear in which context it is better to use the
    90   one or the other, or whether the construction should be reversed.
    91 *}
    92 
    93 
    94 subsection {* Numeral operations *}
    95 
    96 ML {*
    97 structure Dig_Simps = Named_Thms
    98 (
    99   val name = "numeral"
   100   val description = "simplification rules for numerals"
   101 )
   102 *}
   103 
   104 setup Dig_Simps.setup
   105 
   106 instantiation num :: "{plus,times,ord}"
   107 begin
   108 
   109 definition plus_num :: "num \<Rightarrow> num \<Rightarrow> num" where
   110   "m + n = num_of_nat (nat_of_num m + nat_of_num n)"
   111 
   112 definition times_num :: "num \<Rightarrow> num \<Rightarrow> num" where
   113   "m * n = num_of_nat (nat_of_num m * nat_of_num n)"
   114 
   115 definition less_eq_num :: "num \<Rightarrow> num \<Rightarrow> bool" where
   116   "m \<le> n \<longleftrightarrow> nat_of_num m \<le> nat_of_num n"
   117 
   118 definition less_num :: "num \<Rightarrow> num \<Rightarrow> bool" where
   119   "m < n \<longleftrightarrow> nat_of_num m < nat_of_num n"
   120 
   121 instance ..
   122 
   123 end
   124 
   125 lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y"
   126   unfolding plus_num_def
   127   by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos)
   128 
   129 lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y"
   130   unfolding times_num_def
   131   by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos)
   132 
   133 lemma Dig_plus [numeral, simp, code]:
   134   "One + One = Dig0 One"
   135   "One + Dig0 m = Dig1 m"
   136   "One + Dig1 m = Dig0 (m + One)"
   137   "Dig0 n + One = Dig1 n"
   138   "Dig0 n + Dig0 m = Dig0 (n + m)"
   139   "Dig0 n + Dig1 m = Dig1 (n + m)"
   140   "Dig1 n + One = Dig0 (n + One)"
   141   "Dig1 n + Dig0 m = Dig1 (n + m)"
   142   "Dig1 n + Dig1 m = Dig0 (n + m + One)"
   143   by (simp_all add: num_eq_iff nat_of_num_add)
   144 
   145 lemma Dig_times [numeral, simp, code]:
   146   "One * One = One"
   147   "One * Dig0 n = Dig0 n"
   148   "One * Dig1 n = Dig1 n"
   149   "Dig0 n * One = Dig0 n"
   150   "Dig0 n * Dig0 m = Dig0 (n * Dig0 m)"
   151   "Dig0 n * Dig1 m = Dig0 (n * Dig1 m)"
   152   "Dig1 n * One = Dig1 n"
   153   "Dig1 n * Dig0 m = Dig0 (n * Dig0 m + m)"
   154   "Dig1 n * Dig1 m = Dig1 (n * Dig1 m + m)"
   155   by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult
   156                     left_distrib right_distrib)
   157 
   158 lemma less_eq_num_code [numeral, simp, code]:
   159   "One \<le> n \<longleftrightarrow> True"
   160   "Dig0 m \<le> One \<longleftrightarrow> False"
   161   "Dig1 m \<le> One \<longleftrightarrow> False"
   162   "Dig0 m \<le> Dig0 n \<longleftrightarrow> m \<le> n"
   163   "Dig0 m \<le> Dig1 n \<longleftrightarrow> m \<le> n"
   164   "Dig1 m \<le> Dig1 n \<longleftrightarrow> m \<le> n"
   165   "Dig1 m \<le> Dig0 n \<longleftrightarrow> m < n"
   166   using nat_of_num_pos [of n] nat_of_num_pos [of m]
   167   by (auto simp add: less_eq_num_def less_num_def)
   168 
   169 lemma less_num_code [numeral, simp, code]:
   170   "m < One \<longleftrightarrow> False"
   171   "One < One \<longleftrightarrow> False"
   172   "One < Dig0 n \<longleftrightarrow> True"
   173   "One < Dig1 n \<longleftrightarrow> True"
   174   "Dig0 m < Dig0 n \<longleftrightarrow> m < n"
   175   "Dig0 m < Dig1 n \<longleftrightarrow> m \<le> n"
   176   "Dig1 m < Dig1 n \<longleftrightarrow> m < n"
   177   "Dig1 m < Dig0 n \<longleftrightarrow> m < n"
   178   using nat_of_num_pos [of n] nat_of_num_pos [of m]
   179   by (auto simp add: less_eq_num_def less_num_def)
   180 
   181 text {* Rules using @{text One} and @{text inc} as constructors *}
   182 
   183 lemma add_One: "x + One = inc x"
   184   by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
   185 
   186 lemma add_inc: "x + inc y = inc (x + y)"
   187   by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
   188 
   189 lemma mult_One: "x * One = x"
   190   by (simp add: num_eq_iff nat_of_num_mult)
   191 
   192 lemma mult_inc: "x * inc y = x * y + x"
   193   by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc)
   194 
   195 text {* A double-and-decrement function *}
   196 
   197 primrec DigM :: "num \<Rightarrow> num" where
   198   "DigM One = One"
   199 | "DigM (Dig0 n) = Dig1 (DigM n)"
   200 | "DigM (Dig1 n) = Dig1 (Dig0 n)"
   201 
   202 lemma DigM_plus_one: "DigM n + One = Dig0 n"
   203   by (induct n) simp_all
   204 
   205 lemma add_One_commute: "One + n = n + One"
   206   by (induct n) simp_all
   207 
   208 lemma one_plus_DigM: "One + DigM n = Dig0 n"
   209   by (simp add: add_One_commute DigM_plus_one)
   210 
   211 text {* Squaring and exponentiation *}
   212 
   213 primrec square :: "num \<Rightarrow> num" where
   214   "square One = One"
   215 | "square (Dig0 n) = Dig0 (Dig0 (square n))"
   216 | "square (Dig1 n) = Dig1 (Dig0 (square n + n))"
   217 
   218 primrec pow :: "num \<Rightarrow> num \<Rightarrow> num" where
   219   "pow x One = x"
   220 | "pow x (Dig0 y) = square (pow x y)"
   221 | "pow x (Dig1 y) = x * square (pow x y)"
   222 
   223 
   224 subsection {* Binary numerals *}
   225 
   226 text {*
   227   We embed binary representations into a generic algebraic
   228   structure using @{text of_num}.
   229 *}
   230 
   231 class semiring_numeral = semiring + monoid_mult
   232 begin
   233 
   234 primrec of_num :: "num \<Rightarrow> 'a" where
   235   of_num_One [numeral]: "of_num One = 1"
   236 | "of_num (Dig0 n) = of_num n + of_num n"
   237 | "of_num (Dig1 n) = of_num n + of_num n + 1"
   238 
   239 lemma of_num_inc: "of_num (inc n) = of_num n + 1"
   240   by (induct n) (simp_all add: add_ac)
   241 
   242 lemma of_num_add: "of_num (m + n) = of_num m + of_num n"
   243   by (induct n rule: num_induct) (simp_all add: add_One add_inc of_num_inc add_ac)
   244 
   245 lemma of_num_mult: "of_num (m * n) = of_num m * of_num n"
   246   by (induct n rule: num_induct) (simp_all add: mult_One mult_inc of_num_add of_num_inc right_distrib)
   247 
   248 declare of_num.simps [simp del]
   249 
   250 end
   251 
   252 ML {*
   253 fun mk_num k =
   254   if k > 1 then
   255     let
   256       val (l, b) = Integer.div_mod k 2;
   257       val bit = (if b = 0 then @{term Dig0} else @{term Dig1});
   258     in bit $ (mk_num l) end
   259   else if k = 1 then @{term One}
   260   else error ("mk_num " ^ string_of_int k);
   261 
   262 fun dest_num @{term One} = 1
   263   | dest_num (@{term Dig0} $ n) = 2 * dest_num n
   264   | dest_num (@{term Dig1} $ n) = 2 * dest_num n + 1
   265   | dest_num t = raise TERM ("dest_num", [t]);
   266 
   267 fun mk_numeral phi T k = Morphism.term phi (Const (@{const_name of_num}, @{typ num} --> T))
   268   $ mk_num k
   269 
   270 fun dest_numeral phi (u $ t) =
   271   if Term.aconv_untyped (u, Morphism.term phi (Const (@{const_name of_num}, dummyT)))
   272   then (range_type (fastype_of u), dest_num t)
   273   else raise TERM ("dest_numeral", [u, t]);
   274 *}
   275 
   276 syntax
   277   "_Numerals" :: "xnum \<Rightarrow> 'a"    ("_")
   278 
   279 parse_translation {*
   280 let
   281   fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2)
   282      of (0, 1) => Const (@{const_name One}, dummyT)
   283       | (n, 0) => Const (@{const_name Dig0}, dummyT) $ num_of_int n
   284       | (n, 1) => Const (@{const_name Dig1}, dummyT) $ num_of_int n
   285     else raise Match;
   286   fun numeral_tr [Free (num, _)] =
   287         let
   288           val {leading_zeros, value, ...} = Syntax.read_xnum num;
   289           val _ = leading_zeros = 0 andalso value > 0
   290             orelse error ("Bad numeral: " ^ num);
   291         in Const (@{const_name of_num}, @{typ num} --> dummyT) $ num_of_int value end
   292     | numeral_tr ts = raise TERM ("numeral_tr", ts);
   293 in [(@{syntax_const "_Numerals"}, numeral_tr)] end
   294 *}
   295 
   296 typed_print_translation (advanced) {*
   297 let
   298   fun dig b n = b + 2 * n; 
   299   fun int_of_num' (Const (@{const_syntax Dig0}, _) $ n) =
   300         dig 0 (int_of_num' n)
   301     | int_of_num' (Const (@{const_syntax Dig1}, _) $ n) =
   302         dig 1 (int_of_num' n)
   303     | int_of_num' (Const (@{const_syntax One}, _)) = 1;
   304   fun num_tr' ctxt show_sorts T [n] =
   305     let
   306       val k = int_of_num' n;
   307       val t' = Syntax.const @{syntax_const "_Numerals"} $ Syntax.free ("#" ^ string_of_int k);
   308     in
   309       case T of
   310         Type (@{type_name fun}, [_, T']) =>
   311           if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t'
   312           else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T'
   313       | T' => if T' = dummyT then t' else raise Match
   314     end;
   315 in [(@{const_syntax of_num}, num_tr')] end
   316 *}
   317 
   318 
   319 subsection {* Class-specific numeral rules *}
   320 
   321 subsubsection {* Class @{text semiring_numeral} *}
   322 
   323 context semiring_numeral
   324 begin
   325 
   326 abbreviation "Num1 \<equiv> of_num One"
   327 
   328 text {*
   329   Alas, there is still the duplication of @{term 1}, although the
   330   duplicated @{term 0} has disappeared.  We could get rid of it by
   331   replacing the constructor @{term 1} in @{typ num} by two
   332   constructors @{text two} and @{text three}, resulting in a further
   333   blow-up.  But it could be worth the effort.
   334 *}
   335 
   336 lemma of_num_plus_one [numeral]:
   337   "of_num n + 1 = of_num (n + One)"
   338   by (simp only: of_num_add of_num_One)
   339 
   340 lemma of_num_one_plus [numeral]:
   341   "1 + of_num n = of_num (One + n)"
   342   by (simp only: of_num_add of_num_One)
   343 
   344 lemma of_num_plus [numeral]:
   345   "of_num m + of_num n = of_num (m + n)"
   346   by (simp only: of_num_add)
   347 
   348 lemma of_num_times_one [numeral]:
   349   "of_num n * 1 = of_num n"
   350   by simp
   351 
   352 lemma of_num_one_times [numeral]:
   353   "1 * of_num n = of_num n"
   354   by simp
   355 
   356 lemma of_num_times [numeral]:
   357   "of_num m * of_num n = of_num (m * n)"
   358   unfolding of_num_mult ..
   359 
   360 end
   361 
   362 
   363 subsubsection {* Structures with a zero: class @{text semiring_1} *}
   364 
   365 context semiring_1
   366 begin
   367 
   368 subclass semiring_numeral ..
   369 
   370 lemma of_nat_of_num [numeral]: "of_nat (of_num n) = of_num n"
   371   by (induct n)
   372     (simp_all add: semiring_numeral_class.of_num.simps of_num.simps add_ac)
   373 
   374 declare of_nat_1 [numeral]
   375 
   376 lemma Dig_plus_zero [numeral]:
   377   "0 + 1 = 1"
   378   "0 + of_num n = of_num n"
   379   "1 + 0 = 1"
   380   "of_num n + 0 = of_num n"
   381   by simp_all
   382 
   383 lemma Dig_times_zero [numeral]:
   384   "0 * 1 = 0"
   385   "0 * of_num n = 0"
   386   "1 * 0 = 0"
   387   "of_num n * 0 = 0"
   388   by simp_all
   389 
   390 end
   391 
   392 lemma nat_of_num_of_num: "nat_of_num = of_num"
   393 proof
   394   fix n
   395   have "of_num n = nat_of_num n"
   396     by (induct n) (simp_all add: of_num.simps)
   397   then show "nat_of_num n = of_num n" by simp
   398 qed
   399 
   400 
   401 subsubsection {* Equality: class @{text semiring_char_0} *}
   402 
   403 context semiring_char_0
   404 begin
   405 
   406 lemma of_num_eq_iff [numeral]: "of_num m = of_num n \<longleftrightarrow> m = n"
   407   unfolding of_nat_of_num [symmetric] nat_of_num_of_num [symmetric]
   408     of_nat_eq_iff num_eq_iff ..
   409 
   410 lemma of_num_eq_one_iff [numeral]: "of_num n = 1 \<longleftrightarrow> n = One"
   411   using of_num_eq_iff [of n One] by (simp add: of_num_One)
   412 
   413 lemma one_eq_of_num_iff [numeral]: "1 = of_num n \<longleftrightarrow> One = n"
   414   using of_num_eq_iff [of One n] by (simp add: of_num_One)
   415 
   416 end
   417 
   418 
   419 subsubsection {* Comparisons: class @{text linordered_semidom} *}
   420 
   421 text {*
   422   Perhaps the underlying structure could even 
   423   be more general than @{text linordered_semidom}.
   424 *}
   425 
   426 context linordered_semidom
   427 begin
   428 
   429 lemma of_num_pos [numeral]: "0 < of_num n"
   430   by (induct n) (simp_all add: of_num.simps add_pos_pos)
   431 
   432 lemma of_num_not_zero [numeral]: "of_num n \<noteq> 0"
   433   using of_num_pos [of n] by simp
   434 
   435 lemma of_num_less_eq_iff [numeral]: "of_num m \<le> of_num n \<longleftrightarrow> m \<le> n"
   436 proof -
   437   have "of_nat (of_num m) \<le> of_nat (of_num n) \<longleftrightarrow> m \<le> n"
   438     unfolding less_eq_num_def nat_of_num_of_num of_nat_le_iff ..
   439   then show ?thesis by (simp add: of_nat_of_num)
   440 qed
   441 
   442 lemma of_num_less_eq_one_iff [numeral]: "of_num n \<le> 1 \<longleftrightarrow> n \<le> One"
   443   using of_num_less_eq_iff [of n One] by (simp add: of_num_One)
   444 
   445 lemma one_less_eq_of_num_iff [numeral]: "1 \<le> of_num n"
   446   using of_num_less_eq_iff [of One n] by (simp add: of_num_One)
   447 
   448 lemma of_num_less_iff [numeral]: "of_num m < of_num n \<longleftrightarrow> m < n"
   449 proof -
   450   have "of_nat (of_num m) < of_nat (of_num n) \<longleftrightarrow> m < n"
   451     unfolding less_num_def nat_of_num_of_num of_nat_less_iff ..
   452   then show ?thesis by (simp add: of_nat_of_num)
   453 qed
   454 
   455 lemma of_num_less_one_iff [numeral]: "\<not> of_num n < 1"
   456   using of_num_less_iff [of n One] by (simp add: of_num_One)
   457 
   458 lemma one_less_of_num_iff [numeral]: "1 < of_num n \<longleftrightarrow> One < n"
   459   using of_num_less_iff [of One n] by (simp add: of_num_One)
   460 
   461 lemma of_num_nonneg [numeral]: "0 \<le> of_num n"
   462   by (induct n) (simp_all add: of_num.simps add_nonneg_nonneg)
   463 
   464 lemma of_num_less_zero_iff [numeral]: "\<not> of_num n < 0"
   465   by (simp add: not_less of_num_nonneg)
   466 
   467 lemma of_num_le_zero_iff [numeral]: "\<not> of_num n \<le> 0"
   468   by (simp add: not_le of_num_pos)
   469 
   470 end
   471 
   472 context linordered_idom
   473 begin
   474 
   475 lemma minus_of_num_less_of_num_iff: "- of_num m < of_num n"
   476 proof -
   477   have "- of_num m < 0" by (simp add: of_num_pos)
   478   also have "0 < of_num n" by (simp add: of_num_pos)
   479   finally show ?thesis .
   480 qed
   481 
   482 lemma minus_of_num_not_equal_of_num: "- of_num m \<noteq> of_num n"
   483   using minus_of_num_less_of_num_iff [of m n] by simp
   484 
   485 lemma minus_of_num_less_one_iff: "- of_num n < 1"
   486   using minus_of_num_less_of_num_iff [of n One] by (simp add: of_num_One)
   487 
   488 lemma minus_one_less_of_num_iff: "- 1 < of_num n"
   489   using minus_of_num_less_of_num_iff [of One n] by (simp add: of_num_One)
   490 
   491 lemma minus_one_less_one_iff: "- 1 < 1"
   492   using minus_of_num_less_of_num_iff [of One One] by (simp add: of_num_One)
   493 
   494 lemma minus_of_num_le_of_num_iff: "- of_num m \<le> of_num n"
   495   by (simp add: less_imp_le minus_of_num_less_of_num_iff)
   496 
   497 lemma minus_of_num_le_one_iff: "- of_num n \<le> 1"
   498   by (simp add: less_imp_le minus_of_num_less_one_iff)
   499 
   500 lemma minus_one_le_of_num_iff: "- 1 \<le> of_num n"
   501   by (simp add: less_imp_le minus_one_less_of_num_iff)
   502 
   503 lemma minus_one_le_one_iff: "- 1 \<le> 1"
   504   by (simp add: less_imp_le minus_one_less_one_iff)
   505 
   506 lemma of_num_le_minus_of_num_iff: "\<not> of_num m \<le> - of_num n"
   507   by (simp add: not_le minus_of_num_less_of_num_iff)
   508 
   509 lemma one_le_minus_of_num_iff: "\<not> 1 \<le> - of_num n"
   510   by (simp add: not_le minus_of_num_less_one_iff)
   511 
   512 lemma of_num_le_minus_one_iff: "\<not> of_num n \<le> - 1"
   513   by (simp add: not_le minus_one_less_of_num_iff)
   514 
   515 lemma one_le_minus_one_iff: "\<not> 1 \<le> - 1"
   516   by (simp add: not_le minus_one_less_one_iff)
   517 
   518 lemma of_num_less_minus_of_num_iff: "\<not> of_num m < - of_num n"
   519   by (simp add: not_less minus_of_num_le_of_num_iff)
   520 
   521 lemma one_less_minus_of_num_iff: "\<not> 1 < - of_num n"
   522   by (simp add: not_less minus_of_num_le_one_iff)
   523 
   524 lemma of_num_less_minus_one_iff: "\<not> of_num n < - 1"
   525   by (simp add: not_less minus_one_le_of_num_iff)
   526 
   527 lemma one_less_minus_one_iff: "\<not> 1 < - 1"
   528   by (simp add: not_less minus_one_le_one_iff)
   529 
   530 lemmas le_signed_numeral_special [numeral] =
   531   minus_of_num_le_of_num_iff
   532   minus_of_num_le_one_iff
   533   minus_one_le_of_num_iff
   534   minus_one_le_one_iff
   535   of_num_le_minus_of_num_iff
   536   one_le_minus_of_num_iff
   537   of_num_le_minus_one_iff
   538   one_le_minus_one_iff
   539 
   540 lemmas less_signed_numeral_special [numeral] =
   541   minus_of_num_less_of_num_iff
   542   minus_of_num_not_equal_of_num
   543   minus_of_num_less_one_iff
   544   minus_one_less_of_num_iff
   545   minus_one_less_one_iff
   546   of_num_less_minus_of_num_iff
   547   one_less_minus_of_num_iff
   548   of_num_less_minus_one_iff
   549   one_less_minus_one_iff
   550 
   551 end
   552 
   553 subsubsection {* Structures with subtraction: class @{text semiring_1_minus} *}
   554 
   555 class semiring_minus = semiring + minus + zero +
   556   assumes minus_inverts_plus1: "a + b = c \<Longrightarrow> c - b = a"
   557   assumes minus_minus_zero_inverts_plus1: "a + b = c \<Longrightarrow> b - c = 0 - a"
   558 begin
   559 
   560 lemma minus_inverts_plus2: "a + b = c \<Longrightarrow> c - a = b"
   561   by (simp add: add_ac minus_inverts_plus1 [of b a])
   562 
   563 lemma minus_minus_zero_inverts_plus2: "a + b = c \<Longrightarrow> a - c = 0 - b"
   564   by (simp add: add_ac minus_minus_zero_inverts_plus1 [of b a])
   565 
   566 end
   567 
   568 class semiring_1_minus = semiring_1 + semiring_minus
   569 begin
   570 
   571 lemma Dig_of_num_pos:
   572   assumes "k + n = m"
   573   shows "of_num m - of_num n = of_num k"
   574   using assms by (simp add: of_num_plus minus_inverts_plus1)
   575 
   576 lemma Dig_of_num_zero:
   577   shows "of_num n - of_num n = 0"
   578   by (rule minus_inverts_plus1) simp
   579 
   580 lemma Dig_of_num_neg:
   581   assumes "k + m = n"
   582   shows "of_num m - of_num n = 0 - of_num k"
   583   by (rule minus_minus_zero_inverts_plus1) (simp add: of_num_plus assms)
   584 
   585 lemmas Dig_plus_eval =
   586   of_num_plus of_num_eq_iff Dig_plus refl [of One, THEN eqTrueI] num.inject
   587 
   588 simproc_setup numeral_minus ("of_num m - of_num n") = {*
   589   let
   590     (*TODO proper implicit use of morphism via pattern antiquotations*)
   591     fun cdest_of_num ct = (snd o split_last o snd o Drule.strip_comb) ct;
   592     fun cdest_minus ct = case (rev o snd o Drule.strip_comb) ct of [n, m] => (m, n);
   593     fun attach_num ct = (dest_num (Thm.term_of ct), ct);
   594     fun cdifference t = (pairself (attach_num o cdest_of_num) o cdest_minus) t;
   595     val simplify = MetaSimplifier.rewrite false (map mk_meta_eq @{thms Dig_plus_eval});
   596     fun cert ck cl cj = @{thm eqTrueE} OF [@{thm meta_eq_to_obj_eq}
   597       OF [simplify (Drule.list_comb (@{cterm "op = :: num \<Rightarrow> _"},
   598         [Drule.list_comb (@{cterm "op + :: num \<Rightarrow> _"}, [ck, cl]), cj]))]];
   599   in fn phi => fn _ => fn ct => case try cdifference ct
   600    of NONE => (NONE)
   601     | SOME ((k, ck), (l, cl)) => SOME (let val j = k - l in if j = 0
   602         then MetaSimplifier.rewrite false [mk_meta_eq (Morphism.thm phi @{thm Dig_of_num_zero})] ct
   603         else mk_meta_eq (let
   604           val cj = Thm.cterm_of (Thm.theory_of_cterm ct) (mk_num (abs j));
   605         in
   606           (if j > 0 then (Morphism.thm phi @{thm Dig_of_num_pos}) OF [cert cj cl ck]
   607           else (Morphism.thm phi @{thm Dig_of_num_neg}) OF [cert cj ck cl])
   608         end) end)
   609   end
   610 *}
   611 
   612 lemma Dig_of_num_minus_zero [numeral]:
   613   "of_num n - 0 = of_num n"
   614   by (simp add: minus_inverts_plus1)
   615 
   616 lemma Dig_one_minus_zero [numeral]:
   617   "1 - 0 = 1"
   618   by (simp add: minus_inverts_plus1)
   619 
   620 lemma Dig_one_minus_one [numeral]:
   621   "1 - 1 = 0"
   622   by (simp add: minus_inverts_plus1)
   623 
   624 lemma Dig_of_num_minus_one [numeral]:
   625   "of_num (Dig0 n) - 1 = of_num (DigM n)"
   626   "of_num (Dig1 n) - 1 = of_num (Dig0 n)"
   627   by (auto intro: minus_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one)
   628 
   629 lemma Dig_one_minus_of_num [numeral]:
   630   "1 - of_num (Dig0 n) = 0 - of_num (DigM n)"
   631   "1 - of_num (Dig1 n) = 0 - of_num (Dig0 n)"
   632   by (auto intro: minus_minus_zero_inverts_plus1 simp add: DigM_plus_one of_num.simps of_num_plus_one)
   633 
   634 end
   635 
   636 
   637 subsubsection {* Structures with negation: class @{text ring_1} *}
   638 
   639 context ring_1
   640 begin
   641 
   642 subclass semiring_1_minus proof
   643 qed (simp_all add: algebra_simps)
   644 
   645 lemma Dig_zero_minus_of_num [numeral]:
   646   "0 - of_num n = - of_num n"
   647   by simp
   648 
   649 lemma Dig_zero_minus_one [numeral]:
   650   "0 - 1 = - 1"
   651   by simp
   652 
   653 lemma Dig_uminus_uminus [numeral]:
   654   "- (- of_num n) = of_num n"
   655   by simp
   656 
   657 lemma Dig_plus_uminus [numeral]:
   658   "of_num m + - of_num n = of_num m - of_num n"
   659   "- of_num m + of_num n = of_num n - of_num m"
   660   "- of_num m + - of_num n = - (of_num m + of_num n)"
   661   "of_num m - - of_num n = of_num m + of_num n"
   662   "- of_num m - of_num n = - (of_num m + of_num n)"
   663   "- of_num m - - of_num n = of_num n - of_num m"
   664   by (simp_all add: diff_minus add_commute)
   665 
   666 lemma Dig_times_uminus [numeral]:
   667   "- of_num n * of_num m = - (of_num n * of_num m)"
   668   "of_num n * - of_num m = - (of_num n * of_num m)"
   669   "- of_num n * - of_num m = of_num n * of_num m"
   670   by simp_all
   671 
   672 lemma of_int_of_num [numeral]: "of_int (of_num n) = of_num n"
   673 by (induct n)
   674   (simp_all only: of_num.simps semiring_numeral_class.of_num.simps of_int_add, simp_all)
   675 
   676 declare of_int_1 [numeral]
   677 
   678 end
   679 
   680 
   681 subsubsection {* Structures with exponentiation *}
   682 
   683 lemma of_num_square: "of_num (square x) = of_num x * of_num x"
   684 by (induct x)
   685    (simp_all add: of_num.simps of_num_add algebra_simps)
   686 
   687 lemma of_num_pow: "of_num (pow x y) = of_num x ^ of_num y"
   688 by (induct y)
   689    (simp_all add: of_num.simps of_num_square of_num_mult power_add)
   690 
   691 lemma power_of_num [numeral]: "of_num x ^ of_num y = of_num (pow x y)"
   692   unfolding of_num_pow ..
   693 
   694 lemma power_zero_of_num [numeral]:
   695   "0 ^ of_num n = (0::'a::semiring_1)"
   696   using of_num_pos [where n=n and ?'a=nat]
   697   by (simp add: power_0_left)
   698 
   699 lemma power_minus_Dig0 [numeral]:
   700   fixes x :: "'a::ring_1"
   701   shows "(- x) ^ of_num (Dig0 n) = x ^ of_num (Dig0 n)"
   702   by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc)
   703 
   704 lemma power_minus_Dig1 [numeral]:
   705   fixes x :: "'a::ring_1"
   706   shows "(- x) ^ of_num (Dig1 n) = - (x ^ of_num (Dig1 n))"
   707   by (induct n rule: num_induct) (simp_all add: of_num.simps of_num_inc)
   708 
   709 declare power_one [numeral]
   710 
   711 
   712 subsubsection {* Greetings to @{typ nat}. *}
   713 
   714 instance nat :: semiring_1_minus proof
   715 qed simp_all
   716 
   717 lemma Suc_of_num [numeral]: "Suc (of_num n) = of_num (n + One)"
   718   unfolding of_num_plus_one [symmetric] by simp
   719 
   720 lemma nat_number:
   721   "1 = Suc 0"
   722   "of_num One = Suc 0"
   723   "of_num (Dig0 n) = Suc (of_num (DigM n))"
   724   "of_num (Dig1 n) = Suc (of_num (Dig0 n))"
   725   by (simp_all add: of_num.simps DigM_plus_one Suc_of_num)
   726 
   727 declare diff_0_eq_0 [numeral]
   728 
   729 
   730 subsection {* Proof tools setup *}
   731 
   732 subsubsection {* Numeral equations as default simplification rules *}
   733 
   734 declare (in semiring_numeral) of_num_One [simp]
   735 declare (in semiring_numeral) of_num_plus_one [simp]
   736 declare (in semiring_numeral) of_num_one_plus [simp]
   737 declare (in semiring_numeral) of_num_plus [simp]
   738 declare (in semiring_numeral) of_num_times [simp]
   739 
   740 declare (in semiring_1) of_nat_of_num [simp]
   741 
   742 declare (in semiring_char_0) of_num_eq_iff [simp]
   743 declare (in semiring_char_0) of_num_eq_one_iff [simp]
   744 declare (in semiring_char_0) one_eq_of_num_iff [simp]
   745 
   746 declare (in linordered_semidom) of_num_pos [simp]
   747 declare (in linordered_semidom) of_num_not_zero [simp]
   748 declare (in linordered_semidom) of_num_less_eq_iff [simp]
   749 declare (in linordered_semidom) of_num_less_eq_one_iff [simp]
   750 declare (in linordered_semidom) one_less_eq_of_num_iff [simp]
   751 declare (in linordered_semidom) of_num_less_iff [simp]
   752 declare (in linordered_semidom) of_num_less_one_iff [simp]
   753 declare (in linordered_semidom) one_less_of_num_iff [simp]
   754 declare (in linordered_semidom) of_num_nonneg [simp]
   755 declare (in linordered_semidom) of_num_less_zero_iff [simp]
   756 declare (in linordered_semidom) of_num_le_zero_iff [simp]
   757 
   758 declare (in linordered_idom) le_signed_numeral_special [simp]
   759 declare (in linordered_idom) less_signed_numeral_special [simp]
   760 
   761 declare (in semiring_1_minus) Dig_of_num_minus_one [simp]
   762 declare (in semiring_1_minus) Dig_one_minus_of_num [simp]
   763 
   764 declare (in ring_1) Dig_plus_uminus [simp]
   765 declare (in ring_1) of_int_of_num [simp]
   766 
   767 declare power_of_num [simp]
   768 declare power_zero_of_num [simp]
   769 declare power_minus_Dig0 [simp]
   770 declare power_minus_Dig1 [simp]
   771 
   772 declare Suc_of_num [simp]
   773 
   774 
   775 subsubsection {* Reorientation of equalities *}
   776 
   777 setup {*
   778   Reorient_Proc.add
   779     (fn Const(@{const_name of_num}, _) $ _ => true
   780       | Const(@{const_name uminus}, _) $
   781           (Const(@{const_name of_num}, _) $ _) => true
   782       | _ => false)
   783 *}
   784 
   785 simproc_setup reorient_num ("of_num n = x" | "- of_num m = y") = Reorient_Proc.proc
   786 
   787 
   788 subsubsection {* Constant folding for multiplication in semirings *}
   789 
   790 context semiring_numeral
   791 begin
   792 
   793 lemma mult_of_num_commute: "x * of_num n = of_num n * x"
   794 by (induct n)
   795   (simp_all only: of_num.simps left_distrib right_distrib mult_1_left mult_1_right)
   796 
   797 definition
   798   "commutes_with a b \<longleftrightarrow> a * b = b * a"
   799 
   800 lemma commutes_with_commute: "commutes_with a b \<Longrightarrow> a * b = b * a"
   801 unfolding commutes_with_def .
   802 
   803 lemma commutes_with_left_commute: "commutes_with a b \<Longrightarrow> a * (b * c) = b * (a * c)"
   804 unfolding commutes_with_def by (simp only: mult_assoc [symmetric])
   805 
   806 lemma commutes_with_numeral: "commutes_with x (of_num n)" "commutes_with (of_num n) x"
   807 unfolding commutes_with_def by (simp_all add: mult_of_num_commute)
   808 
   809 lemmas mult_ac_numeral =
   810   mult_assoc
   811   commutes_with_commute
   812   commutes_with_left_commute
   813   commutes_with_numeral
   814 
   815 end
   816 
   817 ML {*
   818 structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA =
   819 struct
   820   val assoc_ss = HOL_ss addsimps @{thms mult_ac_numeral}
   821   val eq_reflection = eq_reflection
   822   fun is_numeral (Const(@{const_name of_num}, _) $ _) = true
   823     | is_numeral _ = false;
   824 end;
   825 
   826 structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
   827 *}
   828 
   829 simproc_setup semiring_assoc_fold' ("(a::'a::semiring_numeral) * b") =
   830   {* fn phi => fn ss => fn ct =>
   831     Semiring_Times_Assoc.proc ss (Thm.term_of ct) *}
   832 
   833 
   834 subsection {* Code generator setup for @{typ int} *}
   835 
   836 text {* Reversing standard setup *}
   837 
   838 lemma [code_unfold del]: "(0::int) \<equiv> Numeral0" by simp
   839 lemma [code_unfold del]: "(1::int) \<equiv> Numeral1" by simp
   840 declare zero_is_num_zero [code_unfold del]
   841 declare one_is_num_one [code_unfold del]
   842   
   843 lemma [code, code del]:
   844   "(1 :: int) = 1"
   845   "(op + :: int \<Rightarrow> int \<Rightarrow> int) = op +"
   846   "(uminus :: int \<Rightarrow> int) = uminus"
   847   "(op - :: int \<Rightarrow> int \<Rightarrow> int) = op -"
   848   "(op * :: int \<Rightarrow> int \<Rightarrow> int) = op *"
   849   "(HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool) = HOL.equal"
   850   "(op \<le> :: int \<Rightarrow> int \<Rightarrow> bool) = op \<le>"
   851   "(op < :: int \<Rightarrow> int \<Rightarrow> bool) = op <"
   852   by rule+
   853 
   854 text {* Constructors *}
   855 
   856 definition Pls :: "num \<Rightarrow> int" where
   857   [simp, code_post]: "Pls n = of_num n"
   858 
   859 definition Mns :: "num \<Rightarrow> int" where
   860   [simp, code_post]: "Mns n = - of_num n"
   861 
   862 code_datatype "0::int" Pls Mns
   863 
   864 lemmas [code_unfold] = Pls_def [symmetric] Mns_def [symmetric]
   865 
   866 text {* Auxiliary operations *}
   867 
   868 definition dup :: "int \<Rightarrow> int" where
   869   [simp]: "dup k = k + k"
   870 
   871 lemma Dig_dup [code]:
   872   "dup 0 = 0"
   873   "dup (Pls n) = Pls (Dig0 n)"
   874   "dup (Mns n) = Mns (Dig0 n)"
   875   by (simp_all add: of_num.simps)
   876 
   877 definition sub :: "num \<Rightarrow> num \<Rightarrow> int" where
   878   [simp]: "sub m n = (of_num m - of_num n)"
   879 
   880 lemma Dig_sub [code]:
   881   "sub One One = 0"
   882   "sub (Dig0 m) One = of_num (DigM m)"
   883   "sub (Dig1 m) One = of_num (Dig0 m)"
   884   "sub One (Dig0 n) = - of_num (DigM n)"
   885   "sub One (Dig1 n) = - of_num (Dig0 n)"
   886   "sub (Dig0 m) (Dig0 n) = dup (sub m n)"
   887   "sub (Dig1 m) (Dig1 n) = dup (sub m n)"
   888   "sub (Dig1 m) (Dig0 n) = dup (sub m n) + 1"
   889   "sub (Dig0 m) (Dig1 n) = dup (sub m n) - 1"
   890   by (simp_all add: algebra_simps num_eq_iff nat_of_num_add)
   891 
   892 text {* Implementations *}
   893 
   894 lemma one_int_code [code]:
   895   "1 = Pls One"
   896   by (simp add: of_num_One)
   897 
   898 lemma plus_int_code [code]:
   899   "k + 0 = (k::int)"
   900   "0 + l = (l::int)"
   901   "Pls m + Pls n = Pls (m + n)"
   902   "Pls m + Mns n = sub m n"
   903   "Mns m + Pls n = sub n m"
   904   "Mns m + Mns n = Mns (m + n)"
   905   by simp_all
   906 
   907 lemma uminus_int_code [code]:
   908   "uminus 0 = (0::int)"
   909   "uminus (Pls m) = Mns m"
   910   "uminus (Mns m) = Pls m"
   911   by simp_all
   912 
   913 lemma minus_int_code [code]:
   914   "k - 0 = (k::int)"
   915   "0 - l = uminus (l::int)"
   916   "Pls m - Pls n = sub m n"
   917   "Pls m - Mns n = Pls (m + n)"
   918   "Mns m - Pls n = Mns (m + n)"
   919   "Mns m - Mns n = sub n m"
   920   by simp_all
   921 
   922 lemma times_int_code [code]:
   923   "k * 0 = (0::int)"
   924   "0 * l = (0::int)"
   925   "Pls m * Pls n = Pls (m * n)"
   926   "Pls m * Mns n = Mns (m * n)"
   927   "Mns m * Pls n = Mns (m * n)"
   928   "Mns m * Mns n = Pls (m * n)"
   929   by simp_all
   930 
   931 lemma eq_int_code [code]:
   932   "HOL.equal 0 (0::int) \<longleftrightarrow> True"
   933   "HOL.equal 0 (Pls l) \<longleftrightarrow> False"
   934   "HOL.equal 0 (Mns l) \<longleftrightarrow> False"
   935   "HOL.equal (Pls k) 0 \<longleftrightarrow> False"
   936   "HOL.equal (Pls k) (Pls l) \<longleftrightarrow> HOL.equal k l"
   937   "HOL.equal (Pls k) (Mns l) \<longleftrightarrow> False"
   938   "HOL.equal (Mns k) 0 \<longleftrightarrow> False"
   939   "HOL.equal (Mns k) (Pls l) \<longleftrightarrow> False"
   940   "HOL.equal (Mns k) (Mns l) \<longleftrightarrow> HOL.equal k l"
   941   by (auto simp add: equal dest: sym)
   942 
   943 lemma [code nbe]:
   944   "HOL.equal (k::int) k \<longleftrightarrow> True"
   945   by (fact equal_refl)
   946 
   947 lemma less_eq_int_code [code]:
   948   "0 \<le> (0::int) \<longleftrightarrow> True"
   949   "0 \<le> Pls l \<longleftrightarrow> True"
   950   "0 \<le> Mns l \<longleftrightarrow> False"
   951   "Pls k \<le> 0 \<longleftrightarrow> False"
   952   "Pls k \<le> Pls l \<longleftrightarrow> k \<le> l"
   953   "Pls k \<le> Mns l \<longleftrightarrow> False"
   954   "Mns k \<le> 0 \<longleftrightarrow> True"
   955   "Mns k \<le> Pls l \<longleftrightarrow> True"
   956   "Mns k \<le> Mns l \<longleftrightarrow> l \<le> k"
   957   by simp_all
   958 
   959 lemma less_int_code [code]:
   960   "0 < (0::int) \<longleftrightarrow> False"
   961   "0 < Pls l \<longleftrightarrow> True"
   962   "0 < Mns l \<longleftrightarrow> False"
   963   "Pls k < 0 \<longleftrightarrow> False"
   964   "Pls k < Pls l \<longleftrightarrow> k < l"
   965   "Pls k < Mns l \<longleftrightarrow> False"
   966   "Mns k < 0 \<longleftrightarrow> True"
   967   "Mns k < Pls l \<longleftrightarrow> True"
   968   "Mns k < Mns l \<longleftrightarrow> l < k"
   969   by simp_all
   970 
   971 hide_const (open) sub dup
   972 
   973 text {* Pretty literals *}
   974 
   975 ML {*
   976 local open Code_Thingol in
   977 
   978 fun add_code print target =
   979   let
   980     fun dest_num one' dig0' dig1' thm =
   981       let
   982         fun dest_dig (IConst (c, _)) = if c = dig0' then 0
   983               else if c = dig1' then 1
   984               else Code_Printer.eqn_error thm "Illegal numeral expression: illegal dig"
   985           | dest_dig _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal digit";
   986         fun dest_num (IConst (c, _)) = if c = one' then 1
   987               else Code_Printer.eqn_error thm "Illegal numeral expression: illegal leading digit"
   988           | dest_num (t1 `$ t2) = 2 * dest_num t2 + dest_dig t1
   989           | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term";
   990       in dest_num end;
   991     fun pretty sgn literals [one', dig0', dig1'] _ thm _ _ [(t, _)] =
   992       (Code_Printer.str o print literals o sgn o dest_num one' dig0' dig1' thm) t
   993     fun add_syntax (c, sgn) = Code_Target.add_const_syntax target c
   994       (SOME (Code_Printer.complex_const_syntax
   995         (1, ([@{const_name One}, @{const_name Dig0}, @{const_name Dig1}],
   996           pretty sgn))));
   997   in
   998     add_syntax (@{const_name Pls}, I)
   999     #> add_syntax (@{const_name Mns}, (fn k => ~ k))
  1000   end;
  1001 
  1002 end
  1003 *}
  1004 
  1005 hide_const (open) One Dig0 Dig1
  1006 
  1007 
  1008 subsection {* Toy examples *}
  1009 
  1010 definition "foo \<longleftrightarrow> #4 * #2 + #7 = (#8 :: nat)"
  1011 definition "bar \<longleftrightarrow> #4 * #2 + #7 \<ge> (#8 :: int) - #3"
  1012 
  1013 code_thms foo bar
  1014 export_code foo bar checking SML OCaml? Haskell? Scala?
  1015 
  1016 text {* This is an ad-hoc @{text Code_Integer} setup. *}
  1017 
  1018 setup {*
  1019   fold (add_code Code_Printer.literal_numeral)
  1020     [Code_ML.target_SML, Code_ML.target_OCaml, Code_Haskell.target, Code_Scala.target]
  1021 *}
  1022 
  1023 code_type int
  1024   (SML "IntInf.int")
  1025   (OCaml "Big'_int.big'_int")
  1026   (Haskell "Integer")
  1027   (Scala "BigInt")
  1028   (Eval "int")
  1029 
  1030 code_const "0::int"
  1031   (SML "0/ :/ IntInf.int")
  1032   (OCaml "Big'_int.zero")
  1033   (Haskell "0")
  1034   (Scala "BigInt(0)")
  1035   (Eval "0/ :/ int")
  1036 
  1037 code_const Int.pred
  1038   (SML "IntInf.- ((_), 1)")
  1039   (OCaml "Big'_int.pred'_big'_int")
  1040   (Haskell "!(_/ -/ 1)")
  1041   (Scala "!(_ -/ 1)")
  1042   (Eval "!(_/ -/ 1)")
  1043 
  1044 code_const Int.succ
  1045   (SML "IntInf.+ ((_), 1)")
  1046   (OCaml "Big'_int.succ'_big'_int")
  1047   (Haskell "!(_/ +/ 1)")
  1048   (Scala "!(_ +/ 1)")
  1049   (Eval "!(_/ +/ 1)")
  1050 
  1051 code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
  1052   (SML "IntInf.+ ((_), (_))")
  1053   (OCaml "Big'_int.add'_big'_int")
  1054   (Haskell infixl 6 "+")
  1055   (Scala infixl 7 "+")
  1056   (Eval infixl 8 "+")
  1057 
  1058 code_const "uminus \<Colon> int \<Rightarrow> int"
  1059   (SML "IntInf.~")
  1060   (OCaml "Big'_int.minus'_big'_int")
  1061   (Haskell "negate")
  1062   (Scala "!(- _)")
  1063   (Eval "~/ _")
  1064 
  1065 code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int"
  1066   (SML "IntInf.- ((_), (_))")
  1067   (OCaml "Big'_int.sub'_big'_int")
  1068   (Haskell infixl 6 "-")
  1069   (Scala infixl 7 "-")
  1070   (Eval infixl 8 "-")
  1071 
  1072 code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"
  1073   (SML "IntInf.* ((_), (_))")
  1074   (OCaml "Big'_int.mult'_big'_int")
  1075   (Haskell infixl 7 "*")
  1076   (Scala infixl 8 "*")
  1077   (Eval infixl 9 "*")
  1078 
  1079 code_const pdivmod
  1080   (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)")
  1081   (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
  1082   (Haskell "divMod/ (abs _)/ (abs _)")
  1083   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
  1084   (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
  1085 
  1086 code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
  1087   (SML "!((_ : IntInf.int) = _)")
  1088   (OCaml "Big'_int.eq'_big'_int")
  1089   (Haskell infix 4 "==")
  1090   (Scala infixl 5 "==")
  1091   (Eval infixl 6 "=")
  1092 
  1093 code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
  1094   (SML "IntInf.<= ((_), (_))")
  1095   (OCaml "Big'_int.le'_big'_int")
  1096   (Haskell infix 4 "<=")
  1097   (Scala infixl 4 "<=")
  1098   (Eval infixl 6 "<=")
  1099 
  1100 code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
  1101   (SML "IntInf.< ((_), (_))")
  1102   (OCaml "Big'_int.lt'_big'_int")
  1103   (Haskell infix 4 "<")
  1104   (Scala infixl 4 "<")
  1105   (Eval infixl 6 "<")
  1106 
  1107 code_const Code_Numeral.int_of
  1108   (SML "IntInf.fromInt")
  1109   (OCaml "_")
  1110   (Haskell "toInteger")
  1111   (Scala "!_.as'_BigInt")
  1112   (Eval "_")
  1113 
  1114 export_code foo bar checking SML OCaml? Haskell? Scala?
  1115 
  1116 end