| author | wenzelm | 
| Sun, 13 Mar 2011 15:10:00 +0100 | |
| changeset 41941 | f823f7fae9a2 | 
| parent 40053 | 3fa49ea76cbb | 
| child 42011 | dee23d63d466 | 
| permissions | -rw-r--r-- | 
| 33265 | 1  | 
(* Title: HOL/Tools/Predicate_Compile/predicate_compile.ML  | 
2  | 
Author: Lukas Bulwahn, TU Muenchen  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
|
| 41941 | 4  | 
Entry point for the predicate compiler; definition of Toplevel  | 
5  | 
commands code_pred and values.  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 33265 | 7  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
signature PREDICATE_COMPILE =  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
val setup : theory -> theory  | 
| 
36050
 
88203782cf12
activating the signature of Predicate_Compile again which was deactivated in changeset d93a3cb55068 for debugging purposes
 
bulwahn 
parents: 
36032 
diff
changeset
 | 
11  | 
val preprocess : Predicate_Compile_Aux.options -> term -> theory -> theory  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
12  | 
val present_graph : bool Unsynchronized.ref  | 
| 
36023
 
d606ca1674a7
adding a hook for experiments in the predicate compilation process
 
bulwahn 
parents: 
36022 
diff
changeset
 | 
13  | 
val intro_hook : (theory -> thm -> unit) option Unsynchronized.ref  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
end;  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
|
| 
36050
 
88203782cf12
activating the signature of Predicate_Compile again which was deactivated in changeset d93a3cb55068 for debugging purposes
 
bulwahn 
parents: 
36032 
diff
changeset
 | 
16  | 
structure Predicate_Compile : PREDICATE_COMPILE =  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
struct  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
|
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
19  | 
val present_graph = Unsynchronized.ref false  | 
| 
33108
 
9d9afd478016
added test for higher-order function inductification; added debug messages
 
bulwahn 
parents: 
33107 
diff
changeset
 | 
20  | 
|
| 
36023
 
d606ca1674a7
adding a hook for experiments in the predicate compilation process
 
bulwahn 
parents: 
36022 
diff
changeset
 | 
21  | 
val intro_hook = Unsynchronized.ref NONE : (theory -> thm -> unit) option Unsynchronized.ref  | 
| 
 
d606ca1674a7
adding a hook for experiments in the predicate compilation process
 
bulwahn 
parents: 
36022 
diff
changeset
 | 
22  | 
|
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
open Predicate_Compile_Aux;  | 
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
|
| 33125 | 25  | 
fun print_intross options thy msg intross =  | 
26  | 
if show_intermediate_results options then  | 
|
| 
33376
 
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
 
bulwahn 
parents: 
33375 
diff
changeset
 | 
27  | 
tracing (msg ^  | 
| 
 
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
 
bulwahn 
parents: 
33375 
diff
changeset
 | 
28  | 
(space_implode "\n" (map  | 
| 
 
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
 
bulwahn 
parents: 
33375 
diff
changeset
 | 
29  | 
(fn (c, intros) => "Introduction rule(s) of " ^ c ^ ":\n" ^  | 
| 
 
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
 
bulwahn 
parents: 
33375 
diff
changeset
 | 
30  | 
commas (map (Display.string_of_thm_global thy) intros)) intross)))  | 
| 33125 | 31  | 
else ()  | 
32  | 
||
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
33  | 
fun print_specs options thy specs =  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
34  | 
if show_intermediate_results options then  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
35  | 
map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
36  | 
^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
37  | 
|> space_implode "\n" |> tracing  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
38  | 
else ()  | 
| 
33475
 
42fed8eac357
improved handling of overloaded constants; examples with numerals
 
bulwahn 
parents: 
33473 
diff
changeset
 | 
39  | 
fun overload_const thy s = the_default s (Option.map fst (AxClass.inst_of_param thy s))  | 
| 33404 | 40  | 
|
| 
33146
 
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
 
bulwahn 
parents: 
33142 
diff
changeset
 | 
41  | 
fun map_specs f specs =  | 
| 
 
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
 
bulwahn 
parents: 
33142 
diff
changeset
 | 
42  | 
map (fn (s, ths) => (s, f ths)) specs  | 
| 
 
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
 
bulwahn 
parents: 
33142 
diff
changeset
 | 
43  | 
|
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
44  | 
fun process_specification options specs thy' =  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
45  | 
let  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
46  | 
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
 | 
47  | 
val specs = map (apsnd (map  | 
| 
33140
 
83951822bfd0
cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
 
bulwahn 
parents: 
33139 
diff
changeset
 | 
48  | 
(fn th => if is_equationlike th then Predicate_Compile_Data.normalize_equation thy' th else th))) specs  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
49  | 
val (intross1, thy'') =  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
50  | 
apfst flat (fold_map (Predicate_Compile_Pred.preprocess options) specs thy')  | 
| 33125 | 51  | 
val _ = print_intross options thy'' "Flattened introduction rules: " intross1  | 
| 
33129
 
3085da75ed54
changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
 
bulwahn 
parents: 
33127 
diff
changeset
 | 
52  | 
val _ = print_step options "Replacing functions in introrules..."  | 
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
53  | 
val intross2 =  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
54  | 
if function_flattening options then  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
55  | 
if fail_safe_function_flattening options then  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
56  | 
case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
57  | 
SOME intross => intross  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
58  | 
| NONE =>  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
59  | 
(if show_caught_failures options then tracing "Function replacement failed!" else ();  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
60  | 
intross1)  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
61  | 
else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
62  | 
else  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
63  | 
intross1  | 
| 33125 | 64  | 
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
 | 
65  | 
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
 | 
66  | 
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
 | 
67  | 
val (new_intross, thy'''') =  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
68  | 
if not (null new_defs) then  | 
| 33404 | 69  | 
let  | 
70  | 
val _ = print_step options "Recursively obtaining introduction rules for new definitions..."  | 
|
71  | 
in process_specification options new_defs thy''' end  | 
|
72  | 
else ([], thy''')  | 
|
| 
33121
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
73  | 
in  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
74  | 
(intross3 @ new_intross, thy'''')  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
75  | 
end  | 
| 
 
9b10dc5da0e0
added to process higher-order arguments by adding new constants
 
bulwahn 
parents: 
33120 
diff
changeset
 | 
76  | 
|
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
77  | 
fun preprocess_strong_conn_constnames options gr ts thy =  | 
| 
37003
 
a393a588b82e
moving towards working with proof contexts in the predicate compiler
 
bulwahn 
parents: 
36960 
diff
changeset
 | 
78  | 
if forall (fn (Const (c, _)) =>  | 
| 
40053
 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 
bulwahn 
parents: 
40048 
diff
changeset
 | 
79  | 
Core_Data.is_registered (ProofContext.init_global thy) c) ts then  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
80  | 
thy  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
81  | 
else  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
82  | 
let  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
83  | 
fun get_specs ts = map_filter (fn t =>  | 
| 35404 | 84  | 
Term_Graph.get_node gr t |>  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
85  | 
(fn ths => if null ths then NONE else SOME (fst (dest_Const t), ths)))  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
86  | 
ts  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
87  | 
      val _ = print_step options ("Preprocessing scc of " ^
 | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
88  | 
commas (map (Syntax.string_of_term_global thy) ts))  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
89  | 
      val (prednames, funnames) = List.partition (fn t => body_type (fastype_of t) = @{typ bool}) ts
 | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
90  | 
(* untangle recursion by defining predicates for all functions *)  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
91  | 
val _ = print_step options  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
92  | 
        ("Compiling functions (" ^ commas (map (Syntax.string_of_term_global thy) funnames) ^
 | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
93  | 
") to predicates...")  | 
| 
36032
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
94  | 
val (fun_pred_specs, thy1) =  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
95  | 
(if function_flattening options andalso (not (null funnames)) then  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
96  | 
if fail_safe_function_flattening options then  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
97  | 
case try (Predicate_Compile_Fun.define_predicates (get_specs funnames)) thy of  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
98  | 
SOME (intross, thy) => (intross, thy)  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
99  | 
| NONE => ([], thy)  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
100  | 
else Predicate_Compile_Fun.define_predicates (get_specs funnames) thy  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
101  | 
else ([], thy))  | 
| 
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
102  | 
(*||> Theory.checkpoint*)  | 
| 
36032
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
103  | 
val _ = print_specs options thy1 fun_pred_specs  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
104  | 
val specs = (get_specs prednames) @ fun_pred_specs  | 
| 
36032
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
105  | 
val (intross3, thy2) = process_specification options specs thy1  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
106  | 
val _ = print_intross options thy2 "Introduction rules with new constants: " intross3  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
107  | 
val intross4 = map_specs (maps remove_pointless_clauses) intross3  | 
| 
36032
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
108  | 
val _ = print_intross options thy2 "After removing pointless clauses: " intross4  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
109  | 
val intross5 = map_specs (map (remove_equalities thy2)) intross4  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
110  | 
val _ = print_intross options thy2 "After removing equality premises:" intross5  | 
| 
36022
 
c0fa8499e366
removing simple equalities in introduction rules in the predicate compiler
 
bulwahn 
parents: 
36004 
diff
changeset
 | 
111  | 
val intross6 =  | 
| 
36032
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
112  | 
map (fn (s, ths) => (overload_const thy2 s, map (AxClass.overload thy2) ths)) intross5  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
113  | 
val intross7 = map_specs (map (expand_tuples thy2)) intross6  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
114  | 
val intross8 = map_specs (map (eta_contract_ho_arguments thy2)) intross7  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
115  | 
val _ = case !intro_hook of NONE => () | SOME f => (map_specs (map (f thy2)) intross8; ())  | 
| 
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
116  | 
      val _ = print_step options ("Looking for specialisations in " ^ commas (map fst intross8) ^ "...")
 | 
| 
36248
 
9ed1a37de465
added option for specialisation to the predicate compiler
 
bulwahn 
parents: 
36246 
diff
changeset
 | 
117  | 
val (intross9, thy3) =  | 
| 
 
9ed1a37de465
added option for specialisation to the predicate compiler
 
bulwahn 
parents: 
36246 
diff
changeset
 | 
118  | 
if specialise options then  | 
| 
 
9ed1a37de465
added option for specialisation to the predicate compiler
 
bulwahn 
parents: 
36246 
diff
changeset
 | 
119  | 
Predicate_Compile_Specialisation.find_specialisations [] intross8 thy2  | 
| 
 
9ed1a37de465
added option for specialisation to the predicate compiler
 
bulwahn 
parents: 
36246 
diff
changeset
 | 
120  | 
else (intross8, thy2)  | 
| 
36246
 
43fecedff8cf
added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
 
bulwahn 
parents: 
36050 
diff
changeset
 | 
121  | 
val _ = print_intross options thy3 "introduction rules after specialisations: " intross9  | 
| 
 
43fecedff8cf
added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
 
bulwahn 
parents: 
36050 
diff
changeset
 | 
122  | 
val intross10 = map_specs (map_filter (peephole_optimisation thy3)) intross9  | 
| 
 
43fecedff8cf
added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
 
bulwahn 
parents: 
36050 
diff
changeset
 | 
123  | 
val _ = print_intross options thy3 "introduction rules before registering: " intross10  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
124  | 
val _ = print_step options "Registering introduction rules..."  | 
| 
40053
 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 
bulwahn 
parents: 
40048 
diff
changeset
 | 
125  | 
val thy4 = fold Core_Data.register_intros intross10 thy3  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
126  | 
in  | 
| 
36032
 
dfd30b5b4e73
adding specialisation of predicates to the predicate compiler
 
bulwahn 
parents: 
36027 
diff
changeset
 | 
127  | 
thy4  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
128  | 
end;  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
129  | 
|
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
130  | 
fun preprocess options t thy =  | 
| 
32672
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
131  | 
let  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
132  | 
val _ = print_step options "Fetching definitions from theory..."  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
133  | 
val gr = Output.cond_timeit (!Quickcheck.timing) "preprocess-obtain graph"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
134  | 
(fn () => Predicate_Compile_Data.obtain_specification_graph options thy t  | 
| 35404 | 135  | 
|> (fn gr => Term_Graph.subgraph (member (op =) (Term_Graph.all_succs gr [t])) gr))  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
136  | 
val _ = if !present_graph then Predicate_Compile_Data.present_graph gr else ()  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
137  | 
in  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
138  | 
Output.cond_timeit (!Quickcheck.timing) "preprocess-process"  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
139  | 
(fn () => (fold_rev (preprocess_strong_conn_constnames options gr)  | 
| 35404 | 140  | 
(Term_Graph.strong_conn gr) thy))  | 
| 
32672
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
141  | 
end  | 
| 
 
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
 
bulwahn 
parents: 
32669 
diff
changeset
 | 
142  | 
|
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
143  | 
datatype proposed_modes = Multiple_Preds of (string * (mode * string option) list) list  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
144  | 
| Single_Pred of (mode * string option) list  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
145  | 
|
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
146  | 
fun extract_options lthy (((expected_modes, proposed_modes), (compilation, raw_options)), const) =  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
147  | 
let  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
148  | 
fun chk s = member (op =) raw_options s  | 
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
149  | 
val proposed_modes = case proposed_modes of  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
150  | 
Single_Pred proposed_modes => [(const, proposed_modes)]  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
151  | 
| Multiple_Preds proposed_modes => map  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
152  | 
(apfst (Code.read_const (ProofContext.theory_of lthy))) proposed_modes  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
153  | 
in  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
154  | 
    Options {
 | 
| 
33623
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
155  | 
expected_modes = Option.map (pair const) expected_modes,  | 
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
156  | 
proposed_modes =  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
157  | 
map (apsnd (map fst)) proposed_modes,  | 
| 
33623
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
158  | 
proposed_names =  | 
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
159  | 
maps (fn (predname, ms) => (map_filter  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
160  | 
(fn (m, NONE) => NONE | (m, SOME name) => SOME ((predname, m), name))) ms) proposed_modes,  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
161  | 
show_steps = chk "show_steps",  | 
| 33125 | 162  | 
show_intermediate_results = chk "show_intermediate_results",  | 
| 33127 | 163  | 
show_proof_trace = chk "show_proof_trace",  | 
| 33251 | 164  | 
show_modes = chk "show_modes",  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
165  | 
show_mode_inference = chk "show_mode_inference",  | 
| 
33139
 
9c01ee6f8ee9
added skip_proof option; playing with compilation of depth-limited predicates
 
bulwahn 
parents: 
33134 
diff
changeset
 | 
166  | 
show_compilation = chk "show_compilation",  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
167  | 
show_caught_failures = false,  | 
| 
39383
 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 
bulwahn 
parents: 
39382 
diff
changeset
 | 
168  | 
show_invalid_clauses = chk "show_invalid_clauses",  | 
| 
33139
 
9c01ee6f8ee9
added skip_proof option; playing with compilation of depth-limited predicates
 
bulwahn 
parents: 
33134 
diff
changeset
 | 
169  | 
skip_proof = chk "skip_proof",  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
170  | 
function_flattening = not (chk "no_function_flattening"),  | 
| 
36248
 
9ed1a37de465
added option for specialisation to the predicate compiler
 
bulwahn 
parents: 
36246 
diff
changeset
 | 
171  | 
specialise = chk "specialise",  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
172  | 
fail_safe_function_flattening = false,  | 
| 
35381
 
5038f36b5ea1
adding no_topmost_reordering as new option to the code_pred command
 
bulwahn 
parents: 
35324 
diff
changeset
 | 
173  | 
no_topmost_reordering = (chk "no_topmost_reordering"),  | 
| 
35324
 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 
bulwahn 
parents: 
34948 
diff
changeset
 | 
174  | 
no_higher_order_predicate = [],  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
175  | 
inductify = chk "inductify",  | 
| 36254 | 176  | 
detect_switches = chk "detect_switches",  | 
| 
40048
 
f3a46d524101
adding option smart_depth_limiting to predicate compiler
 
bulwahn 
parents: 
39383 
diff
changeset
 | 
177  | 
smart_depth_limiting = chk "smart_depth_limiting",  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
178  | 
compilation = compilation  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
179  | 
}  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
180  | 
end  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
181  | 
|
| 
33623
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
182  | 
fun code_pred_cmd (((expected_modes, proposed_modes), raw_options), raw_const) lthy =  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
183  | 
let  | 
| 
33132
 
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
 
bulwahn 
parents: 
33131 
diff
changeset
 | 
184  | 
val thy = ProofContext.theory_of lthy  | 
| 
 
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
 
bulwahn 
parents: 
33131 
diff
changeset
 | 
185  | 
val const = Code.read_const thy raw_const  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
186  | 
val T = Sign.the_const_type thy const  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
187  | 
val t = Const (const, T)  | 
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
188  | 
val options = extract_options lthy (((expected_modes, proposed_modes), raw_options), const)  | 
| 
33132
 
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
 
bulwahn 
parents: 
33131 
diff
changeset
 | 
189  | 
in  | 
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38664 
diff
changeset
 | 
190  | 
if is_inductify options then  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
191  | 
let  | 
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38664 
diff
changeset
 | 
192  | 
val lthy' = Local_Theory.background_theory (preprocess options t) lthy  | 
| 
33623
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
193  | 
val const =  | 
| 
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
194  | 
case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
195  | 
SOME c => c  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
196  | 
| NONE => const  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
197  | 
val _ = print_step options "Starting Predicate Compile Core..."  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
198  | 
in  | 
| 
33132
 
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
 
bulwahn 
parents: 
33131 
diff
changeset
 | 
199  | 
Predicate_Compile_Core.code_pred options const lthy'  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
200  | 
end  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
201  | 
else  | 
| 
33132
 
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
 
bulwahn 
parents: 
33131 
diff
changeset
 | 
202  | 
Predicate_Compile_Core.code_pred_cmd options raw_const lthy  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
203  | 
end  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
204  | 
|
| 
33630
 
68e058d061f5
removed unnecessary oracle in the predicate compiler
 
bulwahn 
parents: 
33625 
diff
changeset
 | 
205  | 
val setup = Predicate_Compile_Core.setup  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
206  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
207  | 
|
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
208  | 
(* Parser for mode annotations *)  | 
| 33327 | 209  | 
|
| 
33625
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
210  | 
fun parse_mode_basic_expr xs =  | 
| 
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
211  | 
(Args.$$$ "i" >> K Input || Args.$$$ "o" >> K Output ||  | 
| 
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
212  | 
    Args.$$$ "bool" >> K Bool || Args.$$$ "(" |-- parse_mode_expr --| Args.$$$ ")") xs
 | 
| 
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
213  | 
and parse_mode_tuple_expr xs =  | 
| 
38664
 
7215ae18f44b
added support for xsymbol syntax for mode annotations in code_pred command
 
bulwahn 
parents: 
37003 
diff
changeset
 | 
214  | 
(parse_mode_basic_expr --| (Args.$$$ "*" || Args.$$$ "\<times>") -- parse_mode_tuple_expr >> Pair || parse_mode_basic_expr)  | 
| 
33625
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
215  | 
xs  | 
| 
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
216  | 
and parse_mode_expr xs =  | 
| 
38664
 
7215ae18f44b
added support for xsymbol syntax for mode annotations in code_pred command
 
bulwahn 
parents: 
37003 
diff
changeset
 | 
217  | 
(parse_mode_tuple_expr --| (Args.$$$ "=>" || Args.$$$ "\<Rightarrow>") -- parse_mode_expr >> Fun || parse_mode_tuple_expr) xs  | 
| 
33619
 
d93a3cb55068
first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
 
bulwahn 
parents: 
33481 
diff
changeset
 | 
218  | 
|
| 
33625
 
eefee41ede3a
removed deprecated mode annotation parser; renamed accepted mode annotation parser to nicer naming
 
bulwahn 
parents: 
33623 
diff
changeset
 | 
219  | 
val mode_and_opt_proposal = parse_mode_expr --  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
220  | 
Scan.optional (Args.$$$ "as" |-- Parse.xname >> SOME) NONE  | 
| 
33620
 
b6bf2dc5aed7
added interface of user proposals for names of generated constants
 
bulwahn 
parents: 
33619 
diff
changeset
 | 
221  | 
|
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
222  | 
|
| 
33114
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
223  | 
val opt_modes =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
224  | 
  Scan.optional (Parse.$$$ "(" |-- Args.$$$ "modes" |-- Parse.$$$ ":" |--
 | 
| 
39382
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
225  | 
(((Parse.enum1 "and" (Parse.xname --| Parse.$$$ ":" --  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
226  | 
(Parse.enum "," mode_and_opt_proposal))) >> Multiple_Preds)  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
227  | 
|| ((Parse.enum "," mode_and_opt_proposal) >> Single_Pred))  | 
| 
 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 
bulwahn 
parents: 
38757 
diff
changeset
 | 
228  | 
--| Parse.$$$ ")") (Multiple_Preds [])  | 
| 
33623
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
229  | 
|
| 
 
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
 
bulwahn 
parents: 
33620 
diff
changeset
 | 
230  | 
val opt_expected_modes =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
231  | 
  Scan.optional (Parse.$$$ "(" |-- Args.$$$ "expected_modes" |-- Parse.$$$ ":" |--
 | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
232  | 
Parse.enum "," parse_mode_expr --| Parse.$$$ ")" >> SOME) NONE  | 
| 
33114
 
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
 
bulwahn 
parents: 
33113 
diff
changeset
 | 
233  | 
|
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
234  | 
(* Parser for options *)  | 
| 
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
235  | 
|
| 
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
236  | 
val scan_options =  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
237  | 
let  | 
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
238  | 
val scan_bool_option = foldl1 (op ||) (map Args.$$$ bool_options)  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
239  | 
val scan_compilation = foldl1 (op ||) (map (fn (s, c) => Args.$$$ s >> K c) compilation_names)  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
240  | 
in  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
241  | 
Scan.optional (Parse.$$$ "[" |-- Scan.optional scan_compilation Pred  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
242  | 
-- Parse.enum "," scan_bool_option --| Parse.$$$ "]")  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
243  | 
(Pred, [])  | 
| 
33123
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
244  | 
end  | 
| 
 
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
 
bulwahn 
parents: 
33122 
diff
changeset
 | 
245  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
246  | 
val opt_print_modes =  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
247  | 
  Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
 | 
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
248  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
249  | 
val opt_mode = (Parse.$$$ "_" >> K NONE) || (parse_mode_expr >> SOME)  | 
| 
33479
 
428ddcc16e7b
added optional mode annotations for parameters in the values command
 
bulwahn 
parents: 
33478 
diff
changeset
 | 
250  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
251  | 
val opt_param_modes = Scan.optional (Parse.$$$ "[" |-- Args.$$$ "mode" |-- Parse.$$$ ":" |--  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
252  | 
Parse.enum ", " opt_mode --| Parse.$$$ "]" >> SOME) NONE  | 
| 
33479
 
428ddcc16e7b
added optional mode annotations for parameters in the values command
 
bulwahn 
parents: 
33478 
diff
changeset
 | 
253  | 
|
| 
36027
 
29a15da9c63d
added statistics to values command for random generation
 
bulwahn 
parents: 
36023 
diff
changeset
 | 
254  | 
val stats = Scan.optional (Args.$$$ "stats" >> K true) false  | 
| 
 
29a15da9c63d
added statistics to values command for random generation
 
bulwahn 
parents: 
36023 
diff
changeset
 | 
255  | 
|
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
256  | 
val value_options =  | 
| 
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
257  | 
let  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
258  | 
val expected_values = Scan.optional (Args.$$$ "expected" |-- Parse.term >> SOME) NONE  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
259  | 
val scan_compilation =  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
260  | 
Scan.optional  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
261  | 
(foldl1 (op ||)  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
262  | 
(map (fn (s, c) => Args.$$$ s -- Parse.enum "," Parse.int >> (fn (_, ps) => (c, ps)))  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
263  | 
compilation_names))  | 
| 
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
264  | 
(Pred, [])  | 
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
265  | 
in  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
266  | 
Scan.optional  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
267  | 
(Parse.$$$ "[" |-- (expected_values -- stats) -- scan_compilation --| Parse.$$$ "]")  | 
| 
36027
 
29a15da9c63d
added statistics to values command for random generation
 
bulwahn 
parents: 
36023 
diff
changeset
 | 
268  | 
((NONE, false), (Pred, []))  | 
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
269  | 
end  | 
| 
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
270  | 
|
| 
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
271  | 
(* code_pred command and values command *)  | 
| 
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
272  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
273  | 
val _ = Outer_Syntax.local_theory_to_proof "code_pred"  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
274  | 
"prove equations for predicate specified by intro/elim rules"  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
275  | 
Keyword.thy_goal  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
276  | 
(opt_expected_modes -- opt_modes -- scan_options -- Parse.term_group >> code_pred_cmd)  | 
| 
33478
 
b70fe083694d
moved values command from core to predicate compile
 
bulwahn 
parents: 
33475 
diff
changeset
 | 
277  | 
|
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
278  | 
val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36254 
diff
changeset
 | 
279  | 
(opt_print_modes -- opt_param_modes -- value_options -- Scan.optional Parse.nat ~1 -- Parse.term  | 
| 
34948
 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 
bulwahn 
parents: 
33752 
diff
changeset
 | 
280  | 
>> (fn ((((print_modes, param_modes), options), k), t) => Toplevel.keep  | 
| 
33479
 
428ddcc16e7b
added optional mode annotations for parameters in the values command
 
bulwahn 
parents: 
33478 
diff
changeset
 | 
281  | 
(Predicate_Compile_Core.values_cmd print_modes param_modes options k t)));  | 
| 
32668
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
282  | 
|
| 
 
b2de45007537
added first prototype of the extended predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
283  | 
end  |