src/HOL/Codegenerator_Test/Code_Test_OCaml.thy
author wenzelm
Fri, 18 Aug 2017 20:47:47 +0200
changeset 66453 cc19f7ca2ed6
parent 65552 f533820e7248
child 68155 8b50f29a1992
permissions -rw-r--r--
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59720
f893472fff31 proper headers;
wenzelm
parents: 58626
diff changeset
     1
(*  Title:      HOL/Codegenerator_Test/Code_Test_OCaml.thy
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     2
    Author:     Andreas Lochbihler, ETH Zurich
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     3
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     4
Test case for test_code on OCaml
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     5
*)
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     6
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 65552
diff changeset
     7
theory Code_Test_OCaml imports  "HOL-Library.Code_Test" begin
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     8
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
     9
test_code "14 + 7 * -12 = (140 div -2 :: integer)" in OCaml
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    10
58348
2d47c7d10b62 add target language evaluators for the value command;
Andreas Lochbihler
parents: 58039
diff changeset
    11
value [OCaml] "14 + 7 * -12 :: integer"
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    12
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61856
diff changeset
    13
test_code \<comment> \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
61856
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    14
  "gcd 15 10 = (5 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    15
  "gcd 15 (- 10) = (5 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    16
  "gcd (- 10) 15 = (5 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    17
  "gcd (- 10) (- 15) = (5 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    18
  "gcd 0 (- 10) = (10 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    19
  "gcd (- 10) 0 = (10 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    20
  "gcd 0 0 = (0 :: integer)"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    21
in OCaml
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 59720
diff changeset
    22
58039
469a375212c1 add testing framework for generated code
Andreas Lochbihler
parents:
diff changeset
    23
end