author | wenzelm |
Thu, 11 Jun 2009 14:25:58 +0200 | |
changeset 31556 | ac0badf13d93 |
parent 31459 | ae39b7b2a68a |
child 31739 | 8155c4d94354 |
permissions | -rw-r--r-- |
31378 | 1 |
|
2 |
(* Author: Florian Haftmann, TU Muenchen *) |
|
21917 | 3 |
|
31378 | 4 |
header {* A huge collection of equations to generate code from *} |
21917 | 5 |
|
31378 | 6 |
theory Codegenerator_Candidates |
21917 | 7 |
imports |
27421 | 8 |
Complex_Main |
21917 | 9 |
AssocList |
10 |
Binomial |
|
11 |
Commutative_Ring |
|
27421 | 12 |
Enum |
21917 | 13 |
List_Prefix |
14 |
Nat_Infinity |
|
24433
4a405457e9d6
added explicit equation for equality of nested environments
haftmann
parents:
24423
diff
changeset
|
15 |
Nested_Environment |
26515 | 16 |
Option_ord |
21917 | 17 |
Permutation |
18 |
Primes |
|
19 |
Product_ord |
|
20 |
SetsAndFunctions |
|
31459 | 21 |
Tree |
21917 | 22 |
While_Combinator |
23 |
Word |
|
27421 | 24 |
"~~/src/HOL/ex/Commutative_Ring_Complete" |
25 |
"~~/src/HOL/ex/Records" |
|
21917 | 26 |
begin |
27 |
||
31378 | 28 |
(*temporary Haskell workaround*) |
29 |
declare typerep_fun_def [code inline] |
|
30 |
declare typerep_prod_def [code inline] |
|
31 |
declare typerep_sum_def [code inline] |
|
32 |
declare typerep_cpoint_ext_type_def [code inline] |
|
33 |
declare typerep_itself_def [code inline] |
|
34 |
declare typerep_list_def [code inline] |
|
35 |
declare typerep_option_def [code inline] |
|
36 |
declare typerep_point_ext_type_def [code inline] |
|
37 |
declare typerep_point'_ext_type_def [code inline] |
|
38 |
declare typerep_point''_ext_type_def [code inline] |
|
39 |
declare typerep_pol_def [code inline] |
|
40 |
declare typerep_polex_def [code inline] |
|
41 |
||
42 |
(*avoid popular infixes*) |
|
43 |
code_reserved SML union inter upto |
|
44 |
||
21917 | 45 |
end |