author | wenzelm |
Wed, 28 Dec 2022 12:30:18 +0100 | |
changeset 76798 | 69d8d16c5612 |
parent 74618 | 43142ac556e6 |
permissions | -rw-r--r-- |
31378 | 1 |
|
2 |
(* Author: Florian Haftmann, TU Muenchen *) |
|
24195 | 3 |
|
63167 | 4 |
section \<open>Pervasive test of code generator\<close> |
24195 | 5 |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51143
diff
changeset
|
6 |
theory Generate_Target_Nat |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51143
diff
changeset
|
7 |
imports |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51143
diff
changeset
|
8 |
Candidates |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
9 |
"HOL-Library.AList_Mapping" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
10 |
"HOL-Library.Finite_Lattice" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
11 |
"HOL-Library.Code_Target_Numeral" |
24195 | 12 |
begin |
13 |
||
63167 | 14 |
text \<open> |
37824 | 15 |
If any of the checks fails, inspect the code generated |
63167 | 16 |
by a corresponding \<open>export_code\<close> command. |
17 |
\<close> |
|
37745
6315b6426200
checking generated code for various target languages
haftmann
parents:
37695
diff
changeset
|
18 |
|
50629
264ece81df93
code checking for Scala is mandatory, since Scala is now required anyway for Isabelle
haftmann
parents:
47108
diff
changeset
|
19 |
export_code _ checking SML OCaml? Haskell? Scala |
25616 | 20 |
|
24195 | 21 |
end |