author | haftmann |
Fri Nov 01 18:51:14 2013 +0100 (2013-11-01) | |
changeset 54230 | b1d955791529 |
parent 48891 | c0eafbd55de3 |
child 54489 | 03ff4d1e6784 |
permissions | -rw-r--r-- |
bulwahn@38117 | 1 |
(* Title: HOL/Library/Code_Prolog.thy |
bulwahn@38117 | 2 |
Author: Lukas Bulwahn, TUM 2010 |
bulwahn@38117 | 3 |
*) |
bulwahn@38117 | 4 |
|
bulwahn@38117 | 5 |
header {* Code generation of prolog programs *} |
bulwahn@38117 | 6 |
|
bulwahn@38117 | 7 |
theory Code_Prolog |
bulwahn@38117 | 8 |
imports Main |
bulwahn@38117 | 9 |
begin |
bulwahn@38117 | 10 |
|
wenzelm@48891 | 11 |
ML_file "~~/src/HOL/Tools/Predicate_Compile/code_prolog.ML" |
wenzelm@48891 | 12 |
|
bulwahn@38731 | 13 |
section {* Setup for Numerals *} |
bulwahn@38731 | 14 |
|
huffman@47108 | 15 |
setup {* Predicate_Compile_Data.ignore_consts |
huffman@47108 | 16 |
[@{const_name numeral}, @{const_name neg_numeral}] *} |
huffman@47108 | 17 |
|
huffman@47108 | 18 |
setup {* Predicate_Compile_Data.keep_functions |
huffman@47108 | 19 |
[@{const_name numeral}, @{const_name neg_numeral}] *} |
bulwahn@38731 | 20 |
|
bulwahn@38117 | 21 |
end |