--- a/src/HOL/Codegenerator_Test/Code_Test_PolyML.thy Fri May 11 22:59:00 2018 +0200
+++ b/src/HOL/Codegenerator_Test/Code_Test_PolyML.thy Sat May 12 11:24:11 2018 +0200
@@ -4,10 +4,12 @@
Test case for test_code on PolyML
*)
-theory Code_Test_PolyML imports "HOL-Library.Code_Test" begin
+theory Code_Test_PolyML imports "HOL-Library.Code_Test" Code_Lazy_Test begin
test_code "14 + 7 * -12 = (140 div -2 :: integer)" in PolyML
value [PolyML] "14 + 7 * -12 :: integer"
+test_code "stake 10 (cycle ''ab'') = ''ababababab''" in PolyML
+
end