| author | paulson <lp15@cam.ac.uk> | 
| Wed, 17 Apr 2019 17:48:28 +0100 | |
| changeset 70178 | 4900351361b0 | 
| parent 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 
19281
 
b411f25fff25
added example for operational classes and code generator
 
haftmann 
parents:  
diff
changeset
 | 
1  | 
|
| 31378 | 2  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 
19281
 
b411f25fff25
added example for operational classes and code generator
 
haftmann 
parents:  
diff
changeset
 | 
3  | 
|
| 63167 | 4  | 
section \<open>Pervasive test of code generator\<close>  | 
| 
19281
 
b411f25fff25
added example for operational classes and code generator
 
haftmann 
parents:  
diff
changeset
 | 
5  | 
|
| 
37695
 
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
 
haftmann 
parents: 
34306 
diff
changeset
 | 
6  | 
theory Generate  | 
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
50629 
diff
changeset
 | 
7  | 
imports  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
50629 
diff
changeset
 | 
8  | 
Candidates  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
9  | 
"HOL-Library.AList_Mapping"  | 
| 
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
10  | 
"HOL-Library.Finite_Lattice"  | 
| 31378 | 11  | 
begin  | 
| 24430 | 12  | 
|
| 63167 | 13  | 
text \<open>  | 
| 37824 | 14  | 
If any of the checks fails, inspect the code generated  | 
| 63167 | 15  | 
by a corresponding \<open>export_code\<close> command.  | 
16  | 
\<close>  | 
|
| 
37745
 
6315b6426200
checking generated code for various target languages
 
haftmann 
parents: 
37695 
diff
changeset
 | 
17  | 
|
| 
50629
 
264ece81df93
code checking for Scala is mandatory, since Scala is now required anyway for Isabelle
 
haftmann 
parents: 
40711 
diff
changeset
 | 
18  | 
export_code _ checking SML OCaml? Haskell? Scala  | 
| 
19281
 
b411f25fff25
added example for operational classes and code generator
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
|
| 23266 | 20  | 
end  |