author | paulson <lp15@cam.ac.uk> |
Sat, 04 Dec 2021 20:30:16 +0000 | |
changeset 74878 | 0263787a06b4 |
parent 69605 | a96320074298 |
permissions | -rw-r--r-- |
38117 | 1 |
(* Title: HOL/Library/Code_Prolog.thy |
2 |
Author: Lukas Bulwahn, TUM 2010 |
|
3 |
*) |
|
4 |
||
60500 | 5 |
section \<open>Code generation of prolog programs\<close> |
38117 | 6 |
|
7 |
theory Code_Prolog |
|
8 |
imports Main |
|
55447 | 9 |
keywords "values_prolog" :: diag |
38117 | 10 |
begin |
11 |
||
69605 | 12 |
ML_file \<open>~~/src/HOL/Tools/Predicate_Compile/code_prolog.ML\<close> |
48891 | 13 |
|
60500 | 14 |
section \<open>Setup for Numerals\<close> |
38731
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents:
38117
diff
changeset
|
15 |
|
69593 | 16 |
setup \<open>Predicate_Compile_Data.ignore_consts [\<^const_name>\<open>numeral\<close>]\<close> |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
38731
diff
changeset
|
17 |
|
69593 | 18 |
setup \<open>Predicate_Compile_Data.keep_functions [\<^const_name>\<open>numeral\<close>]\<close> |
38731
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents:
38117
diff
changeset
|
19 |
|
38117 | 20 |
end |