| author | wenzelm | 
| Sun, 14 May 2017 22:07:16 +0200 | |
| changeset 65839 | d081671d4a87 | 
| parent 63167 | 0909deb8059b | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 31378 | 1  | 
|
2  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
|
| 24195 | 3  | 
|
| 63167 | 4  | 
section \<open>Pervasive test of code generator\<close>  | 
| 24195 | 5  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
51143 
diff
changeset
 | 
6  | 
theory Generate_Binary_Nat  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
51143 
diff
changeset
 | 
7  | 
imports  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
51143 
diff
changeset
 | 
8  | 
Candidates  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
51143 
diff
changeset
 | 
9  | 
"~~/src/HOL/Library/AList_Mapping"  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
51143 
diff
changeset
 | 
10  | 
"~~/src/HOL/Library/Finite_Lattice"  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
51143 
diff
changeset
 | 
11  | 
"~~/src/HOL/Library/Code_Binary_Nat"  | 
| 24195 | 12  | 
begin  | 
13  | 
||
| 63167 | 14  | 
text \<open>  | 
| 37824 | 15  | 
If any of the checks fails, inspect the code generated  | 
| 63167 | 16  | 
by a corresponding \<open>export_code\<close> command.  | 
17  | 
\<close>  | 
|
| 
37745
 
6315b6426200
checking generated code for various target languages
 
haftmann 
parents: 
37695 
diff
changeset
 | 
18  | 
|
| 
58678
 
398e05aa84d4
purely algebraic characterization of even and odd
 
haftmann 
parents: 
54228 
diff
changeset
 | 
19  | 
export_code _ checking SML OCaml? Haskell? Scala  | 
| 25616 | 20  | 
|
| 24195 | 21  | 
end  |