author | wenzelm |
Mon, 10 Mar 2014 18:06:23 +0100 | |
changeset 56032 | b034b9f0fa2a |
parent 55757 | 9fc71814b8c1 |
child 56239 | 17df7145a871 |
permissions | -rw-r--r-- |
33265 | 1 |
(* Title: HOL/Tools/Predicate_Compile/predicate_compile_core.ML |
2 |
Author: Lukas Bulwahn, TU Muenchen |
|
32667 | 3 |
|
33265 | 4 |
A compiler from predicates specified by intro/elim rules to equations. |
32667 | 5 |
*) |
6 |
||
7 |
signature PREDICATE_COMPILE_CORE = |
|
8 |
sig |
|
50057 | 9 |
type seed = Random_Engine.seed |
36048 | 10 |
type mode = Predicate_Compile_Aux.mode |
11 |
type options = Predicate_Compile_Aux.options |
|
12 |
type compilation = Predicate_Compile_Aux.compilation |
|
13 |
type compilation_funs = Predicate_Compile_Aux.compilation_funs |
|
55437 | 14 |
|
33478
b70fe083694d
moved values command from core to predicate compile
bulwahn
parents:
33476
diff
changeset
|
15 |
val setup : theory -> theory |
36048 | 16 |
val code_pred : options -> string -> Proof.context -> Proof.state |
17 |
val code_pred_cmd : options -> string -> Proof.context -> Proof.state |
|
55437 | 18 |
val values_cmd : string list -> mode option list option -> |
19 |
((string option * bool) * (compilation * int list)) -> int -> string -> Toplevel.state -> unit |
|
40053
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
bulwahn
parents:
40052
diff
changeset
|
20 |
|
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
21 |
val values_timeout : real Config.T |
55437 | 22 |
|
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
23 |
val print_stored_rules : Proof.context -> unit |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
24 |
val print_all_modes : compilation -> Proof.context -> unit |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
25 |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
26 |
val put_pred_result : (unit -> term Predicate.pred) -> Proof.context -> Proof.context |
50057 | 27 |
val put_pred_random_result : (unit -> seed -> term Predicate.pred * seed) -> |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
28 |
Proof.context -> Proof.context |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
29 |
val put_dseq_result : (unit -> term Limited_Sequence.dseq) -> Proof.context -> Proof.context |
55437 | 30 |
val put_dseq_random_result : |
31 |
(unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> |
|
32 |
term Limited_Sequence.dseq * seed) -> |
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
33 |
Proof.context -> Proof.context |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
34 |
val put_new_dseq_result : (unit -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence) -> |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
35 |
Proof.context -> Proof.context |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
36 |
val put_lseq_random_result : |
55437 | 37 |
(unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> |
38 |
term Lazy_Sequence.lazy_sequence) -> |
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
39 |
Proof.context -> Proof.context |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
40 |
val put_lseq_random_stats_result : |
55437 | 41 |
(unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> |
42 |
(term * Code_Numeral.natural) Lazy_Sequence.lazy_sequence) -> |
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
43 |
Proof.context -> Proof.context |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
44 |
|
33628 | 45 |
val code_pred_intro_attrib : attribute |
55437 | 46 |
(* used by Quickcheck_Generator *) |
32667 | 47 |
(* temporary for testing of the compilation *) |
36048 | 48 |
val add_equations : options -> string list -> theory -> theory |
49 |
val add_depth_limited_random_equations : options -> string list -> theory -> theory |
|
50 |
val add_random_dseq_equations : options -> string list -> theory -> theory |
|
51 |
val add_new_random_dseq_equations : options -> string list -> theory -> theory |
|
40103 | 52 |
val add_generator_dseq_equations : options -> string list -> theory -> theory |
45450 | 53 |
val add_generator_cps_equations : options -> string list -> theory -> theory |
33473
3b275a0bf18c
adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents:
33377
diff
changeset
|
54 |
val mk_tracing : string -> term -> term |
39310
17ef4415b17c
removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents:
39299
diff
changeset
|
55 |
val prepare_intrs : options -> Proof.context -> string list -> thm list -> |
38957
2eb265efa1b0
exporting mode analysis for use in prolog generation
bulwahn
parents:
38864
diff
changeset
|
56 |
((string * typ) list * string list * string list * (string * mode list) list * |
2eb265efa1b0
exporting mode analysis for use in prolog generation
bulwahn
parents:
38864
diff
changeset
|
57 |
(string * (Term.term list * Predicate_Compile_Aux.indprem list) list) list) |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
58 |
type mode_analysis_options = |
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
59 |
{use_generators : bool, |
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
60 |
reorder_premises : bool, |
55437 | 61 |
infer_pos_and_neg_modes : bool} |
38957
2eb265efa1b0
exporting mode analysis for use in prolog generation
bulwahn
parents:
38864
diff
changeset
|
62 |
datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode |
39310
17ef4415b17c
removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents:
39299
diff
changeset
|
63 |
| Mode_Pair of mode_derivation * mode_derivation | Term of mode |
39461
0ed0f015d140
adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
bulwahn
parents:
39403
diff
changeset
|
64 |
val head_mode_of : mode_derivation -> mode |
38957
2eb265efa1b0
exporting mode analysis for use in prolog generation
bulwahn
parents:
38864
diff
changeset
|
65 |
type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list |
2eb265efa1b0
exporting mode analysis for use in prolog generation
bulwahn
parents:
38864
diff
changeset
|
66 |
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list |
2eb265efa1b0
exporting mode analysis for use in prolog generation
bulwahn
parents:
38864
diff
changeset
|
67 |
|
32667 | 68 |
end; |
69 |
||
70 |
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE = |
|
71 |
struct |
|
72 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
73 |
type random_seed = Random_Engine.seed |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
74 |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
75 |
open Predicate_Compile_Aux; |
40052
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
bulwahn
parents:
40051
diff
changeset
|
76 |
open Core_Data; |
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
bulwahn
parents:
40051
diff
changeset
|
77 |
open Mode_Inference; |
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
bulwahn
parents:
40051
diff
changeset
|
78 |
open Predicate_Compile_Proof; |
33108
9d9afd478016
added test for higher-order function inductification; added debug messages
bulwahn
parents:
33106
diff
changeset
|
79 |
|
50057 | 80 |
type seed = Random_Engine.seed; |
81 |
||
82 |
||
32667 | 83 |
(** fundamentals **) |
84 |
||
85 |
(* syntactic operations *) |
|
86 |
||
87 |
fun mk_eq (x, xs) = |
|
88 |
let fun mk_eqs _ [] = [] |
|
89 |
| mk_eqs a (b::cs) = |
|
90 |
HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs |
|
91 |
in mk_eqs x xs end; |
|
92 |
||
33473
3b275a0bf18c
adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents:
33377
diff
changeset
|
93 |
fun mk_tracing s t = |
3b275a0bf18c
adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents:
33377
diff
changeset
|
94 |
Const(@{const_name Code_Evaluation.tracing}, |
3b275a0bf18c
adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents:
33377
diff
changeset
|
95 |
@{typ String.literal} --> (fastype_of t) --> (fastype_of t)) $ (HOLogic.mk_literal s) $ t |
3b275a0bf18c
adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents:
33377
diff
changeset
|
96 |
|
55437 | 97 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
98 |
(* representation of inferred clauses with modes *) |
32667 | 99 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
100 |
type moded_clause = term list * (indprem * mode_derivation) list |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
101 |
|
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:
35267
diff
changeset
|
102 |
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list |
32667 | 103 |
|
55437 | 104 |
|
32667 | 105 |
(* diagnostic display functions *) |
106 |
||
37004 | 107 |
fun print_modes options modes = |
33251 | 108 |
if show_modes options then |
33326
7d0288d90535
improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents:
33268
diff
changeset
|
109 |
tracing ("Inferred modes:\n" ^ |
33251 | 110 |
cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map |
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:
35267
diff
changeset
|
111 |
(fn (p, m) => string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes)) |
33251 | 112 |
else () |
32667 | 113 |
|
37004 | 114 |
fun print_pred_mode_table string_of_entry pred_mode_table = |
32667 | 115 |
let |
50056 | 116 |
fun print_mode pred ((_, mode), entry) = "mode : " ^ string_of_mode mode |
33619
d93a3cb55068
first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents:
33526
diff
changeset
|
117 |
^ string_of_entry pred mode entry |
32667 | 118 |
fun print_pred (pred, modes) = |
119 |
"predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes) |
|
33326
7d0288d90535
improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents:
33268
diff
changeset
|
120 |
val _ = tracing (cat_lines (map print_pred pred_mode_table)) |
32667 | 121 |
in () end; |
122 |
||
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
123 |
fun print_compiled_terms options ctxt = |
33139
9c01ee6f8ee9
added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents:
33138
diff
changeset
|
124 |
if show_compilation options then |
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
125 |
print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term ctxt) |
33139
9c01ee6f8ee9
added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents:
33138
diff
changeset
|
126 |
else K () |
9c01ee6f8ee9
added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents:
33138
diff
changeset
|
127 |
|
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
128 |
fun print_stored_rules ctxt = |
32667 | 129 |
let |
42361 | 130 |
val preds = Graph.keys (PredData.get (Proof_Context.theory_of ctxt)) |
32667 | 131 |
fun print pred () = let |
132 |
val _ = writeln ("predicate: " ^ pred) |
|
133 |
val _ = writeln ("introrules: ") |
|
50056 | 134 |
val _ = fold (fn thm => fn _ => writeln (Display.string_of_thm ctxt thm)) |
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
135 |
(rev (intros_of ctxt pred)) () |
32667 | 136 |
in |
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
137 |
if (has_elim ctxt pred) then |
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
138 |
writeln ("elimrule: " ^ Display.string_of_thm ctxt (the_elim_of ctxt pred)) |
32667 | 139 |
else |
140 |
writeln ("no elimrule defined") |
|
141 |
end |
|
142 |
in |
|
143 |
fold print preds () |
|
144 |
end; |
|
145 |
||
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
146 |
fun print_all_modes compilation ctxt = |
32667 | 147 |
let |
148 |
val _ = writeln ("Inferred modes:") |
|
149 |
fun print (pred, modes) u = |
|
150 |
let |
|
151 |
val _ = writeln ("predicate: " ^ pred) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
152 |
val _ = writeln ("modes: " ^ (commas (map string_of_mode modes))) |
33619
d93a3cb55068
first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents:
33526
diff
changeset
|
153 |
in u end |
32667 | 154 |
in |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
155 |
fold print (all_modes_of compilation ctxt) () |
32667 | 156 |
end |
33129
3085da75ed54
changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents:
33128
diff
changeset
|
157 |
|
33132
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents:
33131
diff
changeset
|
158 |
(* validity checks *) |
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents:
33131
diff
changeset
|
159 |
|
50056 | 160 |
fun check_expected_modes options _ modes = |
55437 | 161 |
(case expected_modes options of |
162 |
SOME (s, ms) => |
|
163 |
(case AList.lookup (op =) modes s of |
|
164 |
SOME modes => |
|
165 |
let |
|
166 |
val modes' = map snd modes |
|
167 |
in |
|
168 |
if not (eq_set eq_mode (ms, modes')) then |
|
169 |
error ("expected modes were not inferred:\n" |
|
170 |
^ " inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes') ^ "\n" |
|
171 |
^ " expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms)) |
|
172 |
else () |
|
173 |
end |
|
174 |
| NONE => ()) |
|
175 |
| NONE => ()) |
|
33752
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
176 |
|
40138 | 177 |
fun check_proposed_modes options preds modes errors = |
55437 | 178 |
map (fn (s, _) => |
179 |
case proposed_modes options s of |
|
180 |
SOME ms => |
|
181 |
(case AList.lookup (op =) modes s of |
|
182 |
SOME inferred_ms => |
|
183 |
let |
|
184 |
val preds_without_modes = map fst (filter (null o snd) modes) |
|
185 |
val modes' = map snd inferred_ms |
|
186 |
in |
|
187 |
if not (eq_set eq_mode (ms, modes')) then |
|
188 |
error ("expected modes were not inferred:\n" |
|
189 |
^ " inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes') ^ "\n" |
|
190 |
^ " expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms) ^ "\n" |
|
191 |
^ (if show_invalid_clauses options then |
|
192 |
("For the following clauses, the following modes could not be inferred: " ^ "\n" |
|
193 |
^ cat_lines errors) else "") ^ |
|
194 |
(if not (null preds_without_modes) then |
|
195 |
"\n" ^ "No mode inferred for the predicates " ^ commas preds_without_modes |
|
196 |
else "")) |
|
197 |
else () |
|
198 |
end |
|
199 |
| NONE => ()) |
|
200 |
| NONE => ()) preds |
|
33132
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents:
33131
diff
changeset
|
201 |
|
39651
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
202 |
fun check_matches_type ctxt predname T ms = |
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
203 |
let |
50056 | 204 |
fun check (Fun (m1, m2)) (Type("fun", [T1,T2])) = check m1 T1 andalso check m2 T2 |
205 |
| check m (Type("fun", _)) = (m = Input orelse m = Output) |
|
39651
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
206 |
| check (Pair (m1, m2)) (Type (@{type_name Product_Type.prod}, [T1, T2])) = |
55437 | 207 |
check m1 T1 andalso check m2 T2 |
50056 | 208 |
| check Input _ = true |
209 |
| check Output _ = true |
|
39651
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
210 |
| check Bool @{typ bool} = true |
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
211 |
| check _ _ = false |
39763
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
212 |
fun check_consistent_modes ms = |
50056 | 213 |
if forall (fn Fun _ => true | _ => false) ms then |
39763
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
214 |
pairself check_consistent_modes (split_list (map (fn Fun (m1, m2) => (m1, m2)) ms)) |
55437 | 215 |
|> (fn (res1, res2) => res1 andalso res2) |
39763
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
216 |
else if forall (fn Input => true | Output => true | Pair _ => true | _ => false) ms then |
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
217 |
true |
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
218 |
else if forall (fn Bool => true | _ => false) ms then |
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
219 |
true |
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
220 |
else |
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
221 |
false |
55437 | 222 |
val _ = |
223 |
map (fn mode => |
|
39763
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
224 |
if length (strip_fun_mode mode) = length (binder_types T) |
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
225 |
andalso (forall (uncurry check) (strip_fun_mode mode ~~ binder_types T)) then () |
55437 | 226 |
else |
227 |
error (string_of_mode mode ^ " is not a valid mode for " ^ |
|
228 |
Syntax.string_of_typ ctxt T ^ " at predicate " ^ predname)) ms |
|
39763
03f95582ef63
weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents:
39762
diff
changeset
|
229 |
val _ = |
55437 | 230 |
if check_consistent_modes ms then () |
231 |
else |
|
232 |
error (commas (map string_of_mode ms) ^ " are inconsistent modes for predicate " ^ predname) |
|
39651
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
233 |
in |
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
234 |
ms |
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
235 |
end |
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
236 |
|
55437 | 237 |
|
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
238 |
(* compilation modifiers *) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
239 |
|
55437 | 240 |
structure Comp_Mod = (* FIXME proper signature *) |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
241 |
struct |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
242 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
243 |
datatype comp_modifiers = Comp_Modifiers of |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
244 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
245 |
compilation : compilation, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
246 |
function_name_prefix : string, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
247 |
compfuns : compilation_funs, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
248 |
mk_random : typ -> term list -> term, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
249 |
modify_funT : typ -> typ, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
250 |
additional_arguments : string list -> term list, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
251 |
wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
252 |
transform_additional_arguments : indprem -> term list -> term list |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
253 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
254 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
255 |
fun dest_comp_modifiers (Comp_Modifiers c) = c |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
256 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
257 |
val compilation = #compilation o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
258 |
val function_name_prefix = #function_name_prefix o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
259 |
val compfuns = #compfuns o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
260 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
261 |
val mk_random = #mk_random o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
262 |
val funT_of' = funT_of o compfuns |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
263 |
val modify_funT = #modify_funT o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
264 |
fun funT_of comp mode = modify_funT comp o funT_of' comp mode |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
265 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
266 |
val additional_arguments = #additional_arguments o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
267 |
val wrap_compilation = #wrap_compilation o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
268 |
val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
269 |
|
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
270 |
fun set_compfuns compfuns' (Comp_Modifiers {compilation, function_name_prefix, compfuns, mk_random, |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
271 |
modify_funT, additional_arguments, wrap_compilation, transform_additional_arguments}) = |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
272 |
(Comp_Modifiers {compilation = compilation, function_name_prefix = function_name_prefix, |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
273 |
compfuns = compfuns', mk_random = mk_random, modify_funT = modify_funT, |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
274 |
additional_arguments = additional_arguments, wrap_compilation = wrap_compilation, |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
275 |
transform_additional_arguments = transform_additional_arguments}) |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
276 |
|
55437 | 277 |
end |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
278 |
|
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
279 |
fun unlimited_compfuns_of true New_Pos_Random_DSeq = |
55437 | 280 |
New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
281 |
| unlimited_compfuns_of false New_Pos_Random_DSeq = |
55437 | 282 |
New_Neg_Random_Sequence_CompFuns.depth_unlimited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
283 |
| unlimited_compfuns_of true Pos_Generator_DSeq = |
55437 | 284 |
New_Pos_DSequence_CompFuns.depth_unlimited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
285 |
| unlimited_compfuns_of false Pos_Generator_DSeq = |
55437 | 286 |
New_Neg_DSequence_CompFuns.depth_unlimited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
287 |
| unlimited_compfuns_of _ c = |
55437 | 288 |
raise Fail ("No unlimited compfuns for compilation " ^ string_of_compilation c) |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
289 |
|
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
290 |
fun limited_compfuns_of true Predicate_Compile_Aux.New_Pos_Random_DSeq = |
55437 | 291 |
New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns |
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
292 |
| limited_compfuns_of false Predicate_Compile_Aux.New_Pos_Random_DSeq = |
55437 | 293 |
New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
294 |
| limited_compfuns_of true Pos_Generator_DSeq = |
55437 | 295 |
New_Pos_DSequence_CompFuns.depth_limited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
296 |
| limited_compfuns_of false Pos_Generator_DSeq = |
55437 | 297 |
New_Neg_DSequence_CompFuns.depth_limited_compfuns |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
298 |
| limited_compfuns_of _ c = |
55437 | 299 |
raise Fail ("No limited compfuns for compilation " ^ string_of_compilation c) |
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
300 |
|
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
301 |
val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
302 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
303 |
compilation = Depth_Limited, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
304 |
function_name_prefix = "depth_limited_", |
45450 | 305 |
compfuns = Predicate_Comp_Funs.compfuns, |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
306 |
mk_random = (fn _ => error "no random generation"), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
307 |
additional_arguments = fn names => |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
308 |
let |
43324
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
43253
diff
changeset
|
309 |
val depth_name = singleton (Name.variant_list names) "depth" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
310 |
in [Free (depth_name, @{typ natural})] end, |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
311 |
modify_funT = (fn T => let val (Ts, U) = strip_type T |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
312 |
val Ts' = [@{typ natural}] in (Ts @ Ts') ---> U end), |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
313 |
wrap_compilation = |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
314 |
fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation => |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
315 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
316 |
val [depth] = additional_arguments |
40139 | 317 |
val (_, Ts) = split_modeT mode (binder_types T) |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
318 |
val T' = mk_monadT compfuns (HOLogic.mk_tupleT Ts) |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
319 |
val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T') |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
320 |
in |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
321 |
if_const $ HOLogic.mk_eq (depth, @{term "0 :: natural"}) |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
322 |
$ mk_empty compfuns (dest_monadT compfuns T') |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
323 |
$ compilation |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
324 |
end, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
325 |
transform_additional_arguments = |
50056 | 326 |
fn _ => fn additional_arguments => |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
327 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
328 |
val [depth] = additional_arguments |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
329 |
val depth' = |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
330 |
Const (@{const_name Groups.minus}, @{typ "natural => natural => natural"}) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
331 |
$ depth $ Const (@{const_name Groups.one}, @{typ "natural"}) |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
332 |
in [depth'] end |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
333 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
334 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
335 |
val random_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
336 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
337 |
compilation = Random, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
338 |
function_name_prefix = "random_", |
45450 | 339 |
compfuns = Predicate_Comp_Funs.compfuns, |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
340 |
mk_random = (fn T => fn additional_arguments => |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
341 |
list_comb (Const(@{const_name Random_Pred.iter}, |
55437 | 342 |
[@{typ natural}, @{typ natural}, @{typ Random.seed}] ---> |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
343 |
Predicate_Comp_Funs.mk_monadT T), additional_arguments)), |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
344 |
modify_funT = (fn T => |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
345 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
346 |
val (Ts, U) = strip_type T |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
347 |
val Ts' = [@{typ natural}, @{typ natural}, @{typ Random.seed}] |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
348 |
in (Ts @ Ts') ---> U end), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
349 |
additional_arguments = (fn names => |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
350 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
351 |
val [nrandom, size, seed] = Name.variant_list names ["nrandom", "size", "seed"] |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
352 |
in |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
353 |
[Free (nrandom, @{typ natural}), Free (size, @{typ natural}), |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
354 |
Free (seed, @{typ Random.seed})] |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
355 |
end), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
356 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
357 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
358 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
359 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
360 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
361 |
val depth_limited_random_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
362 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
363 |
compilation = Depth_Limited_Random, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
364 |
function_name_prefix = "depth_limited_random_", |
45450 | 365 |
compfuns = Predicate_Comp_Funs.compfuns, |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
366 |
mk_random = (fn T => fn additional_arguments => |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
367 |
list_comb (Const(@{const_name Random_Pred.iter}, |
55437 | 368 |
[@{typ natural}, @{typ natural}, @{typ Random.seed}] ---> |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
369 |
Predicate_Comp_Funs.mk_monadT T), tl additional_arguments)), |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
370 |
modify_funT = (fn T => |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
371 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
372 |
val (Ts, U) = strip_type T |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
373 |
val Ts' = [@{typ natural}, @{typ natural}, @{typ natural}, |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
374 |
@{typ Random.seed}] |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
375 |
in (Ts @ Ts') ---> U end), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
376 |
additional_arguments = (fn names => |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
377 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
378 |
val [depth, nrandom, size, seed] = Name.variant_list names ["depth", "nrandom", "size", "seed"] |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
379 |
in |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
380 |
[Free (depth, @{typ natural}), Free (nrandom, @{typ natural}), |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
381 |
Free (size, @{typ natural}), Free (seed, @{typ Random.seed})] |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
382 |
end), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
383 |
wrap_compilation = |
50056 | 384 |
fn compfuns => fn _ => fn T => fn mode => fn additional_arguments => fn compilation => |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
385 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
386 |
val depth = hd (additional_arguments) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
387 |
val (_, Ts) = split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE)) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
388 |
mode (binder_types T) |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
389 |
val T' = mk_monadT compfuns (HOLogic.mk_tupleT Ts) |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
390 |
val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T') |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
391 |
in |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
392 |
if_const $ HOLogic.mk_eq (depth, @{term "0 :: natural"}) |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
393 |
$ mk_empty compfuns (dest_monadT compfuns T') |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
394 |
$ compilation |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
395 |
end, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
396 |
transform_additional_arguments = |
50056 | 397 |
fn _ => fn additional_arguments => |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
398 |
let |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
399 |
val [depth, nrandom, size, seed] = additional_arguments |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
400 |
val depth' = |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
401 |
Const (@{const_name Groups.minus}, @{typ "natural => natural => natural"}) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
402 |
$ depth $ Const (@{const_name Groups.one}, @{typ "natural"}) |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
403 |
in [depth', nrandom, size, seed] end |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
404 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
405 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
406 |
val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
407 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
408 |
compilation = Pred, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
409 |
function_name_prefix = "", |
45450 | 410 |
compfuns = Predicate_Comp_Funs.compfuns, |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
411 |
mk_random = (fn _ => error "no random generation"), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
412 |
modify_funT = I, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
413 |
additional_arguments = K [], |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
414 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
415 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
416 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
417 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
418 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
419 |
val dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
420 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
421 |
compilation = DSeq, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
422 |
function_name_prefix = "dseq_", |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
423 |
compfuns = DSequence_CompFuns.compfuns, |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
424 |
mk_random = (fn _ => error "no random generation in dseq"), |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
425 |
modify_funT = I, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
426 |
additional_arguments = K [], |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
427 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
428 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
429 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
430 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
431 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
432 |
val pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
433 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
434 |
compilation = Pos_Random_DSeq, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
435 |
function_name_prefix = "random_dseq_", |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
436 |
compfuns = Random_Sequence_CompFuns.compfuns, |
50056 | 437 |
mk_random = (fn T => fn _ => |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
438 |
let |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
439 |
val random = Const (@{const_name Quickcheck_Random.random}, |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
440 |
@{typ natural} --> @{typ Random.seed} --> |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
441 |
HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
442 |
in |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
443 |
Const (@{const_name Random_Sequence.Random}, (@{typ natural} --> @{typ Random.seed} --> |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
444 |
HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
445 |
Random_Sequence_CompFuns.mk_random_dseqT T) $ random |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
446 |
end), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
447 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
448 |
modify_funT = I, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
449 |
additional_arguments = K [], |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
450 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
451 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
452 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
453 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
454 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
455 |
val neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
456 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
457 |
compilation = Neg_Random_DSeq, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
458 |
function_name_prefix = "random_dseq_neg_", |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
459 |
compfuns = Random_Sequence_CompFuns.compfuns, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
460 |
mk_random = (fn _ => error "no random generation"), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
461 |
modify_funT = I, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
462 |
additional_arguments = K [], |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
463 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
464 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
465 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
466 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
467 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
468 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
469 |
val new_pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
470 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
471 |
compilation = New_Pos_Random_DSeq, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
472 |
function_name_prefix = "new_random_dseq_", |
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
473 |
compfuns = New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns, |
50056 | 474 |
mk_random = (fn T => fn _ => |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
475 |
let |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
476 |
val random = Const (@{const_name Quickcheck_Random.random}, |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
477 |
@{typ natural} --> @{typ Random.seed} --> |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
478 |
HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
479 |
in |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
480 |
Const (@{const_name Random_Sequence.pos_Random}, (@{typ natural} --> @{typ Random.seed} --> |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
481 |
HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
482 |
New_Pos_Random_Sequence_CompFuns.mk_pos_random_dseqT T) $ random |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
483 |
end), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
484 |
modify_funT = I, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
485 |
additional_arguments = K [], |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
486 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
487 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
488 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
489 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
490 |
|
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
491 |
val new_neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
492 |
{ |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
493 |
compilation = New_Neg_Random_DSeq, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
494 |
function_name_prefix = "new_random_dseq_neg_", |
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
495 |
compfuns = New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns, |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
496 |
mk_random = (fn _ => error "no random generation"), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
497 |
modify_funT = I, |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
498 |
additional_arguments = K [], |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
499 |
wrap_compilation = K (K (K (K (K I)))) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
500 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
501 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
502 |
} |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
503 |
|
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
504 |
val pos_generator_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
505 |
{ |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
506 |
compilation = Pos_Generator_DSeq, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
507 |
function_name_prefix = "generator_dseq_", |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
508 |
compfuns = New_Pos_DSequence_CompFuns.depth_limited_compfuns, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
509 |
mk_random = |
50056 | 510 |
(fn T => fn _ => |
45214 | 511 |
Const (@{const_name "Lazy_Sequence.small_lazy_class.small_lazy"}, |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
512 |
@{typ "natural"} --> Type (@{type_name "Lazy_Sequence.lazy_sequence"}, [T]))), |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
513 |
modify_funT = I, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
514 |
additional_arguments = K [], |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
515 |
wrap_compilation = K (K (K (K (K I)))) |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
516 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
517 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
518 |
} |
55437 | 519 |
|
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
520 |
val neg_generator_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
521 |
{ |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
522 |
compilation = Neg_Generator_DSeq, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
523 |
function_name_prefix = "generator_dseq_neg_", |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
524 |
compfuns = New_Neg_DSequence_CompFuns.depth_limited_compfuns, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
525 |
mk_random = (fn _ => error "no random generation"), |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
526 |
modify_funT = I, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
527 |
additional_arguments = K [], |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
528 |
wrap_compilation = K (K (K (K (K I)))) |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
529 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
530 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
531 |
} |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
532 |
|
45450 | 533 |
val pos_generator_cps_comp_modifiers = Comp_Mod.Comp_Modifiers |
534 |
{ |
|
535 |
compilation = Pos_Generator_CPS, |
|
536 |
function_name_prefix = "generator_cps_pos_", |
|
537 |
compfuns = Pos_Bounded_CPS_Comp_Funs.compfuns, |
|
538 |
mk_random = |
|
50056 | 539 |
(fn T => fn _ => |
45450 | 540 |
Const (@{const_name "Quickcheck_Exhaustive.exhaustive"}, |
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45592
diff
changeset
|
541 |
(T --> @{typ "(bool * term list) option"}) --> |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
542 |
@{typ "natural => (bool * term list) option"})), |
45450 | 543 |
modify_funT = I, |
544 |
additional_arguments = K [], |
|
545 |
wrap_compilation = K (K (K (K (K I)))) |
|
546 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
|
547 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
|
55437 | 548 |
} |
45450 | 549 |
|
550 |
val neg_generator_cps_comp_modifiers = Comp_Mod.Comp_Modifiers |
|
551 |
{ |
|
552 |
compilation = Neg_Generator_CPS, |
|
553 |
function_name_prefix = "generator_cps_neg_", |
|
554 |
compfuns = Neg_Bounded_CPS_Comp_Funs.compfuns, |
|
555 |
mk_random = (fn _ => error "No enumerators"), |
|
556 |
modify_funT = I, |
|
557 |
additional_arguments = K [], |
|
558 |
wrap_compilation = K (K (K (K (K I)))) |
|
559 |
: (compilation_funs -> string -> typ -> mode -> term list -> term -> term), |
|
560 |
transform_additional_arguments = K I : (indprem -> term list -> term list) |
|
561 |
} |
|
55437 | 562 |
|
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
563 |
fun negative_comp_modifiers_of comp_modifiers = |
55437 | 564 |
(case Comp_Mod.compilation comp_modifiers of |
565 |
Pos_Random_DSeq => neg_random_dseq_comp_modifiers |
|
566 |
| Neg_Random_DSeq => pos_random_dseq_comp_modifiers |
|
567 |
| New_Pos_Random_DSeq => new_neg_random_dseq_comp_modifiers |
|
568 |
| New_Neg_Random_DSeq => new_pos_random_dseq_comp_modifiers |
|
569 |
| Pos_Generator_DSeq => neg_generator_dseq_comp_modifiers |
|
570 |
| Neg_Generator_DSeq => pos_generator_dseq_comp_modifiers |
|
571 |
| Pos_Generator_CPS => neg_generator_cps_comp_modifiers |
|
572 |
| Neg_Generator_CPS => pos_generator_cps_comp_modifiers |
|
573 |
| _ => comp_modifiers) |
|
574 |
||
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
575 |
|
32667 | 576 |
(* term construction *) |
577 |
||
55437 | 578 |
fun mk_v (names, vs) s T = |
579 |
(case AList.lookup (op =) vs s of |
|
580 |
NONE => (Free (s, T), (names, (s, [])::vs)) |
|
581 |
| SOME xs => |
|
582 |
let |
|
583 |
val s' = singleton (Name.variant_list names) s; |
|
584 |
val v = Free (s', T) |
|
585 |
in |
|
586 |
(v, (s'::names, AList.update (op =) (s, v::xs) vs)) |
|
587 |
end); |
|
32667 | 588 |
|
589 |
fun distinct_v (Free (s, T)) nvs = mk_v nvs s T |
|
590 |
| distinct_v (t $ u) nvs = |
|
591 |
let |
|
592 |
val (t', nvs') = distinct_v t nvs; |
|
593 |
val (u', nvs'') = distinct_v u nvs'; |
|
594 |
in (t' $ u', nvs'') end |
|
595 |
| distinct_v x nvs = (x, nvs); |
|
596 |
||
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
597 |
|
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
598 |
(** specific rpred functions -- move them to the correct place in this file *) |
50056 | 599 |
fun mk_Eval_of (P as (Free _), T) mode = |
44241 | 600 |
let |
601 |
fun mk_bounds (Type (@{type_name Product_Type.prod}, [T1, T2])) i = |
|
602 |
let |
|
55437 | 603 |
val (bs2, i') = mk_bounds T2 i |
44241 | 604 |
val (bs1, i'') = mk_bounds T1 i' |
605 |
in |
|
606 |
(HOLogic.pair_const T1 T2 $ bs1 $ bs2, i'' + 1) |
|
607 |
end |
|
50056 | 608 |
| mk_bounds _ i = (Bound i, i + 1) |
44241 | 609 |
fun mk_prod ((t1, T1), (t2, T2)) = (HOLogic.pair_const T1 T2 $ t1 $ t2, HOLogic.mk_prodT (T1, T2)) |
610 |
fun mk_tuple [] = (HOLogic.unit, HOLogic.unitT) |
|
611 |
| mk_tuple tTs = foldr1 mk_prod tTs |
|
612 |
fun mk_split_abs (T as Type (@{type_name Product_Type.prod}, [T1, T2])) t = |
|
613 |
absdummy T |
|
614 |
(HOLogic.split_const (T1, T2, @{typ bool}) $ (mk_split_abs T1 (mk_split_abs T2 t))) |
|
615 |
| mk_split_abs T t = absdummy T t |
|
616 |
val args = rev (fst (fold_map mk_bounds (rev (binder_types T)) 0)) |
|
617 |
val (inargs, outargs) = split_mode mode args |
|
50056 | 618 |
val (_, outTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE)) mode (binder_types T) |
45450 | 619 |
val inner_term = Predicate_Comp_Funs.mk_Eval (list_comb (P, inargs), fst (mk_tuple (outargs ~~ outTs))) |
44241 | 620 |
in |
621 |
fold_rev mk_split_abs (binder_types T) inner_term |
|
622 |
end |
|
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
623 |
|
55437 | 624 |
fun compile_arg compilation_modifiers _ _ param_modes arg = |
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
625 |
let |
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
626 |
fun map_params (t as Free (f, T)) = |
55437 | 627 |
(case (AList.lookup (op =) param_modes f) of |
628 |
SOME mode => |
|
629 |
let |
|
630 |
val T' = Comp_Mod.funT_of compilation_modifiers mode T |
|
631 |
in |
|
632 |
mk_Eval_of (Free (f, T'), T) mode |
|
633 |
end |
|
634 |
| NONE => t) |
|
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
635 |
| map_params t = t |
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
636 |
in |
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
637 |
map_aterms map_params arg |
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
638 |
end |
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
639 |
|
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
640 |
fun compile_match compilation_modifiers additional_arguments ctxt param_modes |
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
641 |
eqs eqs' out_ts success_t = |
32667 | 642 |
let |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
643 |
val compfuns = Comp_Mod.compfuns compilation_modifiers |
32667 | 644 |
val eqs'' = maps mk_eq eqs @ eqs' |
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
645 |
val eqs'' = |
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
646 |
map (compile_arg compilation_modifiers additional_arguments ctxt param_modes) eqs'' |
32667 | 647 |
val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) []; |
43324
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
43253
diff
changeset
|
648 |
val name = singleton (Name.variant_list names) "x"; |
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
43253
diff
changeset
|
649 |
val name' = singleton (Name.variant_list (name :: names)) "y"; |
33148
0808f7d0d0d7
removed tuple functions from the predicate compiler
bulwahn
parents:
33147
diff
changeset
|
650 |
val T = HOLogic.mk_tupleT (map fastype_of out_ts); |
32667 | 651 |
val U = fastype_of success_t; |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
652 |
val U' = dest_monadT compfuns U; |
32667 | 653 |
val v = Free (name, T); |
654 |
val v' = Free (name', T); |
|
655 |
in |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
51672
diff
changeset
|
656 |
lambda v (Case_Translation.make_case ctxt Case_Translation.Quiet Name.context v |
33148
0808f7d0d0d7
removed tuple functions from the predicate compiler
bulwahn
parents:
33147
diff
changeset
|
657 |
[(HOLogic.mk_tuple out_ts, |
32667 | 658 |
if null eqs'' then success_t |
659 |
else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $ |
|
660 |
foldr1 HOLogic.mk_conj eqs'' $ success_t $ |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
661 |
mk_empty compfuns U'), |
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
662 |
(v', mk_empty compfuns U')]) |
32667 | 663 |
end; |
664 |
||
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
665 |
fun compile_expr compilation_modifiers ctxt (t, deriv) param_modes additional_arguments = |
32667 | 666 |
let |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
667 |
val compfuns = Comp_Mod.compfuns compilation_modifiers |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
668 |
fun expr_of (t, deriv) = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
669 |
(case (t, deriv) of |
55437 | 670 |
(t, Term Input) => |
671 |
SOME (compile_arg compilation_modifiers additional_arguments ctxt param_modes t) |
|
50056 | 672 |
| (_, Term Output) => NONE |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
673 |
| (Const (name, T), Context mode) => |
55437 | 674 |
(case alternative_compilation_of ctxt name mode of |
675 |
SOME alt_comp => SOME (alt_comp compfuns T) |
|
676 |
| NONE => |
|
677 |
SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers) |
|
678 |
ctxt name mode, |
|
679 |
Comp_Mod.funT_of compilation_modifiers mode T))) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
680 |
| (Free (s, T), Context m) => |
55437 | 681 |
(case (AList.lookup (op =) param_modes s) of |
682 |
SOME _ => SOME (Free (s, Comp_Mod.funT_of compilation_modifiers m T)) |
|
683 |
| NONE => |
|
684 |
let |
|
685 |
val bs = map (pair "x") (binder_types (fastype_of t)) |
|
686 |
val bounds = map Bound (rev (0 upto (length bs) - 1)) |
|
687 |
in SOME (fold_rev Term.abs bs (mk_if compfuns (list_comb (t, bounds)))) end) |
|
50056 | 688 |
| (t, Context _) => |
55437 | 689 |
let |
690 |
val bs = map (pair "x") (binder_types (fastype_of t)) |
|
691 |
val bounds = map Bound (rev (0 upto (length bs) - 1)) |
|
692 |
in SOME (fold_rev Term.abs bs (mk_if compfuns (list_comb (t, bounds)))) end |
|
37391 | 693 |
| (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) => |
55437 | 694 |
(case (expr_of (t1, d1), expr_of (t2, d2)) of |
695 |
(NONE, NONE) => NONE |
|
696 |
| (NONE, SOME t) => SOME t |
|
697 |
| (SOME t, NONE) => SOME t |
|
698 |
| (SOME t1, SOME t2) => SOME (HOLogic.mk_prod (t1, t2))) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
699 |
| (t1 $ t2, Mode_App (deriv1, deriv2)) => |
55437 | 700 |
(case (expr_of (t1, deriv1), expr_of (t2, deriv2)) of |
701 |
(SOME t, NONE) => SOME t |
|
702 |
| (SOME t, SOME u) => SOME (t $ u) |
|
703 |
| _ => error "something went wrong here!")) |
|
32667 | 704 |
in |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
705 |
list_comb (the (expr_of (t, deriv)), additional_arguments) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
706 |
end |
33145 | 707 |
|
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
708 |
fun compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments |
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
709 |
inp (in_ts, out_ts) moded_ps = |
32667 | 710 |
let |
36019
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
711 |
val compfuns = Comp_Mod.compfuns compilation_modifiers |
64bbbd858c39
generalizing the compilation process of the predicate compiler
bulwahn
parents:
36018
diff
changeset
|
712 |
val compile_match = compile_match compilation_modifiers |
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
713 |
additional_arguments ctxt param_modes |
32667 | 714 |
val (in_ts', (all_vs', eqs)) = |
35891
3122bdd95275
contextifying the compilation of the predicate compiler
bulwahn
parents:
35889
diff
changeset
|
715 |
fold_map (collect_non_invertible_subterms ctxt) in_ts (all_vs, []); |
32667 | 716 |
fun compile_prems out_ts' vs names [] = |
717 |
let |
|
718 |
val (out_ts'', (names', eqs')) = |
|
35891
3122bdd95275
contextifying the compilation of the predicate compiler
bulwahn
parents:
35889
diff
changeset
|
719 |
fold_map (collect_non_invertible_subterms ctxt) out_ts' (names, []); |
50056 | 720 |
val (out_ts''', (_, constr_vs)) = fold_map distinct_v |
32667 | 721 |
out_ts'' (names', map (rpair []) vs); |
39762
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents:
39761
diff
changeset
|
722 |
val processed_out_ts = map (compile_arg compilation_modifiers additional_arguments |
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents:
39761
diff
changeset
|
723 |
ctxt param_modes) out_ts |
32667 | 724 |
in |
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
725 |
compile_match constr_vs (eqs @ eqs') out_ts''' |
39762
02fb709ab3cb
handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents:
39761
diff
changeset
|
726 |
(mk_single compfuns (HOLogic.mk_tuple processed_out_ts)) |
32667 | 727 |
end |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
728 |
| compile_prems out_ts vs names ((p, deriv) :: ps) = |
32667 | 729 |
let |
730 |
val vs' = distinct (op =) (flat (vs :: map term_vs out_ts)); |
|
731 |
val (out_ts', (names', eqs)) = |
|
35891
3122bdd95275
contextifying the compilation of the predicate compiler
bulwahn
parents:
35889
diff
changeset
|
732 |
fold_map (collect_non_invertible_subterms ctxt) out_ts (names, []) |
32667 | 733 |
val (out_ts'', (names'', constr_vs')) = fold_map distinct_v |
734 |
out_ts' ((names', map (rpair []) vs)) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
735 |
val mode = head_mode_of deriv |
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
736 |
val additional_arguments' = |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
737 |
Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments |
55437 | 738 |
val (compiled_clause, rest) = |
739 |
(case p of |
|
740 |
Prem t => |
|
741 |
let |
|
742 |
val u = |
|
743 |
compile_expr compilation_modifiers ctxt (t, deriv) |
|
744 |
param_modes additional_arguments' |
|
745 |
val (_, out_ts''') = split_mode mode (snd (strip_comb t)) |
|
746 |
val rest = compile_prems out_ts''' vs' names'' ps |
|
747 |
in |
|
748 |
(u, rest) |
|
749 |
end |
|
750 |
| Negprem t => |
|
751 |
let |
|
752 |
val neg_compilation_modifiers = |
|
753 |
negative_comp_modifiers_of compilation_modifiers |
|
754 |
val u = |
|
755 |
mk_not compfuns |
|
756 |
(compile_expr neg_compilation_modifiers ctxt (t, deriv) |
|
757 |
param_modes additional_arguments') |
|
758 |
val (_, out_ts''') = split_mode mode (snd (strip_comb t)) |
|
759 |
val rest = compile_prems out_ts''' vs' names'' ps |
|
760 |
in |
|
761 |
(u, rest) |
|
762 |
end |
|
763 |
| Sidecond t => |
|
764 |
let |
|
765 |
val t = compile_arg compilation_modifiers additional_arguments |
|
766 |
ctxt param_modes t |
|
767 |
val rest = compile_prems [] vs' names'' ps; |
|
768 |
in |
|
769 |
(mk_if compfuns t, rest) |
|
770 |
end |
|
771 |
| Generator (v, T) => |
|
772 |
let |
|
773 |
val u = Comp_Mod.mk_random compilation_modifiers T additional_arguments |
|
774 |
val rest = compile_prems [Free (v, T)] vs' names'' ps; |
|
775 |
in |
|
776 |
(u, rest) |
|
777 |
end) |
|
32667 | 778 |
in |
33147
180dc60bd88c
improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents:
33146
diff
changeset
|
779 |
compile_match constr_vs' eqs out_ts'' |
32667 | 780 |
(mk_bind compfuns (compiled_clause, rest)) |
781 |
end |
|
55437 | 782 |
val prem_t = compile_prems in_ts' (map fst param_modes) all_vs' moded_ps |
32667 | 783 |
in |
784 |
mk_bind compfuns (mk_single compfuns inp, prem_t) |
|
785 |
end |
|
786 |
||
55437 | 787 |
|
36254 | 788 |
(* switch detection *) |
789 |
||
790 |
(** argument position of an inductive predicates and the executable functions **) |
|
791 |
||
792 |
type position = int * int list |
|
793 |
||
794 |
fun input_positions_pair Input = [[]] |
|
795 |
| input_positions_pair Output = [] |
|
796 |
| input_positions_pair (Fun _) = [] |
|
797 |
| input_positions_pair (Pair (m1, m2)) = |
|
55437 | 798 |
map (cons 1) (input_positions_pair m1) @ map (cons 2) (input_positions_pair m2) |
36254 | 799 |
|
55437 | 800 |
fun input_positions_of_mode mode = |
801 |
flat |
|
802 |
(map_index |
|
803 |
(fn (i, Input) => [(i, [])] |
|
804 |
| (_, Output) => [] |
|
805 |
| (_, Fun _) => [] |
|
806 |
| (i, m as Pair _) => map (pair i) (input_positions_pair m)) |
|
807 |
(Predicate_Compile_Aux.strip_fun_mode mode)) |
|
36254 | 808 |
|
50056 | 809 |
fun argument_position_pair _ [] = [] |
36254 | 810 |
| argument_position_pair (Pair (Fun _, m2)) (2 :: is) = argument_position_pair m2 is |
811 |
| argument_position_pair (Pair (m1, m2)) (i :: is) = |
|
55437 | 812 |
(if eq_mode (m1, Output) andalso i = 2 then |
813 |
argument_position_pair m2 is |
|
814 |
else if eq_mode (m2, Output) andalso i = 1 then |
|
815 |
argument_position_pair m1 is |
|
816 |
else (i :: argument_position_pair (if i = 1 then m1 else m2) is)) |
|
36254 | 817 |
|
818 |
fun argument_position_of mode (i, is) = |
|
819 |
(i - (length (filter (fn Output => true | Fun _ => true | _ => false) |
|
820 |
(List.take (strip_fun_mode mode, i)))), |
|
821 |
argument_position_pair (nth (strip_fun_mode mode) i) is) |
|
822 |
||
823 |
fun nth_pair [] t = t |
|
824 |
| nth_pair (1 :: is) (Const (@{const_name Pair}, _) $ t1 $ _) = nth_pair is t1 |
|
825 |
| nth_pair (2 :: is) (Const (@{const_name Pair}, _) $ _ $ t2) = nth_pair is t2 |
|
826 |
| nth_pair _ _ = raise Fail "unexpected input for nth_tuple" |
|
827 |
||
55437 | 828 |
|
36254 | 829 |
(** switch detection analysis **) |
830 |
||
50056 | 831 |
fun find_switch_test ctxt (i, is) (ts, _) = |
36254 | 832 |
let |
833 |
val t = nth_pair is (nth ts i) |
|
834 |
val T = fastype_of t |
|
835 |
in |
|
55437 | 836 |
(case T of |
36254 | 837 |
TFree _ => NONE |
838 |
| Type (Tcon, _) => |
|
55440 | 839 |
(case Ctr_Sugar.ctr_sugar_of ctxt Tcon of |
55437 | 840 |
NONE => NONE |
55440 | 841 |
| SOME {ctrs, ...} => |
55437 | 842 |
(case strip_comb t of |
843 |
(Var _, []) => NONE |
|
844 |
| (Free _, []) => NONE |
|
55440 | 845 |
| (Const (c, T), _) => |
846 |
if AList.defined (op =) (map_filter (try dest_Const) ctrs) c |
|
847 |
then SOME (c, T) else NONE))) |
|
36254 | 848 |
end |
849 |
||
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
850 |
fun partition_clause ctxt pos moded_clauses = |
36254 | 851 |
let |
852 |
fun insert_list eq (key, value) = AList.map_default eq (key, []) (cons value) |
|
853 |
fun find_switch_test' moded_clause (cases, left) = |
|
55437 | 854 |
(case find_switch_test ctxt pos moded_clause of |
36254 | 855 |
SOME (c, T) => (insert_list (op =) ((c, T), moded_clause) cases, left) |
55437 | 856 |
| NONE => (cases, moded_clause :: left)) |
36254 | 857 |
in |
858 |
fold find_switch_test' moded_clauses ([], []) |
|
859 |
end |
|
860 |
||
861 |
datatype switch_tree = |
|
862 |
Atom of moded_clause list | Node of (position * ((string * typ) * switch_tree) list) * switch_tree |
|
863 |
||
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
864 |
fun mk_switch_tree ctxt mode moded_clauses = |
36254 | 865 |
let |
866 |
fun select_best_switch moded_clauses input_position best_switch = |
|
867 |
let |
|
868 |
val ord = option_ord (rev_order o int_ord o (pairself (length o snd o snd))) |
|
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
869 |
val partition = partition_clause ctxt input_position moded_clauses |
36254 | 870 |
val switch = if (length (fst partition) > 1) then SOME (input_position, partition) else NONE |
871 |
in |
|
55437 | 872 |
(case ord (switch, best_switch) of |
873 |
LESS => best_switch |
|
874 |
| EQUAL => best_switch |
|
875 |
| GREATER => switch) |
|
36254 | 876 |
end |
877 |
fun detect_switches moded_clauses = |
|
55437 | 878 |
(case fold (select_best_switch moded_clauses) (input_positions_of_mode mode) NONE of |
36254 | 879 |
SOME (best_pos, (switched_on, left_clauses)) => |
880 |
Node ((best_pos, map (apsnd detect_switches) switched_on), |
|
881 |
detect_switches left_clauses) |
|
55437 | 882 |
| NONE => Atom moded_clauses) |
36254 | 883 |
in |
884 |
detect_switches moded_clauses |
|
885 |
end |
|
886 |
||
55437 | 887 |
|
36254 | 888 |
(** compilation of detected switches **) |
889 |
||
890 |
fun destruct_constructor_pattern (pat, obj) = |
|
891 |
(case strip_comb pat of |
|
55437 | 892 |
(Free _, []) => cons (pat, obj) |
36254 | 893 |
| (Const (c, T), pat_args) => |
55437 | 894 |
(case strip_comb obj of |
895 |
(Const (c', T'), obj_args) => |
|
896 |
(if c = c' andalso T = T' then |
|
897 |
fold destruct_constructor_pattern (pat_args ~~ obj_args) |
|
898 |
else raise Fail "pattern and object mismatch") |
|
899 |
| _ => raise Fail "unexpected object") |
|
36254 | 900 |
| _ => raise Fail "unexpected pattern") |
901 |
||
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
902 |
fun compile_switch compilation_modifiers ctxt all_vs param_modes additional_arguments mode |
36254 | 903 |
in_ts' outTs switch_tree = |
904 |
let |
|
905 |
val compfuns = Comp_Mod.compfuns compilation_modifiers |
|
42361 | 906 |
val thy = Proof_Context.theory_of ctxt |
36254 | 907 |
fun compile_switch_tree _ _ (Atom []) = NONE |
908 |
| compile_switch_tree all_vs ctxt_eqs (Atom moded_clauses) = |
|
909 |
let |
|
910 |
val in_ts' = map (Pattern.rewrite_term thy ctxt_eqs []) in_ts' |
|
911 |
fun compile_clause' (ts, moded_ps) = |
|
912 |
let |
|
913 |
val (ts, out_ts) = split_mode mode ts |
|
914 |
val subst = fold destruct_constructor_pattern (in_ts' ~~ ts) [] |
|
915 |
val (fsubst, pat') = List.partition (fn (_, Free _) => true | _ => false) subst |
|
916 |
val moded_ps' = (map o apfst o map_indprem) |
|
917 |
(Pattern.rewrite_term thy (map swap fsubst) []) moded_ps |
|
918 |
val inp = HOLogic.mk_tuple (map fst pat') |
|
919 |
val in_ts' = map (Pattern.rewrite_term thy (map swap fsubst) []) (map snd pat') |
|
920 |
val out_ts' = map (Pattern.rewrite_term thy (map swap fsubst) []) out_ts |
|
921 |
in |
|
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
922 |
compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments |
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
923 |
inp (in_ts', out_ts') moded_ps' |
36254 | 924 |
end |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
925 |
in SOME (foldr1 (mk_plus compfuns) (map compile_clause' moded_clauses)) end |
36254 | 926 |
| compile_switch_tree all_vs ctxt_eqs (Node ((position, switched_clauses), left_clauses)) = |
927 |
let |
|
928 |
val (i, is) = argument_position_of mode position |
|
929 |
val inp_var = nth_pair is (nth in_ts' i) |
|
43324
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
43253
diff
changeset
|
930 |
val x = singleton (Name.variant_list all_vs) "x" |
36254 | 931 |
val xt = Free (x, fastype_of inp_var) |
932 |
fun compile_single_case ((c, T), switched) = |
|
933 |
let |
|
934 |
val Ts = binder_types T |
|
935 |
val argnames = Name.variant_list (x :: all_vs) |
|
936 |
(map (fn i => "c" ^ string_of_int i) (1 upto length Ts)) |
|
937 |
val args = map2 (curry Free) argnames Ts |
|
938 |
val pattern = list_comb (Const (c, T), args) |
|
939 |
val ctxt_eqs' = (inp_var, pattern) :: ctxt_eqs |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
940 |
val compilation = the_default (mk_empty compfuns (HOLogic.mk_tupleT outTs)) |
36254 | 941 |
(compile_switch_tree (argnames @ x :: all_vs) ctxt_eqs' switched) |
942 |
in |
|
943 |
(pattern, compilation) |
|
944 |
end |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
51672
diff
changeset
|
945 |
val switch = Case_Translation.make_case ctxt Case_Translation.Quiet Name.context inp_var |
36254 | 946 |
((map compile_single_case switched_clauses) @ |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
947 |
[(xt, mk_empty compfuns (HOLogic.mk_tupleT outTs))]) |
36254 | 948 |
in |
55437 | 949 |
(case compile_switch_tree all_vs ctxt_eqs left_clauses of |
36254 | 950 |
NONE => SOME switch |
55437 | 951 |
| SOME left_comp => SOME (mk_plus compfuns (switch, left_comp))) |
36254 | 952 |
end |
953 |
in |
|
954 |
compile_switch_tree all_vs [] switch_tree |
|
955 |
end |
|
956 |
||
55437 | 957 |
|
36254 | 958 |
(* compilation of predicates *) |
959 |
||
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
960 |
fun compile_pred options compilation_modifiers ctxt all_vs param_vs s T (pol, mode) moded_cls = |
32667 | 961 |
let |
55437 | 962 |
val is_terminating = false (* FIXME: requires an termination analysis *) |
40049
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
963 |
val compilation_modifiers = |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
964 |
(if pol then compilation_modifiers else |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
965 |
negative_comp_modifiers_of compilation_modifiers) |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
966 |
|> (if is_depth_limited_compilation (Comp_Mod.compilation compilation_modifiers) then |
75d9f57123d6
adding decreasing bind and non-decreasing bind; depth-limited and depth-unlimited compilation possible
bulwahn
parents:
39785
diff
changeset
|
967 |
(if is_terminating then |
55437 | 968 |
(Comp_Mod.set_compfuns |
969 |
(unlimited_compfuns_of pol (Comp_Mod.compilation compilation_modifiers))) |
|
970 |
else |
|
971 |
(Comp_Mod.set_compfuns |
|
972 |
(limited_compfuns_of pol (Comp_Mod.compilation compilation_modifiers)))) |
|
973 |
else I) |
|
974 |
val additional_arguments = |
|
975 |
Comp_Mod.additional_arguments compilation_modifiers (all_vs @ param_vs) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
976 |
val compfuns = Comp_Mod.compfuns compilation_modifiers |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
977 |
fun is_param_type (T as Type ("fun",[_ , T'])) = |
55437 | 978 |
is_some (try (dest_monadT compfuns) T) orelse is_param_type T' |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
979 |
| is_param_type T = is_some (try (dest_monadT compfuns) T) |
55437 | 980 |
val (inpTs, outTs) = |
981 |
split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE)) mode |
|
982 |
(binder_types T) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
983 |
val funT = Comp_Mod.funT_of compilation_modifiers mode T |
55437 | 984 |
val (in_ts, _) = |
985 |
fold_map (fold_map_aterms_prodT (curry HOLogic.mk_prod) |
|
986 |
(fn T => fn (param_vs, names) => |
|
987 |
if is_param_type T then |
|
988 |
(Free (hd param_vs, T), (tl param_vs, names)) |
|
989 |
else |
|
990 |
let |
|
991 |
val new = singleton (Name.variant_list names) "x" |
|
992 |
in (Free (new, T), (param_vs, new :: names)) end)) inpTs |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
993 |
(param_vs, (all_vs @ param_vs)) |
55437 | 994 |
val in_ts' = |
995 |
map_filter (map_filter_prod |
|
996 |
(fn t as Free (x, _) => |
|
997 |
if member (op =) param_vs x then NONE else SOME t | t => SOME t)) in_ts |
|
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
998 |
val param_modes = param_vs ~~ ho_arg_modes_of mode |
36254 | 999 |
val compilation = |
1000 |
if detect_switches options then |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
1001 |
the_default (mk_empty compfuns (HOLogic.mk_tupleT outTs)) |
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
1002 |
(compile_switch compilation_modifiers ctxt all_vs param_modes additional_arguments mode |
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
1003 |
in_ts' outTs (mk_switch_tree ctxt mode moded_cls)) |
36254 | 1004 |
else |
1005 |
let |
|
1006 |
val cl_ts = |
|
55437 | 1007 |
map (fn (ts, moded_prems) => |
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
1008 |
compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments |
55437 | 1009 |
(HOLogic.mk_tuple in_ts') (split_mode mode ts) moded_prems) moded_cls |
36254 | 1010 |
in |
1011 |
Comp_Mod.wrap_compilation compilation_modifiers compfuns s T mode additional_arguments |
|
1012 |
(if null cl_ts then |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
1013 |
mk_empty compfuns (HOLogic.mk_tupleT outTs) |
36254 | 1014 |
else |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
1015 |
foldr1 (mk_plus compfuns) cl_ts) |
36254 | 1016 |
end |
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
1017 |
val fun_const = |
55437 | 1018 |
Const (function_name_of (Comp_Mod.compilation compilation_modifiers) ctxt s mode, funT) |
32667 | 1019 |
in |
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
1020 |
HOLogic.mk_Trueprop |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1021 |
(HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation)) |
55437 | 1022 |
end |
1023 |
||
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
1024 |
|
32667 | 1025 |
(* Definition of executable functions and their intro and elim rules *) |
1026 |
||
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55399
diff
changeset
|
1027 |
fun strip_split_abs (Const (@{const_name case_prod}, _) $ t) = strip_split_abs t |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1028 |
| strip_split_abs (Abs (_, _, t)) = strip_split_abs t |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1029 |
| strip_split_abs t = t |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1030 |
|
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37591
diff
changeset
|
1031 |
fun mk_args is_eval (m as Pair (m1, m2), T as Type (@{type_name Product_Type.prod}, [T1, T2])) names = |
55437 | 1032 |
if eq_mode (m, Input) orelse eq_mode (m, Output) then |
1033 |
let |
|
1034 |
val x = singleton (Name.variant_list names) "x" |
|
1035 |
in |
|
1036 |
(Free (x, T), x :: names) |
|
1037 |
end |
|
1038 |
else |
|
1039 |
let |
|
1040 |
val (t1, names') = mk_args is_eval (m1, T1) names |
|
1041 |
val (t2, names'') = mk_args is_eval (m2, T2) names' |
|
1042 |
in |
|
1043 |
(HOLogic.mk_prod (t1, t2), names'') |
|
1044 |
end |
|
1045 |
| mk_args is_eval ((m as Fun _), T) names = |
|
1046 |
let |
|
1047 |
val funT = funT_of Predicate_Comp_Funs.compfuns m T |
|
1048 |
val x = singleton (Name.variant_list names) "x" |
|
1049 |
val (args, _) = fold_map (mk_args is_eval) (strip_fun_mode m ~~ binder_types T) (x :: names) |
|
1050 |
val (inargs, outargs) = split_map_mode (fn _ => fn t => (SOME t, NONE)) m args |
|
1051 |
val t = fold_rev HOLogic.tupled_lambda args (Predicate_Comp_Funs.mk_Eval |
|
1052 |
(list_comb (Free (x, funT), inargs), HOLogic.mk_tuple outargs)) |
|
1053 |
in |
|
1054 |
(if is_eval then t else Free (x, funT), x :: names) |
|
1055 |
end |
|
1056 |
| mk_args _ (_, T) names = |
|
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:
35267
diff
changeset
|
1057 |
let |
43324
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
43253
diff
changeset
|
1058 |
val x = singleton (Name.variant_list names) "x" |
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:
35267
diff
changeset
|
1059 |
in |
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:
35267
diff
changeset
|
1060 |
(Free (x, T), x :: names) |
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:
35267
diff
changeset
|
1061 |
end |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1062 |
|
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1063 |
fun create_intro_elim_rule ctxt mode defthm mode_id funT pred = |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1064 |
let |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1065 |
val funtrm = Const (mode_id, funT) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1066 |
val Ts = binder_types (fastype_of pred) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1067 |
val (args, argnames) = fold_map (mk_args true) (strip_fun_mode mode ~~ Ts) [] |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1068 |
fun strip_eval _ t = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1069 |
let |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1070 |
val t' = strip_split_abs t |
45450 | 1071 |
val (r, _) = Predicate_Comp_Funs.dest_Eval t' |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1072 |
in (SOME (fst (strip_comb r)), NONE) end |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1073 |
val (inargs, outargs) = split_map_mode strip_eval mode args |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1074 |
val eval_hoargs = ho_args_of mode args |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1075 |
val hoargTs = ho_argsT_of mode Ts |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1076 |
val hoarg_names' = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1077 |
Name.variant_list argnames ((map (fn i => "x" ^ string_of_int i)) (1 upto (length hoargTs))) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1078 |
val hoargs' = map2 (curry Free) hoarg_names' hoargTs |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1079 |
val args' = replace_ho_args mode hoargs' args |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1080 |
val predpropI = HOLogic.mk_Trueprop (list_comb (pred, args')) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1081 |
val predpropE = HOLogic.mk_Trueprop (list_comb (pred, args)) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1082 |
val param_eqs = map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) eval_hoargs hoargs' |
45450 | 1083 |
val funpropE = HOLogic.mk_Trueprop (Predicate_Comp_Funs.mk_Eval (list_comb (funtrm, inargs), |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1084 |
if null outargs then Free("y", HOLogic.unitT) else HOLogic.mk_tuple outargs)) |
45450 | 1085 |
val funpropI = HOLogic.mk_Trueprop (Predicate_Comp_Funs.mk_Eval (list_comb (funtrm, inargs), |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1086 |
HOLogic.mk_tuple outargs)) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1087 |
val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI) |
55437 | 1088 |
val simprules = |
1089 |
[defthm, @{thm eval_pred}, |
|
1090 |
@{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}] |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51673
diff
changeset
|
1091 |
val unfolddef_tac = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51673
diff
changeset
|
1092 |
Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simprules) 1 |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1093 |
val introthm = Goal.prove ctxt |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1094 |
(argnames @ hoarg_names' @ ["y"]) [] introtrm (fn _ => unfolddef_tac) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1095 |
val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)); |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1096 |
val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P) |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1097 |
val elimthm = Goal.prove ctxt |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1098 |
(argnames @ ["y", "P"]) [] elimtrm (fn _ => unfolddef_tac) |
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1099 |
val opt_neg_introthm = |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1100 |
if is_all_input mode then |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1101 |
let |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1102 |
val neg_predpropI = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (pred, args'))) |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1103 |
val neg_funpropI = |
45450 | 1104 |
HOLogic.mk_Trueprop (Predicate_Comp_Funs.mk_Eval |
1105 |
(Predicate_Comp_Funs.mk_not (list_comb (funtrm, inargs)), HOLogic.unit)) |
|
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1106 |
val neg_introtrm = Logic.list_implies (neg_predpropI :: param_eqs, neg_funpropI) |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1107 |
val tac = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51673
diff
changeset
|
1108 |
Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps |
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1109 |
(@{thm if_False} :: @{thm Predicate.not_pred_eq} :: simprules)) 1 |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1110 |
THEN rtac @{thm Predicate.singleI} 1 |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1111 |
in SOME (Goal.prove ctxt (argnames @ hoarg_names') [] |
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1112 |
neg_introtrm (fn _ => tac)) |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1113 |
end |
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1114 |
else NONE |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1115 |
in |
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1116 |
((introthm, elimthm), opt_neg_introthm) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1117 |
end |
32667 | 1118 |
|
55437 | 1119 |
fun create_constname_of_mode options thy prefix name _ mode = |
32667 | 1120 |
let |
55437 | 1121 |
val system_proposal = prefix ^ (Long_Name.base_name name) ^ "_" ^ ascii_string_of_mode mode |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1122 |
val name = the_default system_proposal (proposed_names options name mode) |
32667 | 1123 |
in |
33620
b6bf2dc5aed7
added interface of user proposals for names of generated constants
bulwahn
parents:
33619
diff
changeset
|
1124 |
Sign.full_bname thy name |
55437 | 1125 |
end |
32667 | 1126 |
|
33620
b6bf2dc5aed7
added interface of user proposals for names of generated constants
bulwahn
parents:
33619
diff
changeset
|
1127 |
fun create_definitions options preds (name, modes) thy = |
32667 | 1128 |
let |
45450 | 1129 |
val compfuns = Predicate_Comp_Funs.compfuns |
32667 | 1130 |
val T = AList.lookup (op =) preds name |> the |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1131 |
fun create_definition mode thy = |
33752
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1132 |
let |
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1133 |
val mode_cname = create_constname_of_mode options thy "" name T mode |
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1134 |
val mode_cbasename = Long_Name.base_name mode_cname |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1135 |
val funT = funT_of compfuns mode T |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1136 |
val (args, _) = fold_map (mk_args true) ((strip_fun_mode mode) ~~ (binder_types T)) [] |
50056 | 1137 |
fun strip_eval _ t = |
33752
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1138 |
let |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1139 |
val t' = strip_split_abs t |
45450 | 1140 |
val (r, _) = Predicate_Comp_Funs.dest_Eval t' |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1141 |
in (SOME (fst (strip_comb r)), NONE) end |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1142 |
val (inargs, outargs) = split_map_mode strip_eval mode args |
39756
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39685
diff
changeset
|
1143 |
val predterm = fold_rev HOLogic.tupled_lambda inargs |
45450 | 1144 |
(Predicate_Comp_Funs.mk_Enum (HOLogic.tupled_lambda (HOLogic.mk_tuple outargs) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1145 |
(list_comb (Const (name, T), args)))) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1146 |
val lhs = Const (mode_cname, funT) |
33752
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1147 |
val def = Logic.mk_equals (lhs, predterm) |
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1148 |
val ([definition], thy') = thy |> |
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents:
33671
diff
changeset
|
1149 |
Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |> |
46909 | 1150 |
Global_Theory.add_defs false [((Binding.name (Thm.def_name mode_cbasename), def), [])] |
42361 | 1151 |
val ctxt' = Proof_Context.init_global thy' |
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1152 |
val rules as ((intro, elim), _) = |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1153 |
create_intro_elim_rule ctxt' mode definition mode_cname funT (Const (name, T)) |
52788 | 1154 |
in |
1155 |
thy' |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1156 |
|> set_function_name Pred name mode mode_cname |
35884
362bfc2ca0ee
adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents:
35881
diff
changeset
|
1157 |
|> add_predfun_data name mode (definition, rules) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39545
diff
changeset
|
1158 |
|> Global_Theory.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd |
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39545
diff
changeset
|
1159 |
|> Global_Theory.store_thm (Binding.name (mode_cbasename ^ "E"), elim) |> snd |
32667 | 1160 |
end; |
1161 |
in |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1162 |
thy |> defined_function_of Pred name |> fold create_definition modes |
32667 | 1163 |
end; |
1164 |
||
50056 | 1165 |
fun define_functions comp_modifiers _ options preds (name, modes) thy = |
32667 | 1166 |
let |
1167 |
val T = AList.lookup (op =) preds name |> the |
|
1168 |
fun create_definition mode thy = |
|
1169 |
let |
|
33485
0df4f6e46e5e
made definition of functions generically for the different instances
bulwahn
parents:
33484
diff
changeset
|
1170 |
val function_name_prefix = Comp_Mod.function_name_prefix comp_modifiers |
33620
b6bf2dc5aed7
added interface of user proposals for names of generated constants
bulwahn
parents:
33619
diff
changeset
|
1171 |
val mode_cname = create_constname_of_mode options thy function_name_prefix name T mode |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1172 |
val funT = Comp_Mod.funT_of comp_modifiers mode T |
32667 | 1173 |
in |
1174 |
thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)] |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1175 |
|> set_function_name (Comp_Mod.compilation comp_modifiers) name mode mode_cname |
32667 | 1176 |
end; |
1177 |
in |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1178 |
thy |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1179 |
|> defined_function_of (Comp_Mod.compilation comp_modifiers) name |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1180 |
|> fold create_definition modes |
32667 | 1181 |
end; |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32671
diff
changeset
|
1182 |
|
32667 | 1183 |
|
1184 |
(* composition of mode inference, definition, compilation and proof *) |
|
1185 |
||
1186 |
(** auxillary combinators for table of preds and modes **) |
|
1187 |
||
1188 |
fun map_preds_modes f preds_modes_table = |
|
1189 |
map (fn (pred, modes) => |
|
1190 |
(pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table |
|
1191 |
||
1192 |
fun join_preds_modes table1 table2 = |
|
1193 |
map_preds_modes (fn pred => fn mode => fn value => |
|
1194 |
(value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1 |
|
36254 | 1195 |
|
32667 | 1196 |
fun maps_modes preds_modes_table = |
1197 |
map (fn (pred, modes) => |
|
50056 | 1198 |
(pred, map (fn (_, value) => value) modes)) preds_modes_table |
36254 | 1199 |
|
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
1200 |
fun compile_preds options comp_modifiers ctxt all_vs param_vs preds moded_clauses = |
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
1201 |
map_preds_modes (fn pred => compile_pred options comp_modifiers ctxt all_vs param_vs pred |
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
1202 |
(the (AList.lookup (op =) preds pred))) moded_clauses |
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
1203 |
|
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:
35267
diff
changeset
|
1204 |
fun prove options thy clauses preds moded_clauses compiled_terms = |
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:
35267
diff
changeset
|
1205 |
map_preds_modes (prove_pred options thy clauses preds) |
32667 | 1206 |
(join_preds_modes moded_clauses compiled_terms) |
1207 |
||
50056 | 1208 |
fun prove_by_skip _ thy _ _ _ compiled_terms = |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
34974
diff
changeset
|
1209 |
map_preds_modes |
50056 | 1210 |
(fn _ => fn _ => fn t => Drule.export_without_context (Skip_Proof.make_thm thy t)) |
32667 | 1211 |
compiled_terms |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1212 |
|
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1213 |
(* preparation of introduction rules into special datastructures *) |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1214 |
fun dest_prem ctxt params t = |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1215 |
(case strip_comb t of |
50056 | 1216 |
(v as Free _, _) => if member (op =) params v then Prem t else Sidecond t |
55437 | 1217 |
| (c as Const (@{const_name Not}, _), [t]) => |
1218 |
(case dest_prem ctxt params t of |
|
1219 |
Prem t => Negprem t |
|
1220 |
| Negprem _ => error ("Double negation not allowed in premise: " ^ |
|
1221 |
Syntax.string_of_term ctxt (c $ t)) |
|
1222 |
| Sidecond t => Sidecond (c $ t)) |
|
50056 | 1223 |
| (Const (s, _), _) => |
55437 | 1224 |
if is_registered ctxt s then Prem t else Sidecond t |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1225 |
| _ => Sidecond t) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1226 |
|
39310
17ef4415b17c
removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents:
39299
diff
changeset
|
1227 |
fun prepare_intrs options ctxt prednames intros = |
32667 | 1228 |
let |
42361 | 1229 |
val thy = Proof_Context.theory_of ctxt |
33126
bb8806eb5da7
importing of polymorphic introduction rules with different schematic variable names
bulwahn
parents:
33124
diff
changeset
|
1230 |
val intrs = map prop_of intros |
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
1231 |
val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames |
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
1232 |
val (preds, intrs) = unify_consts thy preds intrs |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1233 |
val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] ctxt |
33126
bb8806eb5da7
importing of polymorphic introduction rules with different schematic variable names
bulwahn
parents:
33124
diff
changeset
|
1234 |
val preds = map dest_Const preds |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1235 |
val all_vs = terms_vs intrs |
39201
234e6ef4ff67
using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents:
39194
diff
changeset
|
1236 |
fun generate_modes s T = |
234e6ef4ff67
using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents:
39194
diff
changeset
|
1237 |
if member (op =) (no_higher_order_predicate options) s then |
234e6ef4ff67
using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents:
39194
diff
changeset
|
1238 |
all_smodes_of_typ T |
234e6ef4ff67
using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents:
39194
diff
changeset
|
1239 |
else |
234e6ef4ff67
using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents:
39194
diff
changeset
|
1240 |
all_modes_of_typ T |
55437 | 1241 |
val all_modes = |
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:
35267
diff
changeset
|
1242 |
map (fn (s, T) => |
55437 | 1243 |
(s, |
1244 |
(case proposed_modes options s of |
|
39651
2e06dad03dd3
adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents:
39649
diff
changeset
|
1245 |
SOME ms => check_matches_type ctxt s T ms |
55437 | 1246 |
| NONE => generate_modes s T))) preds |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1247 |
val params = |
55437 | 1248 |
(case intrs of |
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
1249 |
[] => |
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
1250 |
let |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1251 |
val T = snd (hd preds) |
39764
1cf2088cf035
only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents:
39763
diff
changeset
|
1252 |
val one_mode = hd (the (AList.lookup (op =) all_modes (fst (hd preds)))) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1253 |
val paramTs = |
39764
1cf2088cf035
only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents:
39763
diff
changeset
|
1254 |
ho_argsT_of one_mode (binder_types T) |
33482
5029ec373036
strictly respecting the line margin in the predicate compiler core
bulwahn
parents:
33479
diff
changeset
|
1255 |
val param_names = Name.variant_list [] (map (fn i => "p" ^ string_of_int i) |
5029ec373036
strictly respecting the line margin in the predicate compiler core
bulwahn
parents:
33479
diff
changeset
|
1256 |
(1 upto length paramTs)) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1257 |
in |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1258 |
map2 (curry Free) param_names paramTs |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1259 |
end |
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:
35267
diff
changeset
|
1260 |
| (intr :: _) => |
55437 | 1261 |
let |
1262 |
val (p, args) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)) |
|
1263 |
val one_mode = hd (the (AList.lookup (op =) all_modes (fst (dest_Const p)))) |
|
1264 |
in |
|
1265 |
ho_args_of one_mode args |
|
1266 |
end) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1267 |
val param_vs = map (fst o dest_Free) params |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1268 |
fun add_clause intr clauses = |
32667 | 1269 |
let |
55437 | 1270 |
val (Const (name, _), ts) = |
1271 |
strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)) |
|
1272 |
val prems = |
|
1273 |
map (dest_prem ctxt params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr) |
|
32667 | 1274 |
in |
55437 | 1275 |
AList.update op = |
1276 |
(name, these (AList.lookup op = clauses name) @ [(ts, prems)]) |
|
1277 |
clauses |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1278 |
end; |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1279 |
val clauses = fold add_clause intrs [] |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1280 |
in |
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:
35267
diff
changeset
|
1281 |
(preds, all_vs, param_vs, all_modes, clauses) |
55437 | 1282 |
end |
32667 | 1283 |
|
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1284 |
(* sanity check of introduction rules *) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1285 |
(* TODO: rethink check with new modes *) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1286 |
(* |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1287 |
fun check_format_of_intro_rule thy intro = |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1288 |
let |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1289 |
val concl = Logic.strip_imp_concl (prop_of intro) |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1290 |
val (p, args) = strip_comb (HOLogic.dest_Trueprop concl) |
33629
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1291 |
val params = fst (chop (nparams_of thy (fst (dest_Const p))) args) |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1292 |
fun check_arg arg = case HOLogic.strip_tupleT (fastype_of arg) of |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1293 |
(Ts as _ :: _ :: _) => |
33629
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1294 |
if length (HOLogic.strip_tuple arg) = length Ts then |
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1295 |
true |
33114
4785ef554dcc
added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents:
33113
diff
changeset
|
1296 |
else |
33629
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1297 |
error ("Format of introduction rule is invalid: tuples must be expanded:" |
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1298 |
^ (Syntax.string_of_term_global thy arg) ^ " in " ^ |
55437 | 1299 |
(Display.string_of_thm_global thy intro)) |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1300 |
| _ => true |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1301 |
val prems = Logic.strip_imp_prems (prop_of intro) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1302 |
fun check_prem (Prem t) = forall check_arg args |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1303 |
| check_prem (Negprem t) = forall check_arg args |
33106
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1304 |
| check_prem _ = true |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1305 |
in |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1306 |
forall check_arg args andalso |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1307 |
forall (check_prem o dest_prem thy params o HOLogic.dest_Trueprop) prems |
7a1636c3ffc9
extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents:
32740
diff
changeset
|
1308 |
end |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1309 |
*) |
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1310 |
(* |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1311 |
fun check_intros_elim_match thy prednames = |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1312 |
let |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1313 |
fun check predname = |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1314 |
let |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1315 |
val intros = intros_of thy predname |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1316 |
val elim = the_elim_of thy predname |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1317 |
val nparams = nparams_of thy predname |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1318 |
val elim' = |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
34974
diff
changeset
|
1319 |
(Drule.export_without_context o Skip_Proof.make_thm thy) |
42361 | 1320 |
(mk_casesrule (Proof_Context.init_global thy) nparams intros) |
33124
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1321 |
in |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1322 |
if not (Thm.equiv_thm (elim, elim')) then |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1323 |
error "Introduction and elimination rules do not match!" |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1324 |
else true |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1325 |
end |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1326 |
in forall check prednames end |
5378e61add1a
continued cleaning up; moved tuple expanding to core
bulwahn
parents:
33123
diff
changeset
|
1327 |
*) |
33113 | 1328 |
|
55437 | 1329 |
|
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1330 |
(* create code equation *) |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1331 |
|
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
1332 |
fun add_code_equations ctxt preds result_thmss = |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1333 |
let |
33629
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1334 |
fun add_code_equation (predname, T) (pred, result_thms) = |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1335 |
let |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1336 |
val full_mode = fold_rev (curry Fun) (map (K Input) (binder_types T)) Bool |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1337 |
in |
47329
b9e115d4c5da
improved equality check for modes in predicate compiler
bulwahn
parents:
47248
diff
changeset
|
1338 |
if member eq_mode (modes_of Pred ctxt predname) full_mode then |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1339 |
let |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1340 |
val Ts = binder_types T |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1341 |
val arg_names = Name.variant_list [] |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1342 |
(map (fn i => "x" ^ string_of_int i) (1 upto length Ts)) |
33629
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1343 |
val args = map2 (curry Free) arg_names Ts |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1344 |
val predfun = Const (function_name_of Pred ctxt predname full_mode, |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
1345 |
Ts ---> Predicate_Comp_Funs.mk_monadT @{typ unit}) |
33754
f2957bd46faf
adding derived constant Predicate.holds to Predicate theory; adopting the predicate compiler
bulwahn
parents:
33753
diff
changeset
|
1346 |
val rhs = @{term Predicate.holds} $ (list_comb (predfun, args)) |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1347 |
val eq_term = HOLogic.mk_Trueprop |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1348 |
(HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs)) |
35888
d902054e7ac6
contextifying the proof procedure in the predicate compiler
bulwahn
parents:
35887
diff
changeset
|
1349 |
val def = predfun_definition_of ctxt predname full_mode |
33441
99a5f22a967d
eliminated funny record patterns and made SML/NJ happy;
wenzelm
parents:
33377
diff
changeset
|
1350 |
val tac = fn _ => Simplifier.simp_tac |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51673
diff
changeset
|
1351 |
(put_simpset HOL_basic_ss ctxt addsimps [def, @{thm holds_eq}, @{thm eval_pred}]) 1 |
35888
d902054e7ac6
contextifying the proof procedure in the predicate compiler
bulwahn
parents:
35887
diff
changeset
|
1352 |
val eq = Goal.prove ctxt arg_names [] eq_term tac |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1353 |
in |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1354 |
(pred, result_thms @ [eq]) |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1355 |
end |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1356 |
else |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1357 |
(pred, result_thms) |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1358 |
end |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1359 |
in |
33629
5f35cf91c6a4
improving code quality thanks to Florian's code review
bulwahn
parents:
33628
diff
changeset
|
1360 |
map2 add_code_equation preds result_thmss |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1361 |
end |
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1362 |
|
55437 | 1363 |
|
32667 | 1364 |
(** main function of predicate compiler **) |
1365 |
||
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1366 |
datatype steps = Steps of |
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1367 |
{ |
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:
35267
diff
changeset
|
1368 |
define_functions : options -> (string * typ) list -> string * (bool * mode) list -> theory -> theory, |
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:
35267
diff
changeset
|
1369 |
prove : options -> theory -> (string * (term list * indprem list) list) list -> (string * typ) list |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1370 |
-> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table, |
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
1371 |
add_code_equations : Proof.context -> (string * typ) list |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1372 |
-> (string * thm list) list -> (string * thm list) list, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1373 |
comp_modifiers : Comp_Mod.comp_modifiers, |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1374 |
use_generators : bool, |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1375 |
qname : bstring |
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1376 |
} |
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1377 |
|
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:
35267
diff
changeset
|
1378 |
fun add_equations_of steps mode_analysis_options options prednames thy = |
32667 | 1379 |
let |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1380 |
fun dest_steps (Steps s) = s |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1381 |
val compilation = Comp_Mod.compilation (#comp_modifiers (dest_steps steps)) |
42361 | 1382 |
val ctxt = Proof_Context.init_global thy |
55437 | 1383 |
val _ = |
1384 |
print_step options |
|
1385 |
("Starting predicate compiler (compilation: " ^ string_of_compilation compilation ^ |
|
1386 |
") for predicates " ^ commas prednames ^ "...") |
|
1387 |
(*val _ = check_intros_elim_match thy prednames*) |
|
1388 |
(*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*) |
|
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:
35267
diff
changeset
|
1389 |
val _ = |
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:
35267
diff
changeset
|
1390 |
if show_intermediate_results options then |
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
1391 |
tracing (commas (map (Display.string_of_thm ctxt) (maps (intros_of ctxt) prednames))) |
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:
35267
diff
changeset
|
1392 |
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:
35267
diff
changeset
|
1393 |
val (preds, all_vs, param_vs, all_modes, clauses) = |
39310
17ef4415b17c
removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents:
39299
diff
changeset
|
1394 |
prepare_intrs options ctxt prednames (maps (intros_of ctxt) prednames) |
33123
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
bulwahn
parents:
33120
diff
changeset
|
1395 |
val _ = print_step options "Infering modes..." |
39273
92aa2a0f7399
refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents:
39201
diff
changeset
|
1396 |
val (lookup_mode, lookup_neg_mode, needs_random) = (modes_of compilation ctxt, |
92aa2a0f7399
refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents:
39201
diff
changeset
|
1397 |
modes_of (negative_compilation_of compilation) ctxt, needs_random ctxt) |
92aa2a0f7399
refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents:
39201
diff
changeset
|
1398 |
val ((moded_clauses, needs_random), errors) = |
42088
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
bulwahn
parents:
42011
diff
changeset
|
1399 |
cond_timeit (Config.get ctxt Quickcheck.timing) "Infering modes" |
36251
5fd5d732a4ea
only add relevant predicates to the list of extra modes
bulwahn
parents:
36247
diff
changeset
|
1400 |
(fn _ => infer_modes mode_analysis_options |
39273
92aa2a0f7399
refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents:
39201
diff
changeset
|
1401 |
options (lookup_mode, lookup_neg_mode, needs_random) ctxt preds all_modes param_vs clauses) |
92aa2a0f7399
refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents:
39201
diff
changeset
|
1402 |
val thy' = fold (fn (s, ms) => set_needs_random s ms) needs_random thy |
33132
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents:
33131
diff
changeset
|
1403 |
val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses |
40138 | 1404 |
val _ = check_expected_modes options preds modes |
1405 |
val _ = check_proposed_modes options preds modes errors |
|
37004 | 1406 |
val _ = print_modes options modes |
33123
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
bulwahn
parents:
33120
diff
changeset
|
1407 |
val _ = print_step options "Defining executable functions..." |
36252 | 1408 |
val thy'' = |
42088
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
bulwahn
parents:
42011
diff
changeset
|
1409 |
cond_timeit (Config.get ctxt Quickcheck.timing) "Defining executable functions..." |
52788 | 1410 |
(fn _ => fold (#define_functions (dest_steps steps) options preds) modes thy') |
42361 | 1411 |
val ctxt'' = Proof_Context.init_global thy'' |
33123
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
bulwahn
parents:
33120
diff
changeset
|
1412 |
val _ = print_step options "Compiling equations..." |
32667 | 1413 |
val compiled_terms = |
42088
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
bulwahn
parents:
42011
diff
changeset
|
1414 |
cond_timeit (Config.get ctxt Quickcheck.timing) "Compiling equations...." (fn _ => |
36254 | 1415 |
compile_preds options |
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
1416 |
(#comp_modifiers (dest_steps steps)) ctxt'' all_vs param_vs preds moded_clauses) |
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
37004
diff
changeset
|
1417 |
val _ = print_compiled_terms options ctxt'' compiled_terms |
33123
3c7c4372f9ad
cleaned up debugging messages; added options to code_pred command
bulwahn
parents:
33120
diff
changeset
|
1418 |
val _ = print_step options "Proving equations..." |
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:
35267
diff
changeset
|
1419 |
val result_thms = |
42088
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
bulwahn
parents:
42011
diff
changeset
|
1420 |
cond_timeit (Config.get ctxt Quickcheck.timing) "Proving equations...." (fn _ => |
36252 | 1421 |
#prove (dest_steps steps) options thy'' clauses preds moded_clauses compiled_terms) |
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
1422 |
val result_thms' = #add_code_equations (dest_steps steps) ctxt'' preds |
33376
5cb663aa48ee
predicate compiler creates code equations for predicates with full mode
bulwahn
parents:
33375
diff
changeset
|
1423 |
(maps_modes result_thms) |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1424 |
val qname = #qname (dest_steps steps) |
47248 | 1425 |
val attrib = Thm.declaration_attribute (fn thm => Context.mapping (Code.add_eqn thm) I) |
36252 | 1426 |
val thy''' = |
42088
8d00484551fe
making quickcheck's result value more formal; allowing more result information to be returned after timeout; adding output of timing information in quickcheck
bulwahn
parents:
42011
diff
changeset
|
1427 |
cond_timeit (Config.get ctxt Quickcheck.timing) "Setting code equations...." (fn _ => |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39545
diff
changeset
|
1428 |
fold (fn (name, result_thms) => fn thy => snd (Global_Theory.add_thmss |
32667 | 1429 |
[((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms), |
47248 | 1430 |
[attrib])] thy)) |
52788 | 1431 |
result_thms' thy'') |
32667 | 1432 |
in |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1433 |
thy''' |
32667 | 1434 |
end |
55437 | 1435 |
|
33132
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents:
33131
diff
changeset
|
1436 |
fun gen_add_equations steps options names thy = |
32667 | 1437 |
let |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1438 |
fun dest_steps (Steps s) = s |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1439 |
val defined = defined_functions (Comp_Mod.compilation (#comp_modifiers (dest_steps steps))) |
52788 | 1440 |
val thy' = extend_intro_graph names thy; |
32667 | 1441 |
fun strong_conn_of gr keys = |
46614
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
wenzelm
parents:
46219
diff
changeset
|
1442 |
Graph.strong_conn (Graph.restrict (member (op =) (Graph.all_succs gr keys)) gr) |
32667 | 1443 |
val scc = strong_conn_of (PredData.get thy') names |
40053
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
bulwahn
parents:
40052
diff
changeset
|
1444 |
val thy'' = fold preprocess_intros (flat scc) thy' |
39657
5e57675b7e40
moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents:
39653
diff
changeset
|
1445 |
val thy''' = fold_rev |
32667 | 1446 |
(fn preds => fn thy => |
42361 | 1447 |
if not (forall (defined (Proof_Context.init_global thy)) preds) 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:
35267
diff
changeset
|
1448 |
let |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1449 |
val mode_analysis_options = {use_generators = #use_generators (dest_steps steps), |
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:
35267
diff
changeset
|
1450 |
reorder_premises = |
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:
35267
diff
changeset
|
1451 |
not (no_topmost_reordering options andalso not (null (inter (op =) preds names))), |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1452 |
infer_pos_and_neg_modes = #use_generators (dest_steps steps)} |
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:
35267
diff
changeset
|
1453 |
in |
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:
35267
diff
changeset
|
1454 |
add_equations_of steps mode_analysis_options options preds 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:
35267
diff
changeset
|
1455 |
end |
33483
a15a2f7f7560
improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents:
33482
diff
changeset
|
1456 |
else thy) |
52788 | 1457 |
scc thy'' |
39657
5e57675b7e40
moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents:
39653
diff
changeset
|
1458 |
in thy''' end |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1459 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1460 |
val add_equations = gen_add_equations |
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:
35267
diff
changeset
|
1461 |
(Steps { |
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:
35267
diff
changeset
|
1462 |
define_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:
35267
diff
changeset
|
1463 |
fn options => fn preds => fn (s, modes) => |
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:
35267
diff
changeset
|
1464 |
create_definitions |
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:
35267
diff
changeset
|
1465 |
options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes), |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1466 |
prove = prove, |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1467 |
add_code_equations = add_code_equations, |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1468 |
comp_modifiers = predicate_comp_modifiers, |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1469 |
use_generators = false, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1470 |
qname = "equation"}) |
33143
730a2e8a6ec6
modularized the compilation in the predicate compiler
bulwahn
parents:
33141
diff
changeset
|
1471 |
|
33134
88c9c3460fe7
renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents:
33133
diff
changeset
|
1472 |
val add_depth_limited_equations = gen_add_equations |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1473 |
(Steps { |
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1474 |
define_functions = |
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1475 |
fn options => fn preds => fn (s, modes) => |
45450 | 1476 |
define_functions depth_limited_comp_modifiers Predicate_Comp_Funs.compfuns |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1477 |
options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes), |
32667 | 1478 |
prove = prove_by_skip, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1479 |
add_code_equations = K (K I), |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1480 |
comp_modifiers = depth_limited_comp_modifiers, |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1481 |
use_generators = false, |
33330
d6eb7f19bfc6
encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents:
33328
diff
changeset
|
1482 |
qname = "depth_limited_equation"}) |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1483 |
|
35880
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1484 |
val add_random_equations = gen_add_equations |
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1485 |
(Steps { |
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1486 |
define_functions = |
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1487 |
fn options => fn preds => fn (s, modes) => |
45450 | 1488 |
define_functions random_comp_modifiers Predicate_Comp_Funs.compfuns options preds |
35880
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1489 |
(s, map_filter (fn (true, m) => SOME m | _ => NONE) modes), |
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1490 |
comp_modifiers = random_comp_modifiers, |
32667 | 1491 |
prove = prove_by_skip, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1492 |
add_code_equations = K (K I), |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1493 |
use_generators = true, |
33375 | 1494 |
qname = "random_equation"}) |
35880
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1495 |
|
35881
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1496 |
val add_depth_limited_random_equations = gen_add_equations |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1497 |
(Steps { |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1498 |
define_functions = |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1499 |
fn options => fn preds => fn (s, modes) => |
45450 | 1500 |
define_functions depth_limited_random_comp_modifiers Predicate_Comp_Funs.compfuns options preds |
35881
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1501 |
(s, map_filter (fn (true, m) => SOME m | _ => NONE) modes), |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1502 |
comp_modifiers = depth_limited_random_comp_modifiers, |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1503 |
prove = prove_by_skip, |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1504 |
add_code_equations = K (K I), |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1505 |
use_generators = true, |
35881
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1506 |
qname = "depth_limited_random_equation"}) |
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1507 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1508 |
val add_dseq_equations = gen_add_equations |
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:
35267
diff
changeset
|
1509 |
(Steps { |
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:
35267
diff
changeset
|
1510 |
define_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:
35267
diff
changeset
|
1511 |
fn options => fn preds => fn (s, modes) => |
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:
35267
diff
changeset
|
1512 |
define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns |
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:
35267
diff
changeset
|
1513 |
options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes), |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1514 |
prove = prove_by_skip, |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1515 |
add_code_equations = K (K I), |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1516 |
comp_modifiers = dseq_comp_modifiers, |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1517 |
use_generators = false, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1518 |
qname = "dseq_equation"}) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1519 |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1520 |
val add_random_dseq_equations = gen_add_equations |
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:
35267
diff
changeset
|
1521 |
(Steps { |
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:
35267
diff
changeset
|
1522 |
define_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:
35267
diff
changeset
|
1523 |
fn options => fn preds => fn (s, modes) => |
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:
35267
diff
changeset
|
1524 |
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:
35267
diff
changeset
|
1525 |
val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes |
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:
35267
diff
changeset
|
1526 |
val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes |
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:
35267
diff
changeset
|
1527 |
in define_functions pos_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns |
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:
35267
diff
changeset
|
1528 |
options preds (s, pos_modes) |
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:
35267
diff
changeset
|
1529 |
#> define_functions neg_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns |
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:
35267
diff
changeset
|
1530 |
options preds (s, neg_modes) |
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:
35267
diff
changeset
|
1531 |
end, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1532 |
prove = prove_by_skip, |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1533 |
add_code_equations = K (K I), |
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:
35267
diff
changeset
|
1534 |
comp_modifiers = pos_random_dseq_comp_modifiers, |
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1535 |
use_generators = true, |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1536 |
qname = "random_dseq_equation"}) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1537 |
|
36018 | 1538 |
val add_new_random_dseq_equations = gen_add_equations |
1539 |
(Steps { |
|
1540 |
define_functions = |
|
1541 |
fn options => fn preds => fn (s, modes) => |
|
55437 | 1542 |
let |
1543 |
val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes |
|
1544 |
val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes |
|
1545 |
in |
|
1546 |
define_functions new_pos_random_dseq_comp_modifiers |
|
1547 |
New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns |
|
1548 |
options preds (s, pos_modes) #> |
|
1549 |
define_functions new_neg_random_dseq_comp_modifiers |
|
1550 |
New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns |
|
1551 |
options preds (s, neg_modes) |
|
1552 |
end, |
|
36018 | 1553 |
prove = prove_by_skip, |
1554 |
add_code_equations = K (K I), |
|
1555 |
comp_modifiers = new_pos_random_dseq_comp_modifiers, |
|
39761
c2a76ec6e2d9
renaming use_random to use_generators in the predicate compiler
bulwahn
parents:
39760
diff
changeset
|
1556 |
use_generators = true, |
36018 | 1557 |
qname = "new_random_dseq_equation"}) |
32667 | 1558 |
|
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1559 |
val add_generator_dseq_equations = gen_add_equations |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1560 |
(Steps { |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1561 |
define_functions = |
55437 | 1562 |
fn options => fn preds => fn (s, modes) => |
1563 |
let |
|
1564 |
val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes |
|
1565 |
val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes |
|
1566 |
in |
|
1567 |
define_functions pos_generator_dseq_comp_modifiers |
|
1568 |
New_Pos_DSequence_CompFuns.depth_limited_compfuns options preds (s, pos_modes) #> |
|
1569 |
define_functions neg_generator_dseq_comp_modifiers |
|
1570 |
New_Neg_DSequence_CompFuns.depth_limited_compfuns options preds (s, neg_modes) |
|
1571 |
end, |
|
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1572 |
prove = prove_by_skip, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1573 |
add_code_equations = K (K I), |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1574 |
comp_modifiers = pos_generator_dseq_comp_modifiers, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1575 |
use_generators = true, |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1576 |
qname = "generator_dseq_equation"}) |
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1577 |
|
45450 | 1578 |
val add_generator_cps_equations = gen_add_equations |
1579 |
(Steps { |
|
1580 |
define_functions = |
|
55437 | 1581 |
fn options => fn preds => fn (s, modes) => |
1582 |
let |
|
1583 |
val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes |
|
1584 |
val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes |
|
1585 |
in |
|
1586 |
define_functions pos_generator_cps_comp_modifiers Pos_Bounded_CPS_Comp_Funs.compfuns |
|
1587 |
options preds (s, pos_modes) |
|
1588 |
#> define_functions neg_generator_cps_comp_modifiers Neg_Bounded_CPS_Comp_Funs.compfuns |
|
1589 |
options preds (s, neg_modes) |
|
1590 |
end, |
|
45450 | 1591 |
prove = prove_by_skip, |
1592 |
add_code_equations = K (K I), |
|
1593 |
comp_modifiers = pos_generator_cps_comp_modifiers, |
|
1594 |
use_generators = true, |
|
1595 |
qname = "generator_cps_equation"}) |
|
55437 | 1596 |
|
1597 |
||
32667 | 1598 |
(** user interface **) |
1599 |
||
1600 |
(* code_pred_intro attribute *) |
|
1601 |
||
39545
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1602 |
fun attrib' f opt_case_name = |
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1603 |
Thm.declaration_attribute (fn thm => Context.mapping (f (opt_case_name, thm)) I); |
32667 | 1604 |
|
39545
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1605 |
val code_pred_intro_attrib = attrib' add_intro NONE; |
32667 | 1606 |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
1607 |
(*FIXME |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
1608 |
- Naming of auxiliary rules necessary? |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
1609 |
*) |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
1610 |
|
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1611 |
(* values_timeout configuration *) |
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1612 |
|
45452
414732ebf891
increasing values_timeout to avoid failures of isatest with HOL-IMP
bulwahn
parents:
45450
diff
changeset
|
1613 |
val default_values_timeout = if ML_System.is_smlnj then 1200.0 else 40.0 |
43896
8955dcac6c71
values_timeout defaults to 600.0 on SML/NJ -- saves us from cluttering all theories equivalent declarations
krauss
parents:
43619
diff
changeset
|
1614 |
|
55437 | 1615 |
val values_timeout = |
1616 |
Attrib.setup_config_real @{binding values_timeout} (K default_values_timeout) |
|
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1617 |
|
55437 | 1618 |
val setup = |
1619 |
PredData.put (Graph.empty) #> |
|
39545
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1620 |
Attrib.setup @{binding code_pred_intro} (Scan.lift (Scan.option Args.name) >> attrib' add_intro) |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
1621 |
"adding alternative introduction rules for code generation of inductive predicates" |
32667 | 1622 |
|
53380
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1623 |
fun qualified_binding a = |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1624 |
Binding.qualify true (Long_Name.qualifier a) (Binding.name (Long_Name.base_name a)); |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1625 |
|
33522 | 1626 |
(* TODO: make Theory_Data to Generic_Data & remove duplication of local theory and theory *) |
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:
38558
diff
changeset
|
1627 |
(* FIXME ... this is important to avoid changing the background theory below *) |
33132
07efd452a698
moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents:
33131
diff
changeset
|
1628 |
fun generic_code_pred prep_const options raw_const lthy = |
32667 | 1629 |
let |
42361 | 1630 |
val thy = Proof_Context.theory_of lthy |
32667 | 1631 |
val const = prep_const thy raw_const |
40053
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
bulwahn
parents:
40052
diff
changeset
|
1632 |
val lthy' = Local_Theory.background_theory (extend_intro_graph [const]) lthy |
42361 | 1633 |
val thy' = Proof_Context.theory_of lthy' |
1634 |
val ctxt' = Proof_Context.init_global thy' |
|
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
1635 |
val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim ctxt') |
32667 | 1636 |
fun mk_cases const = |
1637 |
let |
|
39299
e98a06145530
directly computing the values of interest instead of composing functions in an unintelligent way that causes exponential much garbage; using the latest theory
bulwahn
parents:
39273
diff
changeset
|
1638 |
val T = Sign.the_const_type thy' const |
33146
bf852ef586f2
now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents:
33145
diff
changeset
|
1639 |
val pred = Const (const, T) |
37007
116670499530
changing operations for accessing data to work with contexts
bulwahn
parents:
37006
diff
changeset
|
1640 |
val intros = intros_of ctxt' const |
55437 | 1641 |
in mk_casesrule lthy' pred intros end |
32667 | 1642 |
val cases_rules = map mk_cases preds |
1643 |
val cases = |
|
39545
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1644 |
map2 (fn pred_name => fn case_rule => Rule_Cases.Case {fixes = [], |
53380
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1645 |
assumes = |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1646 |
(Binding.name "that", tl (Logic.strip_imp_prems case_rule)) :: |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1647 |
map_filter (fn (fact_name, fact) => |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1648 |
Option.map (fn a => (qualified_binding a, [fact])) fact_name) |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1649 |
((SOME (Long_Name.base_name pred_name ^ ".prems") :: names_of ctxt' pred_name) ~~ |
08f3491c50bf
cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents:
53378
diff
changeset
|
1650 |
Logic.strip_imp_prems case_rule), |
39545
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1651 |
binds = [], cases = []}) preds cases_rules |
32667 | 1652 |
val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases |
1653 |
val lthy'' = lthy' |
|
39545
631cf48c7894
code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents:
39477
diff
changeset
|
1654 |
|> fold Variable.auto_fixes cases_rules |
53378
07990ba8c0ea
cases: more position information and PIDE markup;
wenzelm
parents:
52788
diff
changeset
|
1655 |
|> Proof_Context.update_cases true case_env |
32667 | 1656 |
fun after_qed thms goal_ctxt = |
1657 |
let |
|
42361 | 1658 |
val global_thms = Proof_Context.export goal_ctxt |
1659 |
(Proof_Context.init_global (Proof_Context.theory_of goal_ctxt)) (map the_single thms) |
|
32667 | 1660 |
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:
38558
diff
changeset
|
1661 |
goal_ctxt |> Local_Theory.background_theory (fold set_elim global_thms #> |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1662 |
((case compilation options of |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1663 |
Pred => add_equations |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1664 |
| DSeq => add_dseq_equations |
35879
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1665 |
| Pos_Random_DSeq => add_random_dseq_equations |
99818df5b8f5
reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents:
35845
diff
changeset
|
1666 |
| Depth_Limited => add_depth_limited_equations |
35880
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1667 |
| Random => add_random_equations |
35881
aa412e08bfee
adding depth_limited_random compilation to predicate compiler
bulwahn
parents:
35880
diff
changeset
|
1668 |
| Depth_Limited_Random => add_depth_limited_random_equations |
36018 | 1669 |
| New_Pos_Random_DSeq => add_new_random_dseq_equations |
40051
b6acda4d1c29
added generator_dseq compilation for a sound depth-limited compilation with small value generators
bulwahn
parents:
40050
diff
changeset
|
1670 |
| Pos_Generator_DSeq => add_generator_dseq_equations |
45450 | 1671 |
| Pos_Generator_CPS => add_generator_cps_equations |
50056 | 1672 |
| _ => error ("Compilation not supported") |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1673 |
) options [const])) |
33144
1831516747d4
removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents:
33143
diff
changeset
|
1674 |
end |
32667 | 1675 |
in |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36261
diff
changeset
|
1676 |
Proof.theorem NONE after_qed (map (single o (rpair [])) cases_rules) lthy'' |
32667 | 1677 |
end; |
1678 |
||
1679 |
val code_pred = generic_code_pred (K I); |
|
1680 |
val code_pred_cmd = generic_code_pred Code.read_const |
|
1681 |
||
55437 | 1682 |
|
32667 | 1683 |
(* transformation for code generation *) |
1684 |
||
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1685 |
(* FIXME just one data slot (record) per program unit *) |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1686 |
|
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1687 |
structure Pred_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1688 |
( |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1689 |
type T = unit -> term Predicate.pred |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1690 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1691 |
fun init _ () = error "Pred_Result" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1692 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1693 |
val put_pred_result = Pred_Result.put; |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1694 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1695 |
structure Pred_Random_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1696 |
( |
50057 | 1697 |
type T = unit -> seed -> term Predicate.pred * seed |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1698 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1699 |
fun init _ () = error "Pred_Random_Result" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1700 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1701 |
val put_pred_random_result = Pred_Random_Result.put; |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1702 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1703 |
structure Dseq_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1704 |
( |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
1705 |
type T = unit -> term Limited_Sequence.dseq |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1706 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1707 |
fun init _ () = error "Dseq_Result" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1708 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1709 |
val put_dseq_result = Dseq_Result.put; |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1710 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1711 |
structure Dseq_Random_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1712 |
( |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1713 |
type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term Limited_Sequence.dseq * seed |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1714 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1715 |
fun init _ () = error "Dseq_Random_Result" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1716 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1717 |
val put_dseq_random_result = Dseq_Random_Result.put; |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1718 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1719 |
structure New_Dseq_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1720 |
( |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1721 |
type T = unit -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1722 |
(* FIXME avoid user error with non-user text *) |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1723 |
fun init _ () = error "New_Dseq_Random_Result" |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1724 |
); |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1725 |
val put_new_dseq_result = New_Dseq_Result.put; |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1726 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1727 |
structure Lseq_Random_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1728 |
( |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1729 |
type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1730 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1731 |
fun init _ () = error "Lseq_Random_Result" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1732 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1733 |
val put_lseq_random_result = Lseq_Random_Result.put; |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1734 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1735 |
structure Lseq_Random_Stats_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1736 |
( |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1737 |
type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> (term * Code_Numeral.natural) Lazy_Sequence.lazy_sequence |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40845
diff
changeset
|
1738 |
(* FIXME avoid user error with non-user text *) |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1739 |
fun init _ () = error "Lseq_Random_Stats_Result" |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1740 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1741 |
val put_lseq_random_stats_result = Lseq_Random_Stats_Result.put; |
32667 | 1742 |
|
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1743 |
fun dest_special_compr t = |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1744 |
let |
55437 | 1745 |
val (inner_t, T_compr) = |
1746 |
(case t of |
|
1747 |
(Const (@{const_name Collect}, _) $ Abs (_, T, t)) => (t, T) |
|
1748 |
| _ => raise TERM ("dest_special_compr", [t])) |
|
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1749 |
val (Ts, conj) = apfst (map snd) (Predicate_Compile_Aux.strip_ex inner_t) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1750 |
val [eq, body] = HOLogic.dest_conj conj |
55437 | 1751 |
val rhs = |
1752 |
(case HOLogic.dest_eq eq of |
|
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1753 |
(Bound i, rhs) => if i = length Ts then rhs else raise TERM ("dest_special_compr", [t]) |
55437 | 1754 |
| _ => raise TERM ("dest_special_compr", [t])) |
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1755 |
val output_names = Name.variant_list (fold Term.add_free_names [rhs, body] []) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1756 |
(map (fn i => "x" ^ string_of_int i) (1 upto length Ts)) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1757 |
val output_frees = map2 (curry Free) output_names (rev Ts) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1758 |
val body = subst_bounds (output_frees, body) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1759 |
val output = subst_bounds (output_frees, rhs) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1760 |
in |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1761 |
(((body, output), T_compr), output_names) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1762 |
end |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1763 |
|
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1764 |
fun dest_general_compr ctxt t_compr = |
55437 | 1765 |
let |
1766 |
val inner_t = |
|
1767 |
(case t_compr of |
|
1768 |
(Const (@{const_name Collect}, _) $ t) => t |
|
1769 |
| _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr)) |
|
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1770 |
val (body, Ts, fp) = HOLogic.strip_psplits inner_t; |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1771 |
val output_names = Name.variant_list (Term.add_free_names body []) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1772 |
(map (fn i => "x" ^ string_of_int i) (1 upto length Ts)) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1773 |
val output_frees = map2 (curry Free) output_names (rev Ts) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1774 |
val body = subst_bounds (output_frees, body) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1775 |
val T_compr = HOLogic.mk_ptupleT fp Ts |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1776 |
val output = HOLogic.mk_ptuple fp T_compr (rev output_frees) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1777 |
in |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1778 |
(((body, output), T_compr), output_names) |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1779 |
end |
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1780 |
|
32667 | 1781 |
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*) |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1782 |
fun analyze_compr ctxt (comp_modifiers, additional_arguments) param_user_modes |
50056 | 1783 |
(compilation, _) t_compr = |
32667 | 1784 |
let |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1785 |
val compfuns = Comp_Mod.compfuns comp_modifiers |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1786 |
val all_modes_of = all_modes_of compilation |
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1787 |
val (((body, output), T_compr), output_names) = |
55437 | 1788 |
(case try dest_special_compr t_compr of |
1789 |
SOME r => r |
|
1790 |
| NONE => dest_general_compr ctxt t_compr) |
|
50056 | 1791 |
val (Const (name, _), all_args) = |
55437 | 1792 |
(case strip_comb body of |
36031
199fe16cdaab
returning an more understandable user error message in the values command
bulwahn
parents:
36030
diff
changeset
|
1793 |
(Const (name, T), all_args) => (Const (name, T), all_args) |
55437 | 1794 |
| (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1795 |
in |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1796 |
if defined_functions compilation ctxt name then |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32667
diff
changeset
|
1797 |
let |
55437 | 1798 |
fun extract_mode (Const (@{const_name Pair}, _) $ t1 $ t2) = |
1799 |
Pair (extract_mode t1, extract_mode t2) |
|
1800 |
| extract_mode (Free (x, _)) = |
|
1801 |
if member (op =) output_names x then Output else Input |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1802 |
| extract_mode _ = Input |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1803 |
val user_mode = fold_rev (curry Fun) (map extract_mode all_args) Bool |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1804 |
fun valid modes1 modes2 = |
55437 | 1805 |
(case int_ord (length modes1, length modes2) of |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1806 |
GREATER => error "Not enough mode annotations" |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1807 |
| LESS => error "Too many mode annotations" |
55437 | 1808 |
| EQUAL => |
1809 |
forall (fn (_, NONE) => true | (m, SOME m2) => eq_mode (m, m2)) (modes1 ~~ modes2)) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1810 |
fun mode_instance_of (m1, m2) = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1811 |
let |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1812 |
fun instance_of (Fun _, Input) = true |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1813 |
| instance_of (Input, Input) = true |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1814 |
| instance_of (Output, Output) = true |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1815 |
| instance_of (Pair (m1, m2), Pair (m1', m2')) = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1816 |
instance_of (m1, m1') andalso instance_of (m2, m2') |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1817 |
| instance_of (Pair (m1, m2), Input) = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1818 |
instance_of (m1, Input) andalso instance_of (m2, Input) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1819 |
| instance_of (Pair (m1, m2), Output) = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1820 |
instance_of (m1, Output) andalso instance_of (m2, Output) |
37002
34e73e8bab66
improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents:
37001
diff
changeset
|
1821 |
| instance_of (Input, Pair (m1, m2)) = |
34e73e8bab66
improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents:
37001
diff
changeset
|
1822 |
instance_of (Input, m1) andalso instance_of (Input, m2) |
34e73e8bab66
improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents:
37001
diff
changeset
|
1823 |
| instance_of (Output, Pair (m1, m2)) = |
34e73e8bab66
improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents:
37001
diff
changeset
|
1824 |
instance_of (Output, m1) andalso instance_of (Output, m2) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1825 |
| instance_of _ = false |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1826 |
in forall instance_of (strip_fun_mode m1 ~~ strip_fun_mode m2) end |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1827 |
val derivs = all_derivations_of ctxt (all_modes_of ctxt) [] body |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1828 |
|> filter (fn (d, missing_vars) => |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1829 |
let |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1830 |
val (p_mode :: modes) = collect_context_modes d |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1831 |
in |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1832 |
null missing_vars andalso |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1833 |
mode_instance_of (p_mode, user_mode) andalso |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1834 |
the_default true (Option.map (valid modes) param_user_modes) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1835 |
end) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1836 |
|> map fst |
55437 | 1837 |
val deriv = |
1838 |
(case derivs of |
|
1839 |
[] => |
|
1840 |
error ("No mode possible for comprehension " ^ Syntax.string_of_term ctxt t_compr) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1841 |
| [d] => d |
55437 | 1842 |
| d :: _ :: _ => |
1843 |
(warning ("Multiple modes possible for comprehension " ^ |
|
1844 |
Syntax.string_of_term ctxt t_compr); d)) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1845 |
val (_, outargs) = split_mode (head_mode_of deriv) all_args |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1846 |
val t_pred = compile_expr comp_modifiers ctxt |
39648
655307cb8489
rewriting function mk_Eval_of in predicate compiler
bulwahn
parents:
39557
diff
changeset
|
1847 |
(body, deriv) [] additional_arguments; |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45452
diff
changeset
|
1848 |
val T_pred = dest_monadT compfuns (fastype_of t_pred) |
42094
e6867e9c6d10
allowing special set comprehensions in values command; adding an example for special set comprehension in values
bulwahn
parents:
42088
diff
changeset
|
1849 |
val arrange = HOLogic.tupled_lambda (HOLogic.mk_tuple outargs) output |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1850 |
in |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1851 |
if null outargs then t_pred else mk_map compfuns T_pred T_compr arrange t_pred |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1852 |
end |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1853 |
else |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1854 |
error "Evaluation with values is not possible because compilation with code_pred was not invoked" |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1855 |
end |
32667 | 1856 |
|
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1857 |
fun eval ctxt stats param_user_modes (options as (compilation, arguments)) k t_compr = |
32667 | 1858 |
let |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1859 |
fun count xs x = |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1860 |
let |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1861 |
fun count' i [] = i |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1862 |
| count' i (x' :: xs) = if x = x' then count' (i + 1) xs else count' i xs |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1863 |
in count' 0 xs end |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1864 |
fun accumulate xs = (map (fn x => (x, count xs x)) o sort int_ord o distinct (op =)) xs; |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1865 |
val comp_modifiers = |
55437 | 1866 |
(case compilation of |
1867 |
Pred => predicate_comp_modifiers |
|
1868 |
| Random => random_comp_modifiers |
|
1869 |
| Depth_Limited => depth_limited_comp_modifiers |
|
1870 |
| Depth_Limited_Random => depth_limited_random_comp_modifiers |
|
1871 |
(*| Annotated => annotated_comp_modifiers*) |
|
1872 |
| DSeq => dseq_comp_modifiers |
|
1873 |
| Pos_Random_DSeq => pos_random_dseq_comp_modifiers |
|
1874 |
| New_Pos_Random_DSeq => new_pos_random_dseq_comp_modifiers |
|
1875 |
| Pos_Generator_DSeq => pos_generator_dseq_comp_modifiers) |
|
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1876 |
val compfuns = Comp_Mod.compfuns comp_modifiers |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1877 |
val additional_arguments = |
55437 | 1878 |
(case compilation of |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1879 |
Pred => [] |
55437 | 1880 |
| Random => |
1881 |
map (HOLogic.mk_number @{typ "natural"}) arguments @ |
|
1882 |
[@{term "(1, 1) :: natural * natural"}] |
|
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1883 |
| Annotated => [] |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1884 |
| Depth_Limited => [HOLogic.mk_number @{typ "natural"} (hd arguments)] |
55437 | 1885 |
| Depth_Limited_Random => |
1886 |
map (HOLogic.mk_number @{typ "natural"}) arguments @ |
|
1887 |
[@{term "(1, 1) :: natural * natural"}] |
|
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1888 |
| DSeq => [] |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1889 |
| Pos_Random_DSeq => [] |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1890 |
| New_Pos_Random_DSeq => [] |
55437 | 1891 |
| Pos_Generator_DSeq => []) |
1892 |
val t = |
|
1893 |
analyze_compr ctxt (comp_modifiers, additional_arguments) param_user_modes options t_compr |
|
1894 |
val T = dest_monadT compfuns (fastype_of t) |
|
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1895 |
val t' = |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1896 |
if stats andalso compilation = New_Pos_Random_DSeq then |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1897 |
mk_map compfuns T (HOLogic.mk_prodT (HOLogic.termT, @{typ natural})) |
44241 | 1898 |
(absdummy T (HOLogic.mk_prod (HOLogic.term_of_const T $ Bound 0, |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1899 |
@{term natural_of_nat} $ (HOLogic.size_const T $ Bound 0)))) t |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1900 |
else |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1901 |
mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t |
42361 | 1902 |
val thy = Proof_Context.theory_of ctxt |
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1903 |
val time_limit = seconds (Config.get ctxt values_timeout) |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1904 |
val (ts, statistics) = |
40135 | 1905 |
(case compilation of |
35880
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1906 |
(* Random => |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1907 |
fst (Predicate.yieldn k |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1908 |
(Code_Eval.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref) |
33137
0d16c07f8d24
added option to generate random values to values command in the predicate compiler
bulwahn
parents:
33135
diff
changeset
|
1909 |
(fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' [] |
35880
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1910 |
|> Random_Engine.run))*) |
2623b23e41fc
a new simpler random compilation for the predicate compiler
bulwahn
parents:
35879
diff
changeset
|
1911 |
Pos_Random_DSeq => |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1912 |
let |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1913 |
val [nrandom, size, depth] = map Code_Numeral.natural_of_integer arguments |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1914 |
in |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
1915 |
rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Limited_Sequence.yieldn k |
39471 | 1916 |
(Code_Runtime.dynamic_value_strict (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result") |
55757 | 1917 |
ctxt NONE (fn proc => fn g => fn nrandom => fn size => fn s => g nrandom size s |>> Limited_Sequence.map proc) |
39471 | 1918 |
t' [] nrandom size |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1919 |
|> Random_Engine.run) |
40135 | 1920 |
depth true)) ()) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1921 |
end |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1922 |
| DSeq => |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
1923 |
rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Limited_Sequence.yieldn k |
39471 | 1924 |
(Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result") |
55757 | 1925 |
ctxt NONE Limited_Sequence.map t' []) (Code_Numeral.natural_of_integer (the_single arguments)) true)) ()) |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1926 |
| Pos_Generator_DSeq => |
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1927 |
let |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1928 |
val depth = Code_Numeral.natural_of_integer (the_single arguments) |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1929 |
in |
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1930 |
rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Lazy_Sequence.yieldn k |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1931 |
(Code_Runtime.dynamic_value_strict (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Core.put_new_dseq_result") |
55757 | 1932 |
ctxt NONE (fn proc => fn g => fn depth => g depth |> Lazy_Sequence.map proc) |
40135 | 1933 |
t' [] depth))) ()) |
40102
a9e4e94b3022
restructuring values command and adding generator compilation
bulwahn
parents:
40053
diff
changeset
|
1934 |
end |
36020
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents:
36019
diff
changeset
|
1935 |
| New_Pos_Random_DSeq => |
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents:
36019
diff
changeset
|
1936 |
let |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1937 |
val [nrandom, size, depth] = map Code_Numeral.natural_of_integer arguments |
36020
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents:
36019
diff
changeset
|
1938 |
val seed = Random_Engine.next_seed () |
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents:
36019
diff
changeset
|
1939 |
in |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1940 |
if stats then |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1941 |
apsnd (SOME o accumulate o map Code_Numeral.integer_of_natural) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
1942 |
(split_list (TimeLimit.timeLimit time_limit |
40135 | 1943 |
(fn () => fst (Lazy_Sequence.yieldn k |
39471 | 1944 |
(Code_Runtime.dynamic_value_strict |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1945 |
(Lseq_Random_Stats_Result.get, put_lseq_random_stats_result, "Predicate_Compile_Core.put_lseq_random_stats_result") |
55757 | 1946 |
ctxt NONE |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1947 |
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
1948 |
|> Lazy_Sequence.map (apfst proc)) |
40135 | 1949 |
t' [] nrandom size seed depth))) ())) |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1950 |
else rpair NONE |
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1951 |
(TimeLimit.timeLimit time_limit (fn () => fst (Lazy_Sequence.yieldn k |
39471 | 1952 |
(Code_Runtime.dynamic_value_strict |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39310
diff
changeset
|
1953 |
(Lseq_Random_Result.get, put_lseq_random_result, "Predicate_Compile_Core.put_lseq_random_result") |
55757 | 1954 |
ctxt NONE |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1955 |
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
50057
diff
changeset
|
1956 |
|> Lazy_Sequence.map proc) |
40135 | 1957 |
t' [] nrandom size seed depth))) ()) |
36020
3ee4c29ead7f
adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents:
36019
diff
changeset
|
1958 |
end |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1959 |
| _ => |
42141
2c255ba8f299
changing values_timeout from a static option to a dynamic option as required for testing with SML/NJ
bulwahn
parents:
42094
diff
changeset
|
1960 |
rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Predicate.yieldn k |
39471 | 1961 |
(Code_Runtime.dynamic_value_strict (Pred_Result.get, put_pred_result, "Predicate_Compile_Core.put_pred_result") |
55757 | 1962 |
ctxt NONE Predicate.map t' []))) ())) |
40135 | 1963 |
handle TimeLimit.TimeOut => error "Reached timeout during execution of values" |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1964 |
in ((T, ts), statistics) end; |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1965 |
|
43123
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
1966 |
fun values param_user_modes ((raw_expected, stats), comp_options) k t_compr ctxt = |
32667 | 1967 |
let |
37003
a393a588b82e
moving towards working with proof contexts in the predicate compiler
bulwahn
parents:
37002
diff
changeset
|
1968 |
val ((T, ts), statistics) = eval ctxt stats param_user_modes comp_options k t_compr |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1969 |
val setT = HOLogic.mk_setT T |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1970 |
val elems = HOLogic.mk_set T ts |
43123
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
1971 |
val ([dots], ctxt') = |
55437 | 1972 |
Proof_Context.add_fixes [(@{binding dots}, SOME setT, Mixfix ("...", [], 1000))] ctxt |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1973 |
(* check expected values *) |
43123
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
1974 |
val union = Const (@{const_abbrev Set.union}, setT --> setT --> setT) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1975 |
val () = |
55437 | 1976 |
(case raw_expected of |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1977 |
NONE => () |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1978 |
| SOME s => |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1979 |
if eq_set (op =) (HOLogic.dest_set (Syntax.read_term ctxt s), ts) then () |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1980 |
else |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1981 |
error ("expected and computed values do not match:\n" ^ |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1982 |
"expected values: " ^ Syntax.string_of_term ctxt (Syntax.read_term ctxt s) ^ "\n" ^ |
55437 | 1983 |
"computed values: " ^ Syntax.string_of_term ctxt elems ^ "\n")) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1984 |
in |
55437 | 1985 |
((if k = ~1 orelse length ts < k then elems else union $ elems $ Free (dots, setT), statistics), |
1986 |
ctxt') |
|
32667 | 1987 |
end; |
33623
4ec42d38224f
changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents:
33620
diff
changeset
|
1988 |
|
33479
428ddcc16e7b
added optional mode annotations for parameters in the values command
bulwahn
parents:
33478
diff
changeset
|
1989 |
fun values_cmd print_modes param_user_modes options k raw_t state = |
32667 | 1990 |
let |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1991 |
val ctxt = Toplevel.context_of state |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1992 |
val t = Syntax.read_term ctxt raw_t |
43123
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
1993 |
val ((t', stats), ctxt') = values param_user_modes options k t ctxt |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
1994 |
val ty' = Term.type_of t' |
43123
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
1995 |
val ctxt'' = Variable.auto_fixes t' ctxt' |
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
1996 |
val pretty_stat = |
55437 | 1997 |
(case stats of |
1998 |
NONE => [] |
|
1999 |
| SOME xs => |
|
36027
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2000 |
let |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2001 |
val total = fold (curry (op +)) (map snd xs) 0 |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2002 |
fun pretty_entry (s, n) = |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2003 |
[Pretty.str "size", Pretty.brk 1, |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2004 |
Pretty.str (string_of_int s), Pretty.str ":", Pretty.brk 1, |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2005 |
Pretty.str (string_of_int n), Pretty.fbrk] |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2006 |
in |
29a15da9c63d
added statistics to values command for random generation
bulwahn
parents:
36025
diff
changeset
|
2007 |
[Pretty.fbrk, Pretty.str "Statistics:", Pretty.fbrk, |
55437 | 2008 |
Pretty.str "total:", Pretty.brk 1, Pretty.str (string_of_int total), Pretty.fbrk] @ |
2009 |
maps pretty_entry xs |
|
2010 |
end) |
|
2011 |
in |
|
2012 |
Print_Mode.with_modes print_modes (fn () => |
|
43123
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
2013 |
Pretty.block ([Pretty.quote (Syntax.pretty_term ctxt'' t'), Pretty.fbrk, |
28e6351b2f8e
creating a free variable with proper name and local mixfix syntax (cf. db9b9e46131c)
bulwahn
parents:
42616
diff
changeset
|
2014 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt'' ty')] |
55437 | 2015 |
@ pretty_stat)) () |
2016 |
end |> Pretty.writeln |
|
32667 | 2017 |
|
55437 | 2018 |
end |