| author | blanchet | 
| Tue, 09 Sep 2014 20:51:36 +0200 | |
| changeset 58278 | e89c7ac4ce16 | 
| parent 51161 | 6ed12ae3b3e1 | 
| child 58889 | 5b7a9633cfa8 | 
| 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 | |
| 31378 | 4 | header {* Pervasive test of code generator *}
 | 
| 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: 
34306diff
changeset | 6 | theory Generate | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50629diff
changeset | 7 | imports | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50629diff
changeset | 8 | Candidates | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50629diff
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: 
50629diff
changeset | 10 | "~~/src/HOL/Library/Finite_Lattice" | 
| 31378 | 11 | begin | 
| 24430 | 12 | |
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 13 | text {*
 | 
| 37824 | 14 | If any of the checks fails, inspect the code generated | 
| 15 |   by a corresponding @{text export_code} command.
 | |
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 16 | *} | 
| 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 17 | |
| 50629 
264ece81df93
code checking for Scala is mandatory, since Scala is now required anyway for Isabelle
 haftmann parents: 
40711diff
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 | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50629diff
changeset | 21 |