| author | hoelzl | 
| Tue, 04 Dec 2012 18:00:37 +0100 | |
| changeset 50347 | 77e3effa50b6 | 
| parent 47108 | 2a1953f0d20d | 
| child 50629 | 264ece81df93 | 
| permissions | -rw-r--r-- | 
| 31378 | 1 | |
| 2 | (* Author: Florian Haftmann, TU Muenchen *) | |
| 24195 | 3 | |
| 31378 | 4 | header {* Pervasive test of code generator using pretty literals *}
 | 
| 24195 | 5 | |
| 37695 
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
 haftmann parents: 
37475diff
changeset | 6 | theory Generate_Pretty | 
| 
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
 haftmann parents: 
37475diff
changeset | 7 | imports Candidates_Pretty | 
| 24195 | 8 | begin | 
| 9 | ||
| 37695 
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
 haftmann parents: 
37475diff
changeset | 10 | lemma [code, code del]: "nat_of_char = nat_of_char" .. | 
| 
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
 haftmann parents: 
37475diff
changeset | 11 | lemma [code, code del]: "char_of_nat = char_of_nat" .. | 
| 
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
 haftmann parents: 
37475diff
changeset | 12 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43317diff
changeset | 13 | declare Quickcheck_Narrowing.one_code_int_code [code del] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
43317diff
changeset | 14 | declare Quickcheck_Narrowing.int_of_code [code del] | 
| 43317 | 15 | |
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 16 | subsection {* Check whether generated code compiles *}
 | 
| 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 17 | |
| 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 18 | text {*
 | 
| 37824 | 19 | If any of the checks fails, inspect the code generated | 
| 20 |   by a corresponding @{text export_code} command.
 | |
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 21 | *} | 
| 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37695diff
changeset | 22 | |
| 40711 
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
 haftmann parents: 
37825diff
changeset | 23 | export_code _ checking SML OCaml? Haskell? Scala? | 
| 25616 | 24 | |
| 24195 | 25 | end |