| author | Christian Sternagel | 
| Thu, 30 Aug 2012 13:06:04 +0900 | |
| changeset 49088 | 5cd8b4426a57 | 
| parent 48891 | c0eafbd55de3 | 
| child 54489 | 03ff4d1e6784 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/Library/Code_Prolog.thy Author: Lukas Bulwahn, TUM 2010 *) header {* Code generation of prolog programs *} theory Code_Prolog imports Main begin ML_file "~~/src/HOL/Tools/Predicate_Compile/code_prolog.ML" section {* Setup for Numerals *} setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}, @{const_name neg_numeral}] *} setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}, @{const_name neg_numeral}] *} end