| author | wenzelm | 
| Mon, 16 Mar 2015 16:59:59 +0100 | |
| changeset 59720 | f893472fff31 | 
| parent 58626 | 6c473ed0ac70 | 
| child 61856 | 4b1b85f38944 | 
| permissions | -rw-r--r-- | 
| 59720 | 1  | 
(* Title: HOL/Codegenerator_Test/Code_Test_GHC.thy  | 
| 58039 | 2  | 
Author: Andreas Lochbihler, ETH Zurich  | 
3  | 
||
4  | 
Test case for test_code on GHC  | 
|
5  | 
*)  | 
|
6  | 
||
| 58626 | 7  | 
theory Code_Test_GHC imports "../Library/Code_Test" begin  | 
| 58039 | 8  | 
|
| 
58414
 
f945e7af0d27
drop workaround addressed by d0d3c30806b4
 
Andreas Lochbihler 
parents: 
58348 
diff
changeset
 | 
9  | 
test_code "(14 + 7 * -12 :: integer) = 140 div -2" in GHC  | 
| 58039 | 10  | 
|
| 
58348
 
2d47c7d10b62
add target language evaluators for the value command;
 
Andreas Lochbihler 
parents: 
58039 
diff
changeset
 | 
11  | 
value [GHC] "14 + 7 * -12 :: integer"  | 
| 58039 | 12  | 
|
13  | 
end  |