| author | bulwahn | 
| Sat, 24 Oct 2009 16:55:42 +0200 | |
| changeset 33125 | 2fef4f9429f7 | 
| parent 33124 | 5378e61add1a | 
| child 33126 | bb8806eb5da7 | 
| permissions | -rw-r--r-- | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
(* Author: Lukas Bulwahn, TU Muenchen  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
2  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
signature PREDICATE_COMPILE =  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
5  | 
sig  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
val setup : theory -> theory  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
7  | 
val preprocess : Predicate_Compile_Aux.options -> string -> theory -> theory  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
end;  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
structure Predicate_Compile : PREDICATE_COMPILE =  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
struct  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
|
| 
33108
 
9d9afd478016
added test for higher-order function inductification; added debug messages
 
bulwahn 
parents: 
33107 
diff
changeset
 | 
13  | 
(* options *)  | 
| 
33122
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
14  | 
val fail_safe_mode = false  | 
| 
33108
 
9d9afd478016
added test for higher-order function inductification; added debug messages
 
bulwahn 
parents: 
33107 
diff
changeset
 | 
15  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
open Predicate_Compile_Aux;  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
val priority = tracing;  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
(* Some last processing *)  | 
| 33113 | 21  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
fun remove_pointless_clauses intro =  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
  if Logic.strip_imp_prems (prop_of intro) = [@{prop "False"}] then
 | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
[]  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
else [intro]  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
27  | 
fun tracing s = ()  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
28  | 
|
| 33125 | 29  | 
fun print_intross options thy msg intross =  | 
30  | 
if show_intermediate_results options then  | 
|
31  | 
Output.tracing (msg ^  | 
|
| 33120 | 32  | 
(space_implode "; " (map  | 
33  | 
(fn intros => commas (map (Display.string_of_thm_global thy) intros)) intross)))  | 
|
| 33125 | 34  | 
else ()  | 
35  | 
||
| 
33122
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
36  | 
fun print_specs thy specs =  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
37  | 
map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
38  | 
^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
39  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
40  | 
fun process_specification options specs thy' =  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
41  | 
let  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
42  | 
val _ = print_step options "Compiling predicates to flat introrules..."  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
43  | 
val specs = map (apsnd (map  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
44  | 
(fn th => if is_equationlike th then Pred_Compile_Data.normalize_equation thy' th else th))) specs  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
45  | 
val (intross1, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess specs thy')  | 
| 33125 | 46  | 
val _ = print_intross options thy'' "Flattened introduction rules: " intross1  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
47  | 
val _ = "Replacing functions in introrules..."  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
48  | 
val intross2 =  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
49  | 
if fail_safe_mode then  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
50  | 
case try (burrow (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
51  | 
SOME intross => intross  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
52  | 
| NONE => let val _ = warning "Function replacement failed!" in intross1 end  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
53  | 
else burrow (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1  | 
| 33125 | 54  | 
val _ = print_intross options thy'' "Introduction rules with replaced functions: " intross2  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
55  | 
val _ = print_step options "Introducing new constants for abstractions at higher-order argument positions..."  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
56  | 
val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'')  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
57  | 
val (new_intross, thy'''') =  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
58  | 
if not (null new_defs) then  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
59  | 
let  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
60  | 
val _ = print_step options "Recursively obtaining introduction rules for new definitions..."  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
61  | 
in process_specification options new_defs thy''' end  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
62  | 
else ([], thy''')  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
63  | 
in  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
64  | 
(intross3 @ new_intross, thy'''')  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
65  | 
end  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
66  | 
|
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
67  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
68  | 
fun preprocess_strong_conn_constnames options gr constnames thy =  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
69  | 
let  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
70  | 
val get_specs = map (fn k => (k, Graph.get_node gr k))  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
71  | 
    val _ = print_step options ("Preprocessing scc of " ^ commas constnames)
 | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
val (prednames, funnames) = List.partition (is_pred thy) constnames  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
(* untangle recursion by defining predicates for all functions *)  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
74  | 
val _ = print_step options  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
75  | 
      ("Compiling functions (" ^ commas funnames ^ ") to predicates...")
 | 
| 
33122
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
76  | 
val (fun_pred_specs, thy') =  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
77  | 
if not (null funnames) then Predicate_Compile_Fun.define_predicates  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
78  | 
(get_specs funnames) thy else ([], thy)  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
79  | 
val _ = print_specs thy' fun_pred_specs  | 
| 
 
7d01480cc8e3
added first support for higher-order function translation
 
bulwahn 
parents: 
33121 
diff
changeset
 | 
80  | 
val specs = (get_specs prednames) @ fun_pred_specs  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
81  | 
val (intross3, thy''') = process_specification options specs thy'  | 
| 33125 | 82  | 
val _ = print_intross options thy''' "Introduction rules with new constants: " intross3  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
83  | 
val intross4 = map (maps remove_pointless_clauses) intross3  | 
| 33125 | 84  | 
val _ = print_intross options thy''' "After removing pointless clauses: " intross4  | 
| 
33124
 
5378e61add1a
continued cleaning up; moved tuple expanding to core
 
bulwahn 
parents: 
33123 
diff
changeset
 | 
85  | 
val intross5 = map (map (AxClass.overload thy''')) intross4  | 
| 
 
5378e61add1a
continued cleaning up; moved tuple expanding to core
 
bulwahn 
parents: 
33123 
diff
changeset
 | 
86  | 
val intross6 = burrow (map (expand_tuples thy''')) intross5  | 
| 33125 | 87  | 
val _ = print_intross options thy''' "introduction rules before registering: " intross6  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
88  | 
val _ = print_step options "Registering introduction rules..."  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
89  | 
val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
90  | 
in  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
91  | 
thy''''  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
92  | 
end;  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
94  | 
fun preprocess options const thy =  | 
| 
32672
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
95  | 
let  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
96  | 
val _ = print_step options "Fetching definitions from theory..."  | 
| 
32672
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
97  | 
val table = Pred_Compile_Data.make_const_spec_table thy  | 
| 
33107
 
6aa76ce59ef2
added filtering of case constants in the definition retrieval of the predicate compiler
 
bulwahn 
parents: 
32672 
diff
changeset
 | 
98  | 
val gr = Pred_Compile_Data.obtain_specification_graph thy table const  | 
| 
32672
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
99  | 
val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
100  | 
in fold_rev (preprocess_strong_conn_constnames options gr)  | 
| 
32672
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
101  | 
(Graph.strong_conn gr) thy  | 
| 
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
102  | 
end  | 
| 
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
103  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
104  | 
fun extract_options ((modes, raw_options), raw_const) =  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
105  | 
let  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
106  | 
fun chk s = member (op =) raw_options s  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
107  | 
in  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
108  | 
    Options {
 | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
109  | 
show_steps = chk "show_steps",  | 
| 33125 | 110  | 
show_intermediate_results = chk "show_intermediate_results",  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
111  | 
show_mode_inference = chk "show_mode_inference",  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
112  | 
inductify = chk "inductify",  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
113  | 
rpred = chk "rpred"  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
114  | 
}  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
115  | 
end  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
116  | 
|
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
117  | 
fun code_pred_cmd ((modes, raw_options), raw_const) lthy =  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
118  | 
let  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
119  | 
val options = extract_options ((modes, raw_options), raw_const)  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
120  | 
in  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
121  | 
if (is_inductify options) then  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
122  | 
let  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
123  | 
val thy = ProofContext.theory_of lthy  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
124  | 
val const = Code.read_const thy raw_const  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
125  | 
val lthy' = LocalTheory.theory (preprocess options const) lthy  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
126  | 
|> LocalTheory.checkpoint  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
127  | 
|
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
128  | 
val const = case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
129  | 
SOME c => c  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
130  | 
| NONE => const  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
131  | 
val _ = print_step options "Starting Predicate Compile Core..."  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
132  | 
in  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
133  | 
Predicate_Compile_Core.code_pred options modes (is_rpred options) const lthy'  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
134  | 
end  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
135  | 
else  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
136  | 
Predicate_Compile_Core.code_pred_cmd options modes (is_rpred options) raw_const lthy  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
137  | 
end  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
139  | 
val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
|
| 33125 | 141  | 
val bool_options = ["show_steps", "show_intermediate_results", "show_mode_inference", "inductify", "rpred"]  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
142  | 
|
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
143  | 
val _ = List.app OuterKeyword.keyword ("mode" :: bool_options)
 | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
145  | 
local structure P = OuterParse  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
146  | 
in  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
147  | 
|
| 
33114
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
148  | 
val opt_modes =  | 
| 
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
149  | 
  Scan.optional (P.$$$ "(" |-- P.$$$ "mode" |-- P.$$$ ":" |--
 | 
| 
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
150  | 
P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")  | 
| 
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
151  | 
--| P.$$$ ")" >> SOME) NONE  | 
| 
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
152  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
153  | 
val scan_params =  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
154  | 
let  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
155  | 
val scan_bool_param = foldl1 (op ||) (map P.$$$ bool_options)  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
156  | 
in  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
157  | 
Scan.optional (P.$$$ "[" |-- P.enum1 "," scan_bool_param --| P.$$$ "]") []  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
158  | 
end  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
159  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
160  | 
val _ = OuterSyntax.local_theory_to_proof "code_pred"  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
161  | 
"prove equations for predicate specified by intro/elim rules"  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
162  | 
OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >>  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
163  | 
code_pred_cmd)  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
165  | 
end  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
166  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
167  | 
end  |