haftmann@24999: (* Title: HOL/Library/Code_Integer.thy haftmann@24999: ID: $Id$ haftmann@24999: Author: Florian Haftmann, TU Muenchen haftmann@24999: *) haftmann@24999: haftmann@24999: header {* Pretty integer literals for code generation *} haftmann@24999: haftmann@24999: theory Code_Integer haftmann@26466: imports ATP_Linkup haftmann@24999: begin haftmann@24999: haftmann@24999: text {* haftmann@24999: HOL numeral expressions are mapped to integer literals haftmann@24999: in target languages, using predefined target language haftmann@24999: operations for abstract integer operations. haftmann@24999: *} haftmann@24999: haftmann@24999: code_type int haftmann@24999: (SML "IntInf.int") haftmann@24999: (OCaml "Big'_int.big'_int") haftmann@24999: (Haskell "Integer") haftmann@24999: haftmann@24999: code_instance int :: eq haftmann@24999: (Haskell -) haftmann@24999: haftmann@24999: setup {* haftmann@25928: fold (Numeral.add_code @{const_name number_int_inst.number_of_int} haftmann@25928: true true) ["SML", "OCaml", "Haskell"] haftmann@24999: *} haftmann@24999: huffman@26086: code_const "Int.Pls" and "Int.Min" and "Int.Bit0" and "Int.Bit1" haftmann@24999: (SML "raise/ Fail/ \"Pls\"" haftmann@24999: and "raise/ Fail/ \"Min\"" huffman@26086: and "!((_);/ raise/ Fail/ \"Bit0\")" huffman@26086: and "!((_);/ raise/ Fail/ \"Bit1\")") haftmann@24999: (OCaml "failwith/ \"Pls\"" haftmann@24999: and "failwith/ \"Min\"" huffman@26086: and "!((_);/ failwith/ \"Bit0\")" huffman@26086: and "!((_);/ failwith/ \"Bit1\")") haftmann@24999: (Haskell "error/ \"Pls\"" haftmann@24999: and "error/ \"Min\"" huffman@26086: and "error/ \"Bit0\"" huffman@26086: and "error/ \"Bit1\"") haftmann@24999: haftmann@25919: code_const Int.pred haftmann@24999: (SML "IntInf.- ((_), 1)") haftmann@24999: (OCaml "Big'_int.pred'_big'_int") haftmann@24999: (Haskell "!(_/ -/ 1)") haftmann@24999: haftmann@25919: code_const Int.succ haftmann@24999: (SML "IntInf.+ ((_), 1)") haftmann@24999: (OCaml "Big'_int.succ'_big'_int") haftmann@24999: (Haskell "!(_/ +/ 1)") haftmann@24999: haftmann@24999: code_const "op + \ int \ int \ int" haftmann@24999: (SML "IntInf.+ ((_), (_))") haftmann@24999: (OCaml "Big'_int.add'_big'_int") haftmann@24999: (Haskell infixl 6 "+") haftmann@24999: haftmann@24999: code_const "uminus \ int \ int" haftmann@24999: (SML "IntInf.~") haftmann@24999: (OCaml "Big'_int.minus'_big'_int") haftmann@24999: (Haskell "negate") haftmann@24999: haftmann@24999: code_const "op - \ int \ int \ int" haftmann@24999: (SML "IntInf.- ((_), (_))") haftmann@24999: (OCaml "Big'_int.sub'_big'_int") haftmann@24999: (Haskell infixl 6 "-") haftmann@24999: haftmann@24999: code_const "op * \ int \ int \ int" haftmann@24999: (SML "IntInf.* ((_), (_))") haftmann@24999: (OCaml "Big'_int.mult'_big'_int") haftmann@24999: (Haskell infixl 7 "*") haftmann@24999: haftmann@24999: code_const "op = \ int \ int \ bool" haftmann@24999: (SML "!((_ : IntInf.int) = _)") haftmann@24999: (OCaml "Big'_int.eq'_big'_int") haftmann@24999: (Haskell infixl 4 "==") haftmann@24999: haftmann@24999: code_const "op \ \ int \ int \ bool" haftmann@24999: (SML "IntInf.<= ((_), (_))") haftmann@24999: (OCaml "Big'_int.le'_big'_int") haftmann@24999: (Haskell infix 4 "<=") haftmann@24999: haftmann@24999: code_const "op < \ int \ int \ bool" haftmann@24999: (SML "IntInf.< ((_), (_))") haftmann@24999: (OCaml "Big'_int.lt'_big'_int") haftmann@24999: (Haskell infix 4 "<") haftmann@24999: haftmann@24999: code_reserved SML IntInf haftmann@24999: code_reserved OCaml Big_int haftmann@24999: haftmann@24999: end