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