| author | wenzelm | 
| Mon, 28 Dec 2009 22:58:25 +0100 | |
| changeset 34203 | dd2f49d88b47 | 
| parent 32657 | 5f13912245ff | 
| child 34899 | 8674bb6f727b | 
| 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  | 
|
| 
31203
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31192 
diff
changeset
 | 
8  | 
imports Main  | 
| 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")  | 
|
21  | 
||
22  | 
code_instance int :: eq  | 
|
23  | 
(Haskell -)  | 
|
24  | 
||
25  | 
setup {*
 | 
|
| 25928 | 26  | 
  fold (Numeral.add_code @{const_name number_int_inst.number_of_int}
 | 
27  | 
true true) ["SML", "OCaml", "Haskell"]  | 
|
| 24999 | 28  | 
*}  | 
29  | 
||
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26009 
diff
changeset
 | 
30  | 
code_const "Int.Pls" and "Int.Min" and "Int.Bit0" and "Int.Bit1"  | 
| 24999 | 31  | 
(SML "raise/ Fail/ \"Pls\""  | 
32  | 
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
 | 
33  | 
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
 | 
34  | 
and "!((_);/ raise/ Fail/ \"Bit1\")")  | 
| 24999 | 35  | 
(OCaml "failwith/ \"Pls\""  | 
36  | 
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
 | 
37  | 
and "!((_);/ failwith/ \"Bit0\")"  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26009 
diff
changeset
 | 
38  | 
and "!((_);/ failwith/ \"Bit1\")")  | 
| 24999 | 39  | 
(Haskell "error/ \"Pls\""  | 
40  | 
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
 | 
41  | 
and "error/ \"Bit0\""  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26009 
diff
changeset
 | 
42  | 
and "error/ \"Bit1\"")  | 
| 24999 | 43  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25767 
diff
changeset
 | 
44  | 
code_const Int.pred  | 
| 24999 | 45  | 
(SML "IntInf.- ((_), 1)")  | 
46  | 
(OCaml "Big'_int.pred'_big'_int")  | 
|
47  | 
(Haskell "!(_/ -/ 1)")  | 
|
48  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25767 
diff
changeset
 | 
49  | 
code_const Int.succ  | 
| 24999 | 50  | 
(SML "IntInf.+ ((_), 1)")  | 
51  | 
(OCaml "Big'_int.succ'_big'_int")  | 
|
52  | 
(Haskell "!(_/ +/ 1)")  | 
|
53  | 
||
54  | 
code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"  | 
|
55  | 
(SML "IntInf.+ ((_), (_))")  | 
|
56  | 
(OCaml "Big'_int.add'_big'_int")  | 
|
57  | 
(Haskell infixl 6 "+")  | 
|
58  | 
||
59  | 
code_const "uminus \<Colon> int \<Rightarrow> int"  | 
|
60  | 
(SML "IntInf.~")  | 
|
61  | 
(OCaml "Big'_int.minus'_big'_int")  | 
|
62  | 
(Haskell "negate")  | 
|
63  | 
||
64  | 
code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int"  | 
|
65  | 
(SML "IntInf.- ((_), (_))")  | 
|
66  | 
(OCaml "Big'_int.sub'_big'_int")  | 
|
67  | 
(Haskell infixl 6 "-")  | 
|
68  | 
||
69  | 
code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"  | 
|
70  | 
(SML "IntInf.* ((_), (_))")  | 
|
71  | 
(OCaml "Big'_int.mult'_big'_int")  | 
|
72  | 
(Haskell infixl 7 "*")  | 
|
73  | 
||
| 29936 | 74  | 
code_const pdivmod  | 
75  | 
(SML "(fn k => fn l =>/ IntInf.divMod/ (IntInf.abs k,/ IntInf.abs l))")  | 
|
76  | 
(OCaml "(fun k -> fun l ->/ Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int k)/ (Big'_int.abs'_big'_int l))")  | 
|
77  | 
(Haskell "(\\k l ->/ divMod/ (abs k)/ (abs l))")  | 
|
78  | 
||
| 
28346
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28228 
diff
changeset
 | 
79  | 
code_const "eq_class.eq \<Colon> int \<Rightarrow> int \<Rightarrow> bool"  | 
| 24999 | 80  | 
(SML "!((_ : IntInf.int) = _)")  | 
81  | 
(OCaml "Big'_int.eq'_big'_int")  | 
|
82  | 
(Haskell infixl 4 "==")  | 
|
83  | 
||
84  | 
code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool"  | 
|
85  | 
(SML "IntInf.<= ((_), (_))")  | 
|
86  | 
(OCaml "Big'_int.le'_big'_int")  | 
|
87  | 
(Haskell infix 4 "<=")  | 
|
88  | 
||
89  | 
code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool"  | 
|
90  | 
(SML "IntInf.< ((_), (_))")  | 
|
91  | 
(OCaml "Big'_int.lt'_big'_int")  | 
|
92  | 
(Haskell infix 4 "<")  | 
|
93  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
94  | 
code_const Code_Numeral.int_of  | 
| 31192 | 95  | 
(SML "IntInf.fromInt")  | 
| 31377 | 96  | 
(OCaml "_")  | 
| 31192 | 97  | 
(Haskell "toEnum")  | 
98  | 
||
| 24999 | 99  | 
code_reserved SML IntInf  | 
100  | 
||
| 28228 | 101  | 
text {* Evaluation *}
 | 
102  | 
||
| 32657 | 103  | 
code_const "Code_Evaluation.term_of \<Colon> int \<Rightarrow> term"  | 
| 31192 | 104  | 
(Eval "HOLogic.mk'_number/ HOLogic.intT")  | 
| 28228 | 105  | 
|
| 24999 | 106  | 
end  |