src/HOL/Library/Code_Integer.thy
author huffman
Sun, 25 Mar 2012 20:15:39 +0200
changeset 47978 2a1953f0d20d
parent 39499 0b61951d2682
child 47995 a3a64240cd98
permissions -rw-r--r--
merged fork with new numeral representation (see NEWS)
     1 (*  Title:      HOL/Library/Code_Integer.thy
     2     Author:     Florian Haftmann, TU Muenchen
     3 *)
     4 
     5 header {* Pretty integer literals for code generation *}
     6 
     7 theory Code_Integer
     8 imports Main Code_Natural
     9 begin
    10 
    11 text {*
    12   Representation-ignorant code equations for conversions.
    13 *}
    14 
    15 lemma nat_code [code]:
    16   "nat k = (if k \<le> 0 then 0 else
    17      let
    18        (l, j) = divmod_int k 2;
    19        l' = 2 * nat l
    20      in if j = 0 then l' else Suc l')"
    21 proof -
    22   have "2 = nat 2" by simp
    23   show ?thesis
    24     apply (auto simp add: Let_def divmod_int_mod_div not_le
    25      nat_div_distrib nat_mult_distrib mult_div_cancel mod_2_not_eq_zero_eq_one_int)
    26     apply (unfold `2 = nat 2`)
    27     apply (subst nat_mod_distrib [symmetric])
    28     apply simp_all
    29   done
    30 qed
    31 
    32 lemma (in ring_1) of_int_code:
    33   "of_int k = (if k = 0 then 0
    34      else if k < 0 then - of_int (- k)
    35      else let
    36        (l, j) = divmod_int k 2;
    37        l' = 2 * of_int l
    38      in if j = 0 then l' else l' + 1)"
    39 proof -
    40   from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
    41   show ?thesis
    42     by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int
    43       of_int_add [symmetric]) (simp add: * mult_commute)
    44 qed
    45 
    46 declare of_int_code [code]
    47 
    48 text {*
    49   HOL numeral expressions are mapped to integer literals
    50   in target languages, using predefined target language
    51   operations for abstract integer operations.
    52 *}
    53 
    54 code_type int
    55   (SML "IntInf.int")
    56   (OCaml "Big'_int.big'_int")
    57   (Haskell "Integer")
    58   (Scala "BigInt")
    59   (Eval "int")
    60 
    61 code_instance int :: equal
    62   (Haskell -)
    63 
    64 code_const "0::int"
    65   (SML "0")
    66   (OCaml "Big'_int.zero'_big'_int")
    67   (Haskell "0")
    68   (Scala "BigInt(0)")
    69 
    70 setup {*
    71   fold (Numeral.add_code @{const_name Int.Pos}
    72     false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
    73 *}
    74 
    75 setup {*
    76   fold (Numeral.add_code @{const_name Int.Neg}
    77     true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
    78 *}
    79 
    80 code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
    81   (SML "IntInf.+ ((_), (_))")
    82   (OCaml "Big'_int.add'_big'_int")
    83   (Haskell infixl 6 "+")
    84   (Scala infixl 7 "+")
    85   (Eval infixl 8 "+")
    86 
    87 code_const "uminus \<Colon> int \<Rightarrow> int"
    88   (SML "IntInf.~")
    89   (OCaml "Big'_int.minus'_big'_int")
    90   (Haskell "negate")
    91   (Scala "!(- _)")
    92   (Eval "~/ _")
    93 
    94 code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int"
    95   (SML "IntInf.- ((_), (_))")
    96   (OCaml "Big'_int.sub'_big'_int")
    97   (Haskell infixl 6 "-")
    98   (Scala infixl 7 "-")
    99   (Eval infixl 8 "-")
   100 
   101 code_const Int.dup
   102   (SML "IntInf.*/ (2,/ (_))")
   103   (OCaml "Big'_int.mult'_big'_int/ 2")
   104   (Haskell "!(2 * _)")
   105   (Scala "!(2 * _)")
   106   (Eval "!(2 * _)")
   107 
   108 code_const Int.sub
   109   (SML "!(raise/ Fail/ \"sub\")")
   110   (OCaml "failwith/ \"sub\"")
   111   (Haskell "error/ \"sub\"")
   112   (Scala "!error(\"sub\")")
   113 
   114 code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"
   115   (SML "IntInf.* ((_), (_))")
   116   (OCaml "Big'_int.mult'_big'_int")
   117   (Haskell infixl 7 "*")
   118   (Scala infixl 8 "*")
   119   (Eval infixl 9 "*")
   120 
   121 code_const pdivmod
   122   (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)")
   123   (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
   124   (Haskell "divMod/ (abs _)/ (abs _)")
   125   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
   126   (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
   127 
   128 code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
   129   (SML "!((_ : IntInf.int) = _)")
   130   (OCaml "Big'_int.eq'_big'_int")
   131   (Haskell infix 4 "==")
   132   (Scala infixl 5 "==")
   133   (Eval infixl 6 "=")
   134 
   135 code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
   136   (SML "IntInf.<= ((_), (_))")
   137   (OCaml "Big'_int.le'_big'_int")
   138   (Haskell infix 4 "<=")
   139   (Scala infixl 4 "<=")
   140   (Eval infixl 6 "<=")
   141 
   142 code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
   143   (SML "IntInf.< ((_), (_))")
   144   (OCaml "Big'_int.lt'_big'_int")
   145   (Haskell infix 4 "<")
   146   (Scala infixl 4 "<")
   147   (Eval infixl 6 "<")
   148 
   149 code_const Code_Numeral.int_of
   150   (SML "IntInf.fromInt")
   151   (OCaml "_")
   152   (Haskell "toInteger")
   153   (Scala "!_.as'_BigInt")
   154   (Eval "_")
   155 
   156 code_const "Code_Evaluation.term_of \<Colon> int \<Rightarrow> term"
   157   (Eval "HOLogic.mk'_number/ HOLogic.intT")
   158 
   159 end