author | huffman |
Fri, 02 Sep 2011 20:58:31 -0700 | |
changeset 44678 | 21eb31192850 |
parent 39272 | 0b61951d2682 |
child 47108 | 2a1953f0d20d |
permissions | -rw-r--r-- |
24999 | 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 |
|
37968 | 8 |
imports Main Code_Natural |
24999 | 9 |
begin |
10 |
||
11 |
text {* |
|
12 |
HOL numeral expressions are mapped to integer literals |
|
13 |
in target languages, using predefined target language |
|
14 |
operations for abstract integer operations. |
|
15 |
*} |
|
16 |
||
17 |
code_type int |
|
18 |
(SML "IntInf.int") |
|
19 |
(OCaml "Big'_int.big'_int") |
|
20 |
(Haskell "Integer") |
|
34899 | 21 |
(Scala "BigInt") |
38053 | 22 |
(Eval "int") |
24999 | 23 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38773
diff
changeset
|
24 |
code_instance int :: equal |
24999 | 25 |
(Haskell -) |
26 |
||
27 |
setup {* |
|
25928 | 28 |
fold (Numeral.add_code @{const_name number_int_inst.number_of_int} |
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34899
diff
changeset
|
29 |
true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] |
24999 | 30 |
*} |
31 |
||
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
32 |
code_const "Int.Pls" and "Int.Min" and "Int.Bit0" and "Int.Bit1" |
24999 | 33 |
(SML "raise/ Fail/ \"Pls\"" |
34 |
and "raise/ Fail/ \"Min\"" |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
35 |
and "!((_);/ raise/ Fail/ \"Bit0\")" |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
36 |
and "!((_);/ raise/ Fail/ \"Bit1\")") |
24999 | 37 |
(OCaml "failwith/ \"Pls\"" |
38 |
and "failwith/ \"Min\"" |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
39 |
and "!((_);/ failwith/ \"Bit0\")" |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
40 |
and "!((_);/ failwith/ \"Bit1\")") |
24999 | 41 |
(Haskell "error/ \"Pls\"" |
42 |
and "error/ \"Min\"" |
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
43 |
and "error/ \"Bit0\"" |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26009
diff
changeset
|
44 |
and "error/ \"Bit1\"") |
34899 | 45 |
(Scala "!error(\"Pls\")" |
46 |
and "!error(\"Min\")" |
|
47 |
and "!error(\"Bit0\")" |
|
48 |
and "!error(\"Bit1\")") |
|
49 |
||
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25767
diff
changeset
|
50 |
code_const Int.pred |
24999 | 51 |
(SML "IntInf.- ((_), 1)") |
52 |
(OCaml "Big'_int.pred'_big'_int") |
|
53 |
(Haskell "!(_/ -/ 1)") |
|
38773
f9837065b5e8
prevent line breaks after Scala symbolic operators
haftmann
parents:
38053
diff
changeset
|
54 |
(Scala "!(_ -/ 1)") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
55 |
(Eval "!(_/ -/ 1)") |
24999 | 56 |
|
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25767
diff
changeset
|
57 |
code_const Int.succ |
24999 | 58 |
(SML "IntInf.+ ((_), 1)") |
59 |
(OCaml "Big'_int.succ'_big'_int") |
|
60 |
(Haskell "!(_/ +/ 1)") |
|
38773
f9837065b5e8
prevent line breaks after Scala symbolic operators
haftmann
parents:
38053
diff
changeset
|
61 |
(Scala "!(_ +/ 1)") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
62 |
(Eval "!(_/ +/ 1)") |
24999 | 63 |
|
64 |
code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int" |
|
65 |
(SML "IntInf.+ ((_), (_))") |
|
66 |
(OCaml "Big'_int.add'_big'_int") |
|
67 |
(Haskell infixl 6 "+") |
|
34899 | 68 |
(Scala infixl 7 "+") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
69 |
(Eval infixl 8 "+") |
24999 | 70 |
|
71 |
code_const "uminus \<Colon> int \<Rightarrow> int" |
|
72 |
(SML "IntInf.~") |
|
73 |
(OCaml "Big'_int.minus'_big'_int") |
|
74 |
(Haskell "negate") |
|
34899 | 75 |
(Scala "!(- _)") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
76 |
(Eval "~/ _") |
24999 | 77 |
|
78 |
code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int" |
|
79 |
(SML "IntInf.- ((_), (_))") |
|
80 |
(OCaml "Big'_int.sub'_big'_int") |
|
81 |
(Haskell infixl 6 "-") |
|
34899 | 82 |
(Scala infixl 7 "-") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
83 |
(Eval infixl 8 "-") |
24999 | 84 |
|
85 |
code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int" |
|
86 |
(SML "IntInf.* ((_), (_))") |
|
87 |
(OCaml "Big'_int.mult'_big'_int") |
|
88 |
(Haskell infixl 7 "*") |
|
34899 | 89 |
(Scala infixl 8 "*") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
90 |
(Eval infixl 9 "*") |
24999 | 91 |
|
29936 | 92 |
code_const pdivmod |
34899 | 93 |
(SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") |
94 |
(OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") |
|
34944
970e1466028d
code literals: distinguish numeral classes by different entries
haftmann
parents:
34899
diff
changeset
|
95 |
(Haskell "divMod/ (abs _)/ (abs _)") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
96 |
(Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") |
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
97 |
(Eval "Integer.div'_mod/ (abs _)/ (abs _)") |
29936 | 98 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38773
diff
changeset
|
99 |
code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool" |
24999 | 100 |
(SML "!((_ : IntInf.int) = _)") |
101 |
(OCaml "Big'_int.eq'_big'_int") |
|
39272 | 102 |
(Haskell infix 4 "==") |
34899 | 103 |
(Scala infixl 5 "==") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
104 |
(Eval infixl 6 "=") |
24999 | 105 |
|
106 |
code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool" |
|
107 |
(SML "IntInf.<= ((_), (_))") |
|
108 |
(OCaml "Big'_int.le'_big'_int") |
|
109 |
(Haskell infix 4 "<=") |
|
34899 | 110 |
(Scala infixl 4 "<=") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
111 |
(Eval infixl 6 "<=") |
24999 | 112 |
|
113 |
code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool" |
|
114 |
(SML "IntInf.< ((_), (_))") |
|
115 |
(OCaml "Big'_int.lt'_big'_int") |
|
116 |
(Haskell infix 4 "<") |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
117 |
(Scala infixl 4 "<") |
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
118 |
(Eval infixl 6 "<") |
24999 | 119 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
120 |
code_const Code_Numeral.int_of |
31192 | 121 |
(SML "IntInf.fromInt") |
31377 | 122 |
(OCaml "_") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
123 |
(Haskell "toInteger") |
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
124 |
(Scala "!_.as'_BigInt") |
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
125 |
(Eval "_") |
24999 | 126 |
|
28228 | 127 |
text {* Evaluation *} |
128 |
||
32657 | 129 |
code_const "Code_Evaluation.term_of \<Colon> int \<Rightarrow> term" |
31192 | 130 |
(Eval "HOLogic.mk'_number/ HOLogic.intT") |
28228 | 131 |
|
24999 | 132 |
end |