| author | blanchet | 
| Fri, 18 Mar 2011 12:20:32 +0100 | |
| changeset 41997 | 33b7d118e9dc | 
| parent 41952 | c7297638599b | 
| child 42091 | 6fe4abb9437b | 
| permissions | -rw-r--r-- | 
| 38073 | 1 | (* Title: HOL/Tools/Predicate_Compile/code_prolog.ML | 
| 2 | Author: Lukas Bulwahn, TU Muenchen | |
| 3 | ||
| 41941 | 4 | Prototype of an code generator for logic programming languages | 
| 5 | (a.k.a. Prolog). | |
| 38073 | 6 | *) | 
| 7 | ||
| 8 | signature CODE_PROLOG = | |
| 9 | sig | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 10 | datatype prolog_system = SWI_PROLOG | YAP | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 11 | type code_options = | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 12 |     {ensure_groundness : bool,
 | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 13 | limit_globally : int option, | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 14 | limited_types : (typ * int) list, | 
| 38959 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 15 | limited_predicates : (string list * int) list, | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 16 | replacing : ((string * string) * string) list, | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 17 | manual_reorder : ((string * int) * int list) list} | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 18 | val set_ensure_groundness : code_options -> code_options | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 19 | val map_limit_predicates : ((string list * int) list -> (string list * int) list) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 20 | -> code_options -> code_options | 
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 21 | val code_options_of : theory -> code_options | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 22 | val map_code_options : (code_options -> code_options) -> theory -> theory | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 23 | |
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 24 | datatype arith_op = Plus | Minus | 
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 25 | datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list | 
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 26 | | Number of int | ArithOp of arith_op * prol_term list; | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 27 | datatype prem = Conj of prem list | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 28 | | Rel of string * prol_term list | NotRel of string * prol_term list | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 29 | | Eq of prol_term * prol_term | NotEq of prol_term * prol_term | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 30 | | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 31 | | Ground of string * typ; | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 32 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 33 | type clause = ((string * prol_term list) * prem); | 
| 38073 | 34 | type logic_program = clause list; | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 35 | type constant_table = (string * string) list | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 36 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 37 | val generate : Predicate_Compile_Aux.mode option * bool -> | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 38 | Proof.context -> string -> (logic_program * constant_table) | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 39 | val write_program : logic_program -> string | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 40 | val run : (Time.time * prolog_system) -> logic_program -> (string * prol_term list) -> string list -> int option -> prol_term list list | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 41 | |
| 40924 | 42 | val quickcheck : Proof.context -> term -> int -> term list option * Quickcheck.report option | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 43 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 44 | val trace : bool Unsynchronized.ref | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 45 | |
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 46 | val replace : ((string * string) * string) -> logic_program -> logic_program | 
| 38073 | 47 | end; | 
| 48 | ||
| 49 | structure Code_Prolog : CODE_PROLOG = | |
| 50 | struct | |
| 51 | ||
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 52 | (* diagnostic tracing *) | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 53 | |
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 54 | val trace = Unsynchronized.ref false | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 55 | |
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 56 | fun tracing s = if !trace then Output.tracing s else () | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 57 | |
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 58 | (* code generation options *) | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 59 | |
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 60 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 61 | type code_options = | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 62 |   {ensure_groundness : bool,
 | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 63 | limit_globally : int option, | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 64 | limited_types : (typ * int) list, | 
| 38959 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 65 | limited_predicates : (string list * int) list, | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 66 | replacing : ((string * string) * string) list, | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 67 | manual_reorder : ((string * int) * int list) list} | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 68 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 69 | |
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 70 | fun set_ensure_groundness {ensure_groundness, limit_globally, limited_types, limited_predicates,
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 71 | replacing, manual_reorder} = | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 72 |   {ensure_groundness = true, limit_globally = limit_globally, limited_types = limited_types,
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 73 | limited_predicates = limited_predicates, replacing = replacing, | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 74 | manual_reorder = manual_reorder} | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 75 | |
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 76 | fun map_limit_predicates f {ensure_groundness, limit_globally, limited_types, limited_predicates,
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 77 | replacing, manual_reorder} = | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 78 |   {ensure_groundness = ensure_groundness, limit_globally = limit_globally, limited_types = limited_types,
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 79 | limited_predicates = f limited_predicates, replacing = replacing, | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 80 | manual_reorder = manual_reorder} | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 81 | |
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 82 | fun merge_global_limit (NONE, NONE) = NONE | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 83 | | merge_global_limit (NONE, SOME n) = SOME n | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 84 | | merge_global_limit (SOME n, NONE) = SOME n | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41307diff
changeset | 85 | | merge_global_limit (SOME n, SOME m) = SOME (Int.max (n, m)) (* FIXME odd merge *) | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 86 | |
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 87 | structure Options = Theory_Data | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 88 | ( | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 89 | type T = code_options | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 90 |   val empty = {ensure_groundness = false, limit_globally = NONE,
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 91 | limited_types = [], limited_predicates = [], replacing = [], manual_reorder = []} | 
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 92 | val extend = I; | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 93 | fun merge | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 94 |     ({ensure_groundness = ensure_groundness1, limit_globally = limit_globally1,
 | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 95 | limited_types = limited_types1, limited_predicates = limited_predicates1, | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 96 | replacing = replacing1, manual_reorder = manual_reorder1}, | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 97 |      {ensure_groundness = ensure_groundness2, limit_globally = limit_globally2,
 | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 98 | limited_types = limited_types2, limited_predicates = limited_predicates2, | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 99 | replacing = replacing2, manual_reorder = manual_reorder2}) = | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41307diff
changeset | 100 |     {ensure_groundness = ensure_groundness1 orelse ensure_groundness2 (* FIXME odd merge *),
 | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 101 | limit_globally = merge_global_limit (limit_globally1, limit_globally2), | 
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 102 | limited_types = AList.merge (op =) (K true) (limited_types1, limited_types2), | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 103 | limited_predicates = AList.merge (op =) (K true) (limited_predicates1, limited_predicates2), | 
| 38960 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 104 | manual_reorder = AList.merge (op =) (K true) (manual_reorder1, manual_reorder2), | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 105 | replacing = Library.merge (op =) (replacing1, replacing2)}; | 
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 106 | ); | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 107 | |
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 108 | val code_options_of = Options.get | 
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 109 | |
| 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 110 | val map_code_options = Options.map | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 111 | |
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 112 | (* system configuration *) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 113 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 114 | datatype prolog_system = SWI_PROLOG | YAP | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 115 | |
| 39462 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 116 | fun string_of_system SWI_PROLOG = "swiprolog" | 
| 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 117 | | string_of_system YAP = "yap" | 
| 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 118 | |
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 119 | type system_configuration = {timeout : Time.time, prolog_system : prolog_system}
 | 
| 39462 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 120 | |
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 121 | structure System_Config = Generic_Data | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 122 | ( | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 123 | type T = system_configuration | 
| 40301 | 124 |   val empty = {timeout = seconds 10.0, prolog_system = SWI_PROLOG}
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 125 | val extend = I; | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41307diff
changeset | 126 | fun merge (a, _) = a | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 127 | ) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 128 | |
| 38073 | 129 | (* general string functions *) | 
| 130 | ||
| 40924 | 131 | val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o raw_explode; | 
| 132 | val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o raw_explode; | |
| 38073 | 133 | |
| 134 | (* internal program representation *) | |
| 135 | ||
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 136 | datatype arith_op = Plus | Minus | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 137 | |
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 138 | datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list | 
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 139 | | Number of int | ArithOp of arith_op * prol_term list; | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 140 | |
| 38735 | 141 | fun dest_Var (Var v) = v | 
| 142 | ||
| 143 | fun add_vars (Var v) = insert (op =) v | |
| 144 | | add_vars (ArithOp (_, ts)) = fold add_vars ts | |
| 145 | | add_vars (AppF (_, ts)) = fold add_vars ts | |
| 146 | | add_vars _ = I | |
| 147 | ||
| 148 | fun map_vars f (Var v) = Var (f v) | |
| 149 | | map_vars f (ArithOp (opr, ts)) = ArithOp (opr, map (map_vars f) ts) | |
| 150 | | map_vars f (AppF (fs, ts)) = AppF (fs, map (map_vars f) ts) | |
| 151 | | map_vars f t = t | |
| 152 | ||
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 153 | fun maybe_AppF (c, []) = Cons c | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 154 | | maybe_AppF (c, xs) = AppF (c, xs) | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 155 | |
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 156 | fun is_Var (Var _) = true | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 157 | | is_Var _ = false | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 158 | |
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 159 | fun is_arith_term (Var _) = true | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 160 | | is_arith_term (Number _) = true | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 161 | | is_arith_term (ArithOp (_, operands)) = forall is_arith_term operands | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 162 | | is_arith_term _ = false | 
| 38073 | 163 | |
| 38081 
8b02ce312893
removing pointless type information in internal prolog terms
 bulwahn parents: 
38080diff
changeset | 164 | fun string_of_prol_term (Var s) = "Var " ^ s | 
| 38075 | 165 | | string_of_prol_term (Cons s) = "Cons " ^ s | 
| 166 |   | string_of_prol_term (AppF (f, args)) = f ^ "(" ^ commas (map string_of_prol_term args) ^ ")" 
 | |
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 167 | | string_of_prol_term (Number n) = "Number " ^ string_of_int n | 
| 38075 | 168 | |
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 169 | datatype prem = Conj of prem list | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 170 | | Rel of string * prol_term list | NotRel of string * prol_term list | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 171 | | Eq of prol_term * prol_term | NotEq of prol_term * prol_term | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 172 | | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 173 | | Ground of string * typ; | 
| 38735 | 174 | |
| 38073 | 175 | fun dest_Rel (Rel (c, ts)) = (c, ts) | 
| 38735 | 176 | |
| 177 | fun map_term_prem f (Conj prems) = Conj (map (map_term_prem f) prems) | |
| 178 | | map_term_prem f (Rel (r, ts)) = Rel (r, map f ts) | |
| 179 | | map_term_prem f (NotRel (r, ts)) = NotRel (r, map f ts) | |
| 180 | | map_term_prem f (Eq (l, r)) = Eq (f l, f r) | |
| 181 | | map_term_prem f (NotEq (l, r)) = NotEq (f l, f r) | |
| 182 | | map_term_prem f (ArithEq (l, r)) = ArithEq (f l, f r) | |
| 183 | | map_term_prem f (NotArithEq (l, r)) = NotArithEq (f l, f r) | |
| 184 | | map_term_prem f (Ground (v, T)) = Ground (dest_Var (f (Var v)), T) | |
| 185 | ||
| 186 | fun fold_prem_terms f (Conj prems) = fold (fold_prem_terms f) prems | |
| 187 | | fold_prem_terms f (Rel (_, ts)) = fold f ts | |
| 188 | | fold_prem_terms f (NotRel (_, ts)) = fold f ts | |
| 189 | | fold_prem_terms f (Eq (l, r)) = f l #> f r | |
| 190 | | fold_prem_terms f (NotEq (l, r)) = f l #> f r | |
| 191 | | fold_prem_terms f (ArithEq (l, r)) = f l #> f r | |
| 192 | | fold_prem_terms f (NotArithEq (l, r)) = f l #> f r | |
| 193 | | fold_prem_terms f (Ground (v, T)) = f (Var v) | |
| 194 | ||
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 195 | type clause = ((string * prol_term list) * prem); | 
| 38073 | 196 | |
| 197 | type logic_program = clause list; | |
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 198 | |
| 38073 | 199 | (* translation from introduction rules to internal representation *) | 
| 200 | ||
| 38958 
08eb0ffa2413
improving clash-free naming of variables and preds in code_prolog
 bulwahn parents: 
38956diff
changeset | 201 | fun mk_conform f empty avoid name = | 
| 38956 | 202 | let | 
| 203 | fun dest_Char (Symbol.Char c) = c | |
| 204 | val name' = space_implode "" (map (dest_Char o Symbol.decode) | |
| 205 | (filter (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s) | |
| 206 | (Symbol.explode name))) | |
| 38958 
08eb0ffa2413
improving clash-free naming of variables and preds in code_prolog
 bulwahn parents: 
38956diff
changeset | 207 | val name'' = f (if name' = "" then empty else name') | 
| 
08eb0ffa2413
improving clash-free naming of variables and preds in code_prolog
 bulwahn parents: 
38956diff
changeset | 208 | in (if member (op =) avoid name'' then Name.variant avoid name'' else name'') end | 
| 38956 | 209 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 210 | (** constant table **) | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 211 | |
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 212 | type constant_table = (string * string) list | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 213 | |
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 214 | fun declare_consts consts constant_table = | 
| 38956 | 215 | let | 
| 216 | fun update' c table = | |
| 217 | if AList.defined (op =) table c then table else | |
| 218 | let | |
| 38958 
08eb0ffa2413
improving clash-free naming of variables and preds in code_prolog
 bulwahn parents: 
38956diff
changeset | 219 | val c' = mk_conform first_lower "pred" (map snd table) (Long_Name.base_name c) | 
| 38956 | 220 | in | 
| 221 | AList.update (op =) (c, c') table | |
| 222 | end | |
| 223 | in | |
| 224 | fold update' consts constant_table | |
| 225 | end | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 226 | |
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 227 | fun translate_const constant_table c = | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 228 | case AList.lookup (op =) constant_table c of | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 229 | SOME c' => c' | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 230 |   | NONE => error ("No such constant: " ^ c)
 | 
| 38073 | 231 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 232 | fun inv_lookup _ [] _ = NONE | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 233 | | inv_lookup eq ((key, value)::xs) value' = | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 234 | if eq (value', value) then SOME key | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 235 | else inv_lookup eq xs value'; | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 236 | |
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 237 | fun restore_const constant_table c = | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 238 | case inv_lookup (op =) constant_table c of | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 239 | SOME c' => c' | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 240 |   | NONE => error ("No constant corresponding to "  ^ c)
 | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 241 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 242 | (** translation of terms, literals, premises, and clauses **) | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 243 | |
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 244 | fun translate_arith_const @{const_name "Groups.plus_class.plus"} = SOME Plus
 | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 245 |   | translate_arith_const @{const_name "Groups.minus_class.minus"} = SOME Minus
 | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 246 | | translate_arith_const _ = NONE | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 247 | |
| 38734 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 248 | fun mk_nat_term constant_table n = | 
| 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 249 | let | 
| 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 250 |     val zero = translate_const constant_table @{const_name "Groups.zero_class.zero"}
 | 
| 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 251 |     val Suc = translate_const constant_table @{const_name "Suc"}
 | 
| 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 252 | in funpow n (fn t => AppF (Suc, [t])) (Cons zero) end | 
| 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 253 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 254 | fun translate_term ctxt constant_table t = | 
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 255 | case try HOLogic.dest_number t of | 
| 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 256 |     SOME (@{typ "int"}, n) => Number n
 | 
| 38734 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 257 |   | SOME (@{typ "nat"}, n) => mk_nat_term constant_table n
 | 
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 258 | | NONE => | 
| 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 259 | (case strip_comb t of | 
| 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 260 | (Free (v, T), []) => Var v | 
| 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 261 | | (Const (c, _), []) => Cons (translate_const constant_table c) | 
| 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 262 | | (Const (c, _), args) => | 
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 263 | (case translate_arith_const c of | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 264 | SOME aop => ArithOp (aop, map (translate_term ctxt constant_table) args) | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 265 | | NONE => | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 266 | AppF (translate_const constant_table c, map (translate_term ctxt constant_table) args)) | 
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 267 |       | _ => error ("illegal term for translation: " ^ Syntax.string_of_term ctxt t))
 | 
| 38073 | 268 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 269 | fun translate_literal ctxt constant_table t = | 
| 38073 | 270 | case strip_comb t of | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38797diff
changeset | 271 |     (Const (@{const_name HOL.eq}, _), [l, r]) =>
 | 
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 272 | let | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 273 | val l' = translate_term ctxt constant_table l | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 274 | val r' = translate_term ctxt constant_table r | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 275 | in | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 276 | (if is_Var l' andalso is_arith_term r' andalso not (is_Var r') then ArithEq else Eq) (l', r') | 
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 277 | end | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 278 | | (Const (c, _), args) => | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 279 | Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args) | 
| 38073 | 280 |   | _ => error ("illegal literal for translation: " ^ Syntax.string_of_term ctxt t)
 | 
| 281 | ||
| 282 | fun NegRel_of (Rel lit) = NotRel lit | |
| 283 | | NegRel_of (Eq eq) = NotEq eq | |
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 284 | | NegRel_of (ArithEq eq) = NotArithEq eq | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 285 | |
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 286 | fun mk_groundness_prems t = map Ground (Term.add_frees t []) | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 287 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 288 | fun translate_prem ensure_groundness ctxt constant_table t = | 
| 38073 | 289 | case try HOLogic.dest_not t of | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 290 | SOME t => | 
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 291 | if ensure_groundness then | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 292 | Conj (mk_groundness_prems t @ [NegRel_of (translate_literal ctxt constant_table t)]) | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 293 | else | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 294 | NegRel_of (translate_literal ctxt constant_table t) | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 295 | | NONE => translate_literal ctxt constant_table t | 
| 38114 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 296 | |
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 297 | fun imp_prems_conv cv ct = | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 298 | case Thm.term_of ct of | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 299 |     Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
 | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 300 | | _ => Conv.all_conv ct | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 301 | |
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 302 | fun Trueprop_conv cv ct = | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 303 | case Thm.term_of ct of | 
| 38558 | 304 |     Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
 | 
| 38114 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 305 | | _ => raise Fail "Trueprop_conv" | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 306 | |
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 307 | fun preprocess_intro thy rule = | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 308 | Conv.fconv_rule | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 309 | (imp_prems_conv | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 310 |       (Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq}))))
 | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 311 | (Thm.transfer thy rule) | 
| 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 312 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 313 | fun translate_intros ensure_groundness ctxt gr const constant_table = | 
| 38073 | 314 | let | 
| 38114 
0f06e3cc04a6
adding preprocessing of introduction rules to replace the constant Predicate.eq in the prolog code generation
 bulwahn parents: 
38113diff
changeset | 315 | val intros = map (preprocess_intro (ProofContext.theory_of ctxt)) (Graph.get_node gr const) | 
| 38073 | 316 | val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 317 | val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table | 
| 38073 | 318 | fun translate_intro intro = | 
| 319 | let | |
| 320 | val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro) | |
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 321 | val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro) | 
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 322 | val prems' = Conj (map (translate_prem ensure_groundness ctxt' constant_table') prems) | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 323 | val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems') | 
| 38073 | 324 | in clause end | 
| 39724 | 325 | in | 
| 326 | (map translate_intro intros', constant_table') | |
| 327 | end | |
| 38073 | 328 | |
| 38731 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 329 | fun depending_preds_of (key, intros) = | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 330 | fold Term.add_const_names (map Thm.prop_of intros) [] | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 331 | |
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 332 | fun add_edges edges_of key G = | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 333 | let | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 334 | fun extend' key (G, visited) = | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 335 | case try (Graph.get_node G) key of | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 336 | SOME v => | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 337 | let | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 338 | val new_edges = filter (fn k => is_some (try (Graph.get_node G) k)) (edges_of (key, v)) | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 339 | val (G', visited') = fold extend' | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 340 | (subtract (op =) (key :: visited) new_edges) (G, key :: visited) | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 341 | in | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 342 | (fold (Graph.add_edge o (pair key)) new_edges G', visited') | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 343 | end | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 344 | | NONE => (G, visited) | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 345 | in | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 346 | fst (extend' key (G, [])) | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 347 | end | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 348 | |
| 39183 
512c10416590
adding dependencies to IsaMakefile; increasing negative search limit for predicate_compile_quickcheck; adding tracing of introduction rules in code_prolog
 bulwahn parents: 
38961diff
changeset | 349 | fun print_intros ctxt gr consts = | 
| 
512c10416590
adding dependencies to IsaMakefile; increasing negative search limit for predicate_compile_quickcheck; adding tracing of introduction rules in code_prolog
 bulwahn parents: 
38961diff
changeset | 350 | tracing (cat_lines (map (fn const => | 
| 
512c10416590
adding dependencies to IsaMakefile; increasing negative search limit for predicate_compile_quickcheck; adding tracing of introduction rules in code_prolog
 bulwahn parents: 
38961diff
changeset | 351 | "Constant " ^ const ^ "has intros:\n" ^ | 
| 
512c10416590
adding dependencies to IsaMakefile; increasing negative search limit for predicate_compile_quickcheck; adding tracing of introduction rules in code_prolog
 bulwahn parents: 
38961diff
changeset | 352 | cat_lines (map (Display.string_of_thm ctxt) (Graph.get_node gr const))) consts)) | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 353 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 354 | (* translation of moded predicates *) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 355 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 356 | (** generating graph of moded predicates **) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 357 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 358 | (* could be moved to Predicate_Compile_Core *) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 359 | fun requires_modes polarity cls = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 360 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 361 | fun req_mode_of pol (t, derivation) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 362 | (case fst (strip_comb t) of | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 363 | Const (c, _) => SOME (c, (pol, Predicate_Compile_Core.head_mode_of derivation)) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 364 | | _ => NONE) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 365 | fun req (Predicate_Compile_Aux.Prem t, derivation) = req_mode_of polarity (t, derivation) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 366 | | req (Predicate_Compile_Aux.Negprem t, derivation) = req_mode_of (not polarity) (t, derivation) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 367 | | req _ = NONE | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 368 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 369 | maps (fn (_, prems) => map_filter req prems) cls | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 370 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 371 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 372 | structure Mode_Graph = Graph(type key = string * (bool * Predicate_Compile_Aux.mode) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 373 | val ord = prod_ord fast_string_ord (prod_ord bool_ord Predicate_Compile_Aux.mode_ord)); | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 374 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 375 | fun mk_moded_clauses_graph ctxt scc gr = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 376 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 377 | val options = Predicate_Compile_Aux.default_options | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 378 | val mode_analysis_options = | 
| 39761 
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
 bulwahn parents: 
39724diff
changeset | 379 |       {use_generators = true, reorder_premises = true, infer_pos_and_neg_modes = true}
 | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 380 | fun infer prednames (gr, (pos_modes, neg_modes, random)) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 381 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 382 | val (lookup_modes, lookup_neg_modes, needs_random) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 383 | ((fn s => the (AList.lookup (op =) pos_modes s)), | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 384 | (fn s => the (AList.lookup (op =) neg_modes s)), | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 385 | (fn s => member (op =) (the (AList.lookup (op =) random s)))) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 386 | val (preds, all_vs, param_vs, all_modes, clauses) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 387 | Predicate_Compile_Core.prepare_intrs options ctxt prednames | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39798diff
changeset | 388 | (maps (Core_Data.intros_of ctxt) prednames) | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 389 | val ((moded_clauses, random'), _) = | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39798diff
changeset | 390 | Mode_Inference.infer_modes mode_analysis_options options | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 391 | (lookup_modes, lookup_neg_modes, needs_random) ctxt preds all_modes param_vs clauses | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 392 | val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 393 | val pos_modes' = map (apsnd (map_filter (fn (true, m) => SOME m | _ => NONE))) modes | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 394 | val neg_modes' = map (apsnd (map_filter (fn (false, m) => SOME m | _ => NONE))) modes | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 395 |         val _ = tracing ("Inferred modes:\n" ^
 | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 396 | cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 397 | (fn (p, m) => Predicate_Compile_Aux.string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes)) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 398 | val gr' = gr | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 399 | |> fold (fn (p, mps) => fold (fn (mode, cls) => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 400 | Mode_Graph.new_node ((p, mode), cls)) mps) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 401 | moded_clauses | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 402 | |> fold (fn (p, mps) => fold (fn (mode, cls) => fold (fn req => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 403 | Mode_Graph.add_edge ((p, mode), req)) (requires_modes (fst mode) cls)) mps) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 404 | moded_clauses | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 405 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 406 | (gr', (AList.merge (op =) (op =) (pos_modes, pos_modes'), | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 407 | AList.merge (op =) (op =) (neg_modes, neg_modes'), | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 408 | AList.merge (op =) (op =) (random, random'))) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 409 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 410 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 411 | fst (fold infer (rev scc) (Mode_Graph.empty, ([], [], []))) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 412 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 413 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 414 | fun declare_moded_predicate moded_preds table = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 415 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 416 | fun update' (p as (pred, (pol, mode))) table = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 417 | if AList.defined (op =) table p then table else | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 418 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 419 | val name = Long_Name.base_name pred ^ (if pol then "p" else "n") | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 420 | ^ Predicate_Compile_Aux.ascii_string_of_mode mode | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 421 | val p' = mk_conform first_lower "pred" (map snd table) name | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 422 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 423 | AList.update (op =) (p, p') table | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 424 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 425 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 426 | fold update' moded_preds table | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 427 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 428 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 429 | fun mk_program ctxt moded_gr moded_preds (prog, (moded_pred_table, constant_table)) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 430 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 431 | val moded_pred_table' = declare_moded_predicate moded_preds moded_pred_table | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 432 | fun mk_literal pol derivation constant_table' t = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 433 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 434 | val (p, args) = strip_comb t | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 435 | val mode = Predicate_Compile_Core.head_mode_of derivation | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 436 | val name = fst (dest_Const p) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 437 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 438 | val p' = the (AList.lookup (op =) moded_pred_table' (name, (pol, mode))) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 439 | val args' = map (translate_term ctxt constant_table') args | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 440 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 441 | Rel (p', args') | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 442 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 443 | fun mk_prem pol (indprem, derivation) constant_table = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 444 | case indprem of | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 445 | Predicate_Compile_Aux.Generator (s, T) => (Ground (s, T), constant_table) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 446 | | _ => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 447 | declare_consts (Term.add_const_names (Predicate_Compile_Aux.dest_indprem indprem) []) constant_table | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 448 | |> (fn constant_table' => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 449 | (case indprem of Predicate_Compile_Aux.Negprem t => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 450 | NegRel_of (mk_literal (not pol) derivation constant_table' t) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 451 | | _ => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 452 | mk_literal pol derivation constant_table' (Predicate_Compile_Aux.dest_indprem indprem), constant_table')) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 453 | fun mk_clause pred_name pol (ts, prems) (prog, constant_table) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 454 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 455 | val constant_table' = declare_consts (fold Term.add_const_names ts []) constant_table | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 456 | val args = map (translate_term ctxt constant_table') ts | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 457 | val (prems', constant_table'') = fold_map (mk_prem pol) prems constant_table' | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 458 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 459 | (((pred_name, args), Conj prems') :: prog, constant_table'') | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 460 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 461 | fun mk_clauses (pred, mode as (pol, _)) = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 462 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 463 | val clauses = Mode_Graph.get_node moded_gr (pred, mode) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 464 | val pred_name = the (AList.lookup (op =) moded_pred_table' (pred, mode)) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 465 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 466 | fold (mk_clause pred_name pol) clauses | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 467 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 468 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 469 | apsnd (pair moded_pred_table') (fold mk_clauses moded_preds (prog, constant_table)) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 470 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 471 | |
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 472 | fun generate (use_modes, ensure_groundness) ctxt const = | 
| 38073 | 473 | let | 
| 38731 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 474 | fun strong_conn_of gr keys = | 
| 38073 | 475 | Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr) | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39798diff
changeset | 476 | val gr = Core_Data.intros_graph_of ctxt | 
| 38731 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 477 | val gr' = add_edges depending_preds_of const gr | 
| 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38729diff
changeset | 478 | val scc = strong_conn_of gr' [const] | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 479 | val initial_constant_table = | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 480 |       declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}] []
 | 
| 38073 | 481 | in | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 482 | case use_modes of | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 483 | SOME mode => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 484 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 485 | val moded_gr = mk_moded_clauses_graph ctxt scc gr | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 486 | val moded_gr' = Mode_Graph.subgraph | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 487 | (member (op =) (Mode_Graph.all_succs moded_gr [(const, (true, mode))])) moded_gr | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 488 | val scc = Mode_Graph.strong_conn moded_gr' | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 489 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 490 | apfst rev (apsnd snd | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 491 | (fold (mk_program ctxt moded_gr') (rev scc) ([], ([], initial_constant_table)))) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 492 | end | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 493 | | NONE => | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 494 | let | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 495 | val _ = print_intros ctxt gr (flat scc) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 496 | val constant_table = declare_consts (flat scc) initial_constant_table | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 497 | in | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 498 | apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 499 | end | 
| 38073 | 500 | end | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 501 | |
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 502 | (* implementation for fully enumerating predicates and | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 503 | for size-limited predicates for enumerating the values of a datatype upto a specific size *) | 
| 38073 | 504 | |
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 505 | fun add_ground_typ (Conj prems) = fold add_ground_typ prems | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 506 | | add_ground_typ (Ground (_, T)) = insert (op =) T | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 507 | | add_ground_typ _ = I | 
| 38073 | 508 | |
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 509 | fun mk_relname (Type (Tcon, Targs)) = | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 510 | first_lower (Long_Name.base_name Tcon) ^ space_implode "_" (map mk_relname Targs) | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 511 | | mk_relname _ = raise Fail "unexpected type" | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 512 | |
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 513 | fun mk_lim_relname T = "lim_" ^ mk_relname T | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 514 | |
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 515 | (* This is copied from "pat_completeness.ML" *) | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 516 | fun inst_constrs_of thy (T as Type (name, _)) = | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 517 | map (fn (Cn,CT) => | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 518 | Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 519 | (the (Datatype.get_constrs thy name)) | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 520 |   | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
 | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 521 | |
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 522 | fun is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 523 | |
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 524 | fun mk_ground_impl ctxt limited_types (T as Type (Tcon, Targs)) (seen, constant_table) = | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 525 | if member (op =) seen T then ([], (seen, constant_table)) | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 526 | else | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 527 | let | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 528 | val (limited, size) = case AList.lookup (op =) limited_types T of | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 529 | SOME s => (true, s) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 530 | | NONE => (false, 0) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 531 | val rel_name = (if limited then mk_lim_relname else mk_relname) T | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 532 | fun mk_impl (Const (constr_name, cT), recursive) (seen, constant_table) = | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 533 | let | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 534 | val constant_table' = declare_consts [constr_name] constant_table | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 535 | val Ts = binder_types cT | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 536 | val (rec_clauses, (seen', constant_table'')) = | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 537 | fold_map (mk_ground_impl ctxt limited_types) Ts (seen, constant_table') | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 538 |           val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto (length Ts))
 | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 539 | val lim_var = | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 540 | if limited then | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 541 |               if recursive then [AppF ("suc", [Var "Lim"])]              
 | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 542 | else [Var "Lim"] | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 543 | else [] | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 544 | fun mk_prem v T' = | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 545 | if limited andalso T' = T then Rel (mk_lim_relname T', [Var "Lim", v]) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 546 | else Rel (mk_relname T', [v]) | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 547 | val clause = | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 548 | ((rel_name, lim_var @ [maybe_AppF (translate_const constant_table'' constr_name, vars)]), | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 549 | Conj (map2 mk_prem vars Ts)) | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 550 | in | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 551 | (clause :: flat rec_clauses, (seen', constant_table'')) | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 552 | end | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 553 | val constrs = inst_constrs_of (ProofContext.theory_of ctxt) T | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 554 | val constrs' = (constrs ~~ map (is_recursive_constr T) constrs) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 555 | |> (fn cs => filter_out snd cs @ filter snd cs) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 556 | val (clauses, constant_table') = | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 557 | apfst flat (fold_map mk_impl constrs' (T :: seen, constant_table)) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 558 |       val size_term = funpow size (fn t => AppF ("suc", [t])) (Cons "zero")
 | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 559 | in | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 560 | ((if limited then | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 561 | cons ((mk_relname T, [Var "x"]), Rel (mk_lim_relname T, [size_term, Var "x"])) | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 562 | else I) clauses, constant_table') | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 563 | end | 
| 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 564 | | mk_ground_impl ctxt _ T (seen, constant_table) = | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 565 |    raise Fail ("unexpected type :" ^ Syntax.string_of_typ ctxt T)
 | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 566 | |
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 567 | fun replace_ground (Conj prems) = Conj (map replace_ground prems) | 
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 568 | | replace_ground (Ground (x, T)) = | 
| 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 569 | Rel (mk_relname T, [Var x]) | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 570 | | replace_ground p = p | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 571 | |
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 572 | fun add_ground_predicates ctxt limited_types (p, constant_table) = | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 573 | let | 
| 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 574 | val ground_typs = fold (add_ground_typ o snd) p [] | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 575 | val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt limited_types) ground_typs ([], constant_table) | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 576 | val p' = map (apsnd replace_ground) p | 
| 38073 | 577 | in | 
| 38727 
c7f5f0b7dc7f
adding very basic transformation to ensure groundness before negations
 bulwahn parents: 
38558diff
changeset | 578 | ((flat grs) @ p', constant_table') | 
| 38073 | 579 | end | 
| 38789 
d171840881fd
added generation of predicates for size-limited enumeration of values
 bulwahn parents: 
38735diff
changeset | 580 | |
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 581 | (* make depth-limited version of predicate *) | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 582 | |
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 583 | fun mk_lim_rel_name rel_name = "lim_" ^ rel_name | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 584 | |
| 38959 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 585 | fun mk_depth_limited rel_names ((rel_name, ts), prem) = | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 586 | let | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 587 | fun has_positive_recursive_prems (Conj prems) = exists has_positive_recursive_prems prems | 
| 38959 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 588 | | has_positive_recursive_prems (Rel (rel, ts)) = member (op =) rel_names rel | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 589 | | has_positive_recursive_prems _ = false | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 590 | fun mk_lim_prem (Conj prems) = Conj (map mk_lim_prem prems) | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 591 | | mk_lim_prem (p as Rel (rel, ts)) = | 
| 38959 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 592 | if member (op =) rel_names rel then Rel (mk_lim_rel_name rel, Var "Lim" :: ts) else p | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 593 | | mk_lim_prem p = p | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 594 | in | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 595 | if has_positive_recursive_prems prem then | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 596 |       ((mk_lim_rel_name rel_name, (AppF ("suc", [Var "Lim"]))  :: ts), mk_lim_prem prem)
 | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 597 | else | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 598 | ((mk_lim_rel_name rel_name, (Var "Lim") :: ts), prem) | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 599 | end | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 600 | |
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 601 | fun nat_term_of n = funpow n (fn t => AppF ("suc", [t])) (Cons "zero")
 | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 602 | |
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 603 | fun add_limited_predicates limited_predicates (p, constant_table) = | 
| 38956 | 604 | let | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 605 | fun add (rel_names, limit) p = | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 606 | let | 
| 38959 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 607 | val clauses = filter (fn ((rel, _), _) => member (op =) rel_names rel) p | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 608 | val clauses' = map (mk_depth_limited rel_names) clauses | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 609 | fun mk_entry_clause rel_name = | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 610 | let | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 611 | val nargs = length (snd (fst | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 612 | (the (find_first (fn ((rel, _), _) => rel = rel_name) clauses)))) | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 613 |             val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto nargs)        
 | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 614 | in | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 615 |             (("limited_" ^ rel_name, vars), Rel ("lim_" ^ rel_name, nat_term_of limit :: vars))
 | 
| 
706ab66e3464
towards support of limited predicates for mutually recursive predicates
 bulwahn parents: 
38958diff
changeset | 616 | end | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 617 | in (p @ (map mk_entry_clause rel_names) @ clauses') end | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 618 | in | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 619 | (fold add limited_predicates p, constant_table) | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 620 | end | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 621 | |
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 622 | |
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 623 | (* replace predicates in clauses *) | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 624 | |
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 625 | (* replace (A, B, C) p = replace A by B in clauses of C *) | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 626 | fun replace ((from, to), location) p = | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 627 | let | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 628 | fun replace_prem (Conj prems) = Conj (map replace_prem prems) | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 629 | | replace_prem (r as Rel (rel, ts)) = | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 630 | if rel = from then Rel (to, ts) else r | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 631 | | replace_prem r = r | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 632 | in | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 633 | map (fn ((rel, args), prem) => ((rel, args), (if rel = location then replace_prem else I) prem)) p | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 634 | end | 
| 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 635 | |
| 38960 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 636 | |
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 637 | (* reorder manually : reorder premises of ith clause of predicate p by a permutation perm *) | 
| 38947 
6ed1cffd9d4e
added quite adhoc logic program transformations limited_predicates and replacements of predicates
 bulwahn parents: 
38864diff
changeset | 638 | |
| 38960 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 639 | fun reorder_manually reorder p = | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 640 | let | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 641 | fun reorder' (clause as ((rel, args), prem)) seen = | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 642 | let | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 643 | val seen' = AList.map_default (op =) (rel, 0) (fn x => x + 1) seen | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 644 | val i = the (AList.lookup (op =) seen' rel) | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 645 | val perm = AList.lookup (op =) reorder (rel, i) | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 646 | val prem' = (case perm of | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 647 | SOME p => (case prem of Conj prems => Conj (map (nth prems) p) | _ => prem) | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 648 | | NONE => prem) | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 649 | in (((rel, args), prem'), seen') end | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 650 | in | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 651 | fst (fold_map reorder' p []) | 
| 
363bfb245917
adding manual reordering of premises to prolog generation
 bulwahn parents: 
38959diff
changeset | 652 | end | 
| 39462 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 653 | |
| 38735 | 654 | (* rename variables to prolog-friendly names *) | 
| 655 | ||
| 656 | fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v)) | |
| 657 | ||
| 658 | fun rename_vars_prem renaming = map_term_prem (rename_vars_term renaming) | |
| 659 | ||
| 660 | fun is_prolog_conform v = | |
| 661 | forall (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s) (Symbol.explode v) | |
| 662 | ||
| 663 | fun mk_renaming v renaming = | |
| 38958 
08eb0ffa2413
improving clash-free naming of variables and preds in code_prolog
 bulwahn parents: 
38956diff
changeset | 664 | (v, mk_conform first_upper "Var" (map snd renaming) v) :: renaming | 
| 38735 | 665 | |
| 666 | fun rename_vars_clause ((rel, args), prem) = | |
| 667 | let | |
| 668 | val vars = fold_prem_terms add_vars prem (fold add_vars args []) | |
| 669 | val renaming = fold mk_renaming vars [] | |
| 670 | in ((rel, map (rename_vars_term renaming) args), rename_vars_prem renaming prem) end | |
| 671 | ||
| 672 | val rename_vars_program = map rename_vars_clause | |
| 38956 | 673 | |
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 674 | (* limit computation globally by some threshold *) | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 675 | |
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 676 | fun limit_globally ctxt limit const_name (p, constant_table) = | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 677 | let | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 678 | val rel_names = fold (fn ((r, _), _) => insert (op =) r) p [] | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 679 | val p' = map (mk_depth_limited rel_names) p | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 680 | val rel_name = translate_const constant_table const_name | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 681 | val nargs = length (snd (fst | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 682 | (the (find_first (fn ((rel, _), _) => rel = rel_name) p)))) | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 683 |     val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto nargs)
 | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 684 |     val entry_clause = ((rel_name, vars), Rel ("lim_" ^ rel_name, nat_term_of limit :: vars))
 | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 685 | val p'' = filter_out (fn ((rel, _), _) => rel = rel_name) p | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 686 | in | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 687 | (entry_clause :: p' @ p'', constant_table) | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 688 | end | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 689 | |
| 39542 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 690 | (* post processing of generated prolog program *) | 
| 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 691 | |
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 692 | fun post_process ctxt options const_name (p, constant_table) = | 
| 39542 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 693 | (p, constant_table) | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 694 | |> (case #limit_globally options of | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 695 | SOME limit => limit_globally ctxt limit const_name | 
| 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 696 | | NONE => I) | 
| 39542 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 697 | |> (if #ensure_groundness options then | 
| 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 698 | add_ground_predicates ctxt (#limited_types options) | 
| 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 699 | else I) | 
| 39724 | 700 | |> tap (fn _ => tracing "Adding limited predicates...") | 
| 39542 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 701 | |> add_limited_predicates (#limited_predicates options) | 
| 39724 | 702 | |> tap (fn _ => tracing "Replacing predicates...") | 
| 39542 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 703 | |> apfst (fold replace (#replacing options)) | 
| 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 704 | |> apfst (reorder_manually (#manual_reorder options)) | 
| 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 705 | |> apfst rename_vars_program | 
| 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 706 | |
| 38073 | 707 | (* code printer *) | 
| 708 | ||
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 709 | fun write_arith_op Plus = "+" | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 710 | | write_arith_op Minus = "-" | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 711 | |
| 38735 | 712 | fun write_term (Var v) = v | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 713 | | write_term (Cons c) = c | 
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 714 |   | write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
 | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 715 | | write_term (ArithOp (oper, [a1, a2])) = write_term a1 ^ " " ^ write_arith_op oper ^ " " ^ write_term a2 | 
| 38112 
cf08f4780938
adding numbers as basic term in prolog code generation
 bulwahn parents: 
38082diff
changeset | 716 | | write_term (Number n) = string_of_int n | 
| 38073 | 717 | |
| 718 | fun write_rel (pred, args) = | |
| 719 |   pred ^ "(" ^ space_implode ", " (map write_term args) ^ ")" 
 | |
| 720 | ||
| 721 | fun write_prem (Conj prems) = space_implode ", " (map write_prem prems) | |
| 722 | | write_prem (Rel p) = write_rel p | |
| 723 |   | write_prem (NotRel p) = "not(" ^ write_rel p ^ ")"
 | |
| 724 | | write_prem (Eq (l, r)) = write_term l ^ " = " ^ write_term r | |
| 725 | | write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r | |
| 38113 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 726 | | write_prem (ArithEq (l, r)) = write_term l ^ " is " ^ write_term r | 
| 
81f08bbb3be7
adding basic arithmetic support for prolog code generation
 bulwahn parents: 
38112diff
changeset | 727 | | write_prem (NotArithEq (l, r)) = write_term l ^ " =\\= " ^ write_term r | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 728 | | write_prem _ = raise Fail "Not a valid prolog premise" | 
| 38073 | 729 | |
| 730 | fun write_clause (head, prem) = | |
| 731 | write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".") | |
| 732 | ||
| 733 | fun write_program p = | |
| 734 | cat_lines (map write_clause p) | |
| 735 | ||
| 38790 | 736 | (* query templates *) | 
| 38078 | 737 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 738 | (** query and prelude for swi-prolog **) | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 739 | |
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 740 | fun swi_prolog_query_first (rel, args) vnames = | 
| 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 741 |   "eval :- once("  ^ rel ^ "(" ^ space_implode ", " (map write_term args) ^ ")),\n" ^
 | 
| 38082 | 742 |   "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
 | 
| 743 | "\\n', [" ^ space_implode ", " vnames ^ "]).\n" | |
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 744 | |
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 745 | fun swi_prolog_query_firstn n (rel, args) vnames = | 
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 746 |   "eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
 | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 747 |     rel ^ "(" ^ space_implode ", " (map write_term args) ^ "), Sols), writelist(Sols).\n" ^
 | 
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 748 | "writelist([]).\n" ^ | 
| 39546 
bfe10da7d764
renaming variable name to decrease likelyhood of nameclash
 bulwahn parents: 
39542diff
changeset | 749 |     "writelist([(" ^ space_implode ", " vnames ^ ")|SolutionTail]) :- " ^
 | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 750 |     "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
 | 
| 39546 
bfe10da7d764
renaming variable name to decrease likelyhood of nameclash
 bulwahn parents: 
39542diff
changeset | 751 | "\\n', [" ^ space_implode ", " vnames ^ "]), writelist(SolutionTail).\n" | 
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 752 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 753 | val swi_prolog_prelude = | 
| 41952 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 wenzelm parents: 
41941diff
changeset | 754 | "#!/usr/bin/swipl -q -t main -f\n\n" ^ (* FIXME hardwired executable!? *) | 
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 755 |   ":- use_module(library('dialect/ciao/aggregates')).\n" ^
 | 
| 38729 
9c9d14827380
improving output of set comprehensions; adding style_check flags
 bulwahn parents: 
38728diff
changeset | 756 | ":- style_check(-singleton).\n" ^ | 
| 41067 | 757 | ":- style_check(-discontiguous).\n" ^ | 
| 38729 
9c9d14827380
improving output of set comprehensions; adding style_check flags
 bulwahn parents: 
38728diff
changeset | 758 | ":- style_check(-atom).\n\n" ^ | 
| 38073 | 759 | "main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^ | 
| 760 | "main :- halt(1).\n" | |
| 38075 | 761 | |
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 762 | (** query and prelude for yap **) | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 763 | |
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 764 | fun yap_query_first (rel, args) vnames = | 
| 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 765 |   "eval :- once(" ^ rel ^ "(" ^ space_implode ", " (map write_term args) ^ ")),\n" ^
 | 
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 766 |   "format('" ^ space_implode ";" (map (fn v => v ^ " = ~w") vnames) ^
 | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 767 | "\\n', [" ^ space_implode ", " vnames ^ "]).\n" | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 768 | |
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 769 | val yap_prelude = | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 770 | "#!/usr/bin/yap -L\n\n" ^ | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 771 | ":- initialization(eval).\n" | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 772 | |
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 773 | (* system-dependent query, prelude and invocation *) | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 774 | |
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 775 | fun query system nsols = | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 776 | case system of | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 777 | SWI_PROLOG => | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 778 | (case nsols of NONE => swi_prolog_query_first | SOME n => swi_prolog_query_firstn n) | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 779 | | YAP => | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 780 | case nsols of NONE => yap_query_first | SOME n => | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 781 | error "No support for querying multiple solutions in the prolog system yap" | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 782 | |
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 783 | fun prelude system = | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 784 | case system of SWI_PROLOG => swi_prolog_prelude | YAP => yap_prelude | 
| 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 785 | |
| 41940 | 786 | fun invoke system file = | 
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 787 | let | 
| 41940 | 788 | val (env_var, cmd) = | 
| 789 | (case system of | |
| 41952 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 wenzelm parents: 
41941diff
changeset | 790 |         SWI_PROLOG => ("ISABELLE_SWIPL", "\"$ISABELLE_SWIPL\" -f ")
 | 
| 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 wenzelm parents: 
41941diff
changeset | 791 |       | YAP => ("ISABELLE_YAP", "\"$ISABELLE_YAP\" -L "))
 | 
| 39462 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 792 | in | 
| 41940 | 793 | if getenv env_var = "" then | 
| 39462 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 794 | (warning (env_var ^ " not set; could not execute code for " ^ string_of_system system); "") | 
| 41940 | 795 | else fst (bash_output (cmd ^ File.shell_path file)) | 
| 39462 
3a86194d1534
registering code_prolog as component; using environment variable; adding settings file for prolog code generation
 bulwahn parents: 
39461diff
changeset | 796 | end | 
| 38792 
970508a5119f
added support for yet another prolog system (yap); generate has only one option ensure_groundness; added one example of yap invocation in example theory
 bulwahn parents: 
38790diff
changeset | 797 | |
| 41952 
c7297638599b
cleanup of former settings GHC_PATH, EXEC_GHC, EXEC_OCAML, EXEC_SWIPL, EXEC_YAP -- discontinued implicit detection;
 wenzelm parents: 
41941diff
changeset | 798 | |
| 38075 | 799 | (* parsing prolog solution *) | 
| 38790 | 800 | |
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 801 | val scan_number = | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 802 | Scan.many1 Symbol.is_ascii_digit | 
| 38075 | 803 | |
| 804 | val scan_atom = | |
| 38728 
182b180e9804
improving ensure_groundness in prolog generation; added further example
 bulwahn parents: 
38727diff
changeset | 805 | Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s) | 
| 38075 | 806 | |
| 807 | val scan_var = | |
| 38078 | 808 | Scan.many1 | 
| 809 | (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s) | |
| 38075 | 810 | |
| 38076 | 811 | val scan_ident = | 
| 812 | Scan.repeat (Scan.one | |
| 813 | (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)) | |
| 814 | ||
| 38075 | 815 | fun dest_Char (Symbol.Char s) = s | 
| 816 | ||
| 817 | val string_of = concat o map (dest_Char o Symbol.decode) | |
| 818 | ||
| 38076 | 819 | val is_atom_ident = forall Symbol.is_ascii_lower | 
| 820 | ||
| 821 | val is_var_ident = | |
| 822 | forall (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s) | |
| 38078 | 823 | |
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 824 | fun int_of_symbol_list xs = fold (fn x => fn s => s * 10 + (ord x - ord "0")) xs 0 | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 825 | |
| 38078 | 826 | fun scan_terms xs = (((scan_term --| $$ ",") ::: scan_terms) | 
| 827 | || (scan_term >> single)) xs | |
| 828 | and scan_term xs = | |
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 829 | ((scan_number >> (Number o int_of_symbol_list)) | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 830 | || (scan_var >> (Var o string_of)) | 
| 38078 | 831 |   || ((scan_atom -- ($$ "(" |-- scan_terms --| $$ ")"))
 | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 832 | >> (fn (f, ts) => AppF (string_of f, ts))) | 
| 38078 | 833 | || (scan_atom >> (Cons o string_of))) xs | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 834 | |
| 38075 | 835 | val parse_term = fst o Scan.finite Symbol.stopper | 
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 836 | (Scan.error (!! (fn _ => raise Fail "parsing prolog output failed")) scan_term) | 
| 40924 | 837 | o raw_explode | 
| 38075 | 838 | |
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 839 | fun parse_solutions sol = | 
| 38075 | 840 | let | 
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 841 | fun dest_eq s = case space_explode "=" s of | 
| 38075 | 842 | (l :: r :: []) => parse_term (unprefix " " r) | 
| 38078 | 843 | | _ => raise Fail "unexpected equation in prolog output" | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 844 | fun parse_solution s = map dest_eq (space_explode ";" s) | 
| 38961 
8c2f59171647
handling the quickcheck result no counterexample more correctly
 bulwahn parents: 
38960diff
changeset | 845 | val sols = case space_explode "\n" sol of [] => [] | s => fst (split_last s) | 
| 38075 | 846 | in | 
| 38961 
8c2f59171647
handling the quickcheck result no counterexample more correctly
 bulwahn parents: 
38960diff
changeset | 847 | map parse_solution sols | 
| 38075 | 848 | end | 
| 38073 | 849 | |
| 850 | (* calling external interpreter and getting results *) | |
| 851 | ||
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 852 | fun run (timeout, system) p (query_rel, args) vnames nsols = | 
| 38073 | 853 | let | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 854 | val renaming = fold mk_renaming (fold add_vars args vnames) [] | 
| 38735 | 855 | val vnames' = map (fn v => the (AList.lookup (op =) renaming v)) vnames | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 856 | val args' = map (rename_vars_term renaming) args | 
| 39542 
a50c0ae2312c
moving renaming_vars to post_processing; removing clone in values and quickcheck of code_prolog
 bulwahn parents: 
39541diff
changeset | 857 | val prog = prelude system ^ query system nsols (query_rel, args') vnames' ^ write_program p | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 858 |     val _ = tracing ("Generated prolog program:\n" ^ prog)
 | 
| 41307 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 wenzelm parents: 
41067diff
changeset | 859 | val solution = TimeLimit.timeLimit timeout (fn prog => | 
| 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 wenzelm parents: 
41067diff
changeset | 860 | Isabelle_System.with_tmp_file "prolog_file" (fn prolog_file => | 
| 41940 | 861 | (File.write prolog_file prog; invoke system prolog_file))) prog | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 862 |     val _ = tracing ("Prolog returned solution(s):\n" ^ solution)
 | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 863 | val tss = parse_solutions solution | 
| 38073 | 864 | in | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 865 | tss | 
| 38073 | 866 | end | 
| 867 | ||
| 38790 | 868 | (* restoring types in terms *) | 
| 38075 | 869 | |
| 38081 
8b02ce312893
removing pointless type information in internal prolog terms
 bulwahn parents: 
38080diff
changeset | 870 | fun restore_term ctxt constant_table (Var s, T) = Free (s, T) | 
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 871 |   | restore_term ctxt constant_table (Number n, @{typ "int"}) = HOLogic.mk_number @{typ "int"} n
 | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 872 | | restore_term ctxt constant_table (Number n, _) = raise (Fail "unexpected type for number") | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 873 | | restore_term ctxt constant_table (Cons s, T) = Const (restore_const constant_table s, T) | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 874 | | restore_term ctxt constant_table (AppF (f, args), T) = | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 875 | let | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 876 | val thy = ProofContext.theory_of ctxt | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 877 | val c = restore_const constant_table f | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 878 | val cT = Sign.the_const_type thy c | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 879 | val (argsT, resT) = strip_type cT | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 880 | val subst = Sign.typ_match thy (resT, T) Vartab.empty | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 881 | val argsT' = map (Envir.subst_type subst) argsT | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 882 | in | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 883 | list_comb (Const (c, Envir.subst_type subst cT), | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 884 | map (restore_term ctxt constant_table) (args ~~ argsT')) | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 885 | end | 
| 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 886 | |
| 39465 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 887 | |
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 888 | (* restore numerals in natural numbers *) | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 889 | |
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 890 | fun restore_nat_numerals t = | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 891 |   if fastype_of t = @{typ nat} andalso is_some (try HOLogic.dest_nat t) then
 | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 892 |     HOLogic.mk_number @{typ nat} (HOLogic.dest_nat t)
 | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 893 | else | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 894 | (case t of | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 895 | t1 $ t2 => restore_nat_numerals t1 $ restore_nat_numerals t2 | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 896 | | t => t) | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 897 | |
| 38790 | 898 | (* values command *) | 
| 899 | ||
| 900 | val preprocess_options = Predicate_Compile_Aux.Options {
 | |
| 901 | expected_modes = NONE, | |
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39187diff
changeset | 902 | proposed_modes = [], | 
| 38790 | 903 | proposed_names = [], | 
| 904 | show_steps = false, | |
| 905 | show_intermediate_results = false, | |
| 906 | show_proof_trace = false, | |
| 907 | show_modes = false, | |
| 908 | show_mode_inference = false, | |
| 909 | show_compilation = false, | |
| 910 | show_caught_failures = false, | |
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39187diff
changeset | 911 | show_invalid_clauses = false, | 
| 38790 | 912 | skip_proof = true, | 
| 913 | no_topmost_reordering = false, | |
| 914 | function_flattening = true, | |
| 915 | specialise = false, | |
| 916 | fail_safe_function_flattening = false, | |
| 917 | no_higher_order_predicate = [], | |
| 918 | inductify = false, | |
| 919 | detect_switches = true, | |
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39798diff
changeset | 920 | smart_depth_limiting = true, | 
| 38790 | 921 | compilation = Predicate_Compile_Aux.Pred | 
| 922 | } | |
| 923 | ||
| 38075 | 924 | fun values ctxt soln t_compr = | 
| 925 | let | |
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 926 | val options = code_options_of (ProofContext.theory_of ctxt) | 
| 38075 | 927 |     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
 | 
| 928 |       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
 | |
| 929 | val (body, Ts, fp) = HOLogic.strip_psplits split; | |
| 930 | val output_names = Name.variant_list (Term.add_free_names body []) | |
| 931 | (map (fn i => "x" ^ string_of_int i) (1 upto length Ts)) | |
| 38080 
8c20eb9a388d
cleaning example file; more natural ordering of variable names
 bulwahn parents: 
38079diff
changeset | 932 | val output_frees = rev (map2 (curry Free) output_names Ts) | 
| 38075 | 933 | val body = subst_bounds (output_frees, body) | 
| 934 | val (pred as Const (name, T), all_args) = | |
| 935 | case strip_comb body of | |
| 936 | (Const (name, T), all_args) => (Const (name, T), all_args) | |
| 937 |       | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
 | |
| 38732 
3371dbc806ae
moving preprocessing to values in prolog generation
 bulwahn parents: 
38731diff
changeset | 938 | val _ = tracing "Preprocessing specification..." | 
| 
3371dbc806ae
moving preprocessing to values in prolog generation
 bulwahn parents: 
38731diff
changeset | 939 | val T = Sign.the_const_type (ProofContext.theory_of ctxt) name | 
| 
3371dbc806ae
moving preprocessing to values in prolog generation
 bulwahn parents: 
38731diff
changeset | 940 | val t = Const (name, T) | 
| 38755 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 941 | val thy' = | 
| 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 942 | Theory.copy (ProofContext.theory_of ctxt) | 
| 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 943 | |> Predicate_Compile.preprocess preprocess_options t | 
| 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 944 | val ctxt' = ProofContext.init_global thy' | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 945 | val _ = tracing "Generating prolog program..." | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 946 | val (p, constant_table) = generate (NONE, #ensure_groundness options) ctxt' name (* FIXME *) | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 947 | |> post_process ctxt' options name | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 948 | val constant_table' = declare_consts (fold Term.add_const_names all_args []) constant_table | 
| 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 949 | val args' = map (translate_term ctxt constant_table') all_args | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 950 | val _ = tracing "Running prolog program..." | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 951 | val system_config = System_Config.get (Context.Proof ctxt) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 952 | val tss = run (#timeout system_config, #prolog_system system_config) | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 953 | p (translate_const constant_table' name, args') output_names soln | 
| 38079 
7fb011dd51de
improving translation to prolog; restoring terms from prolog output; adding tracing support
 bulwahn parents: 
38078diff
changeset | 954 | val _ = tracing "Restoring terms..." | 
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 955 |     val empty = Const("Orderings.bot_class.bot", fastype_of t_compr)
 | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 956 | fun mk_insert x S = | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 957 |       Const (@{const_name "Set.insert"}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S 
 | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 958 | fun mk_set_compr in_insert [] xs = | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 959 |        rev ((Free ("...", fastype_of t_compr)) ::
 | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 960 | (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)) | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 961 | | mk_set_compr in_insert (t :: ts) xs = | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 962 | let | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 963 | val frees = Term.add_frees t [] | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 964 | in | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 965 | if null frees then | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 966 | mk_set_compr (t :: in_insert) ts xs | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 967 | else | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 968 | let | 
| 38755 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 969 |               val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt' [t]) ("uu", fastype_of t)
 | 
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 970 | val set_compr = | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 971 | HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t)) | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 972 |                   frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"})))
 | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 973 | in | 
| 38729 
9c9d14827380
improving output of set comprehensions; adding style_check flags
 bulwahn parents: 
38728diff
changeset | 974 | mk_set_compr [] ts | 
| 
9c9d14827380
improving output of set comprehensions; adding style_check flags
 bulwahn parents: 
38728diff
changeset | 975 | (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)) | 
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 976 | end | 
| 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 977 | end | 
| 38075 | 978 | in | 
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 979 |       foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr []
 | 
| 39465 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 980 | (map (fn ts => HOLogic.mk_tuple | 
| 
fcff6903595f
adding restoring of numerals for natural numbers for values command
 bulwahn parents: 
39464diff
changeset | 981 | (map (restore_nat_numerals o restore_term ctxt' constant_table) (ts ~~ Ts))) tss) []) | 
| 38075 | 982 | end | 
| 983 | ||
| 984 | fun values_cmd print_modes soln raw_t state = | |
| 985 | let | |
| 986 | val ctxt = Toplevel.context_of state | |
| 987 | val t = Syntax.read_term ctxt raw_t | |
| 988 | val t' = values ctxt soln t | |
| 989 | val ty' = Term.type_of t' | |
| 990 | val ctxt' = Variable.auto_fixes t' ctxt | |
| 38115 
987edb27f449
adding parsing of numbers; improving output of solution without free variables in prolog code generation
 bulwahn parents: 
38114diff
changeset | 991 | val _ = tracing "Printing terms..." | 
| 38075 | 992 | val p = Print_Mode.with_modes print_modes (fn () => | 
| 993 | Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, | |
| 994 | Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); | |
| 995 | in Pretty.writeln p end; | |
| 996 | ||
| 997 | ||
| 998 | (* renewing the values command for Prolog queries *) | |
| 999 | ||
| 1000 | val opt_print_modes = | |
| 1001 |   Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
 | |
| 1002 | ||
| 1003 | val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag | |
| 38077 
46ff55ace18c
querying for multiple solutions in values command for prolog execution
 bulwahn parents: 
38076diff
changeset | 1004 | (opt_print_modes -- Scan.optional (Parse.nat >> SOME) NONE -- Parse.term | 
| 38075 | 1005 | >> (fn ((print_modes, soln), t) => Toplevel.keep | 
| 38504 
76965c356d2a
removed Code_Prolog: modifies global environment setup non-conservatively
 haftmann parents: 
38115diff
changeset | 1006 | (values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*) | 
| 38075 | 1007 | |
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1008 | (* quickcheck generator *) | 
| 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1009 | |
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39465diff
changeset | 1010 | (* FIXME: a small clone of Predicate_Compile_Quickcheck - maybe refactor out commons *) | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1011 | |
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 1012 | fun quickcheck ctxt t size = | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1013 | let | 
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38948diff
changeset | 1014 | val options = code_options_of (ProofContext.theory_of ctxt) | 
| 38755 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 1015 | val thy = Theory.copy (ProofContext.theory_of ctxt) | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39465diff
changeset | 1016 | val ((((full_constname, constT), vs'), intro), thy1) = | 
| 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39465diff
changeset | 1017 | Predicate_Compile_Aux.define_quickcheck_predicate t thy | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1018 | val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39465diff
changeset | 1019 | val thy3 = Predicate_Compile.preprocess preprocess_options (Const (full_constname, constT)) thy2 | 
| 38755 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38735diff
changeset | 1020 | val ctxt' = ProofContext.init_global thy3 | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1021 | val _ = tracing "Generating prolog program..." | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 1022 | val (p, constant_table) = generate (NONE, true) ctxt' full_constname | 
| 39798 
9e7a0a9d194e
adding option to globally limit the prolog execution
 bulwahn parents: 
39761diff
changeset | 1023 | |> post_process ctxt' (set_ensure_groundness options) full_constname | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1024 | val _ = tracing "Running prolog program..." | 
| 39461 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 1025 | val system_config = System_Config.get (Context.Proof ctxt) | 
| 
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
 bulwahn parents: 
39383diff
changeset | 1026 | val tss = run (#timeout system_config, #prolog_system system_config) | 
| 39464 
13493d3c28d0
values command for prolog supports complex terms and not just variables
 bulwahn parents: 
39462diff
changeset | 1027 | p (translate_const constant_table full_constname, map (Var o fst) vs') (map fst vs') (SOME 1) | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1028 | val _ = tracing "Restoring terms..." | 
| 38961 
8c2f59171647
handling the quickcheck result no counterexample more correctly
 bulwahn parents: 
38960diff
changeset | 1029 | val res = | 
| 
8c2f59171647
handling the quickcheck result no counterexample more correctly
 bulwahn parents: 
38960diff
changeset | 1030 | case tss of | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39465diff
changeset | 1031 | [ts] => SOME (map (restore_term ctxt' constant_table) (ts ~~ map snd vs')) | 
| 38961 
8c2f59171647
handling the quickcheck result no counterexample more correctly
 bulwahn parents: 
38960diff
changeset | 1032 | | _ => NONE | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1033 | in | 
| 40924 | 1034 | (res, NONE) | 
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38732diff
changeset | 1035 | end; | 
| 38732 
3371dbc806ae
moving preprocessing to values in prolog generation
 bulwahn parents: 
38731diff
changeset | 1036 | |
| 38073 | 1037 | end; |