author | wenzelm |
Mon, 27 Mar 2023 21:48:47 +0200 | |
changeset 77723 | b761c91c2447 |
parent 74561 | 8e6c973003c8 |
child 80636 | 4041e7c8059d |
permissions | -rw-r--r-- |
33265 | 1 |
(* Title: HOL/Tools/Predicate_Compile/predicate_compile_data.ML |
2 |
Author: Lukas Bulwahn, TU Muenchen |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
3 |
|
33265 | 4 |
Book-keeping datastructure for the predicate compiler. |
5 |
*) |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
6 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
7 |
signature PREDICATE_COMPILE_DATA = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
8 |
sig |
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
|
9 |
val ignore_consts : string 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
|
10 |
val keep_functions : string 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
|
11 |
val keep_function : theory -> string -> bool |
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
|
12 |
val processed_specs : theory -> string -> (string * thm list) list option |
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
|
13 |
val store_processed_specs : (string * (string * thm list) list) -> theory -> theory |
55437 | 14 |
|
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
|
15 |
val get_specification : Predicate_Compile_Aux.options -> theory -> term -> thm list |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
16 |
val obtain_specification_graph : |
35404 | 17 |
Predicate_Compile_Aux.options -> theory -> term -> thm list Term_Graph.T |
55437 | 18 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
19 |
val normalize_equation : theory -> thm -> thm |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
20 |
end; |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
21 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
22 |
structure Predicate_Compile_Data : PREDICATE_COMPILE_DATA = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
23 |
struct |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
24 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
25 |
open Predicate_Compile_Aux; |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
26 |
|
33522 | 27 |
structure Data = Theory_Data |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
28 |
( |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
29 |
type T = |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
30 |
{ignore_consts : Symset.T, |
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
31 |
keep_functions : Symset.T, |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35267
diff
changeset
|
32 |
processed_specs : ((string * thm list) list) Symtab.table}; |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
33 |
val empty = |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
34 |
{ignore_consts = Symset.empty, |
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
35 |
keep_functions = Symset.empty, |
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
|
36 |
processed_specs = Symtab.empty}; |
33522 | 37 |
fun merge |
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
|
38 |
({ignore_consts = c1, keep_functions = k1, processed_specs = s1}, |
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
|
39 |
{ignore_consts = c2, keep_functions = k2, processed_specs = s2}) = |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
40 |
{ignore_consts = Symset.merge (c1, c2), |
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
41 |
keep_functions = Symset.merge (k1, k2), |
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
|
42 |
processed_specs = Symtab.merge (K true) (s1, s2)} |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
43 |
); |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
44 |
|
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
|
45 |
|
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
|
46 |
|
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
|
47 |
fun mk_data (c, k, s) = {ignore_consts = c, keep_functions = k, processed_specs = s} |
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
|
48 |
fun map_data f {ignore_consts = c, keep_functions = k, processed_specs = s} = mk_data (f (c, k, s)) |
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
|
49 |
|
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
57962
diff
changeset
|
50 |
fun ignore_consts cs = |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
51 |
Data.map (map_data (@{apply 3(1)} (fold Symset.insert cs))) |
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
|
52 |
|
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
57962
diff
changeset
|
53 |
fun keep_functions cs = |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
54 |
Data.map (map_data (@{apply 3(2)} (fold Symset.insert cs))) |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
55 |
|
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
56 |
fun keep_function thy = Symset.member (#keep_functions (Data.get thy)) |
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
|
57 |
|
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
|
58 |
fun processed_specs thy = Symtab.lookup (#processed_specs (Data.get 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
|
59 |
|
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
|
60 |
fun store_processed_specs (constname, specs) = |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
57962
diff
changeset
|
61 |
Data.map (map_data (@{apply 3(3)} (Symtab.update_new (constname, specs)))) |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35267
diff
changeset
|
62 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
63 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
64 |
fun defining_term_of_introrule_term t = |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
65 |
let |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
66 |
val _ $ u = Logic.strip_imp_concl t |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
67 |
in fst (strip_comb u) end |
55437 | 68 |
(* |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
69 |
in case pred of |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
70 |
Const (c, T) => c |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
71 |
| _ => raise TERM ("defining_const_of_introrule_term failed: Not a constant", [t]) |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
72 |
end |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
73 |
*) |
59582 | 74 |
val defining_term_of_introrule = defining_term_of_introrule_term o Thm.prop_of |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
75 |
|
40142
128f8a1611e6
relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents:
40053
diff
changeset
|
76 |
fun defining_const_of_introrule th = |
55437 | 77 |
(case defining_term_of_introrule th of |
78 |
Const (c, _) => c |
|
59582 | 79 |
| _ => raise TERM ("defining_const_of_introrule failed: Not a constant", [Thm.prop_of th])) |
40142
128f8a1611e6
relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents:
40053
diff
changeset
|
80 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
81 |
(*TODO*) |
50056 | 82 |
fun is_introlike_term _ = true |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
83 |
|
59582 | 84 |
val is_introlike = is_introlike_term o Thm.prop_of |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
85 |
|
69593 | 86 |
fun check_equation_format_term (t as (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ u $ _)) = |
55437 | 87 |
(case strip_comb u of |
88 |
(Const (_, T), args) => |
|
89 |
if (length (binder_types T) = length args) then |
|
90 |
true |
|
91 |
else |
|
92 |
raise TERM ("check_equation_format_term failed: Number of arguments mismatch", [t]) |
|
93 |
| _ => raise TERM ("check_equation_format_term failed: Not a constant", [t])) |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
94 |
| check_equation_format_term t = |
55437 | 95 |
raise TERM ("check_equation_format_term failed: Not an equation", [t]) |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
96 |
|
59582 | 97 |
val check_equation_format = check_equation_format_term o Thm.prop_of |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
98 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
99 |
|
69593 | 100 |
fun defining_term_of_equation_term (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ u $ _) = fst (strip_comb u) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
101 |
| defining_term_of_equation_term t = |
55437 | 102 |
raise TERM ("defining_const_of_equation_term failed: Not an equation", [t]) |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
103 |
|
59582 | 104 |
val defining_term_of_equation = defining_term_of_equation_term o Thm.prop_of |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
105 |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
106 |
fun defining_const_of_equation th = |
55437 | 107 |
(case defining_term_of_equation th of |
108 |
Const (c, _) => c |
|
59582 | 109 |
| _ => raise TERM ("defining_const_of_equation failed: Not a constant", [Thm.prop_of th])) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
110 |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
111 |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
112 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
113 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
114 |
(* Normalizing equations *) |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
115 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
116 |
fun mk_meta_equation th = |
59582 | 117 |
(case Thm.prop_of th of |
69593 | 118 |
Const (\<^const_name>\<open>Trueprop\<close>, _) $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _) => |
55437 | 119 |
th RS @{thm eq_reflection} |
120 |
| _ => th) |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
121 |
|
62958
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
61268
diff
changeset
|
122 |
val meta_fun_cong = @{lemma "\<And>f :: 'a::{} \<Rightarrow> 'b::{}.f == g ==> f x == g x" by simp} |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
123 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
124 |
fun full_fun_cong_expand th = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
125 |
let |
59582 | 126 |
val (f, args) = strip_comb (fst (Logic.dest_equals (Thm.prop_of th))) |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
127 |
val i = length (binder_types (fastype_of f)) - length args |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
128 |
in funpow i (fn th => th RS meta_fun_cong) th end; |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
129 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
130 |
fun declare_names s xs ctxt = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
131 |
let |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43277
diff
changeset
|
132 |
val res = Name.invent_names ctxt s xs |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
133 |
in (res, fold Name.declare (map fst res) ctxt) end |
55437 | 134 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
135 |
fun split_all_pairs thy th = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
136 |
let |
51552
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
50056
diff
changeset
|
137 |
val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) |
50056 | 138 |
val ((_, [th']), _) = Variable.import true [th] ctxt |
59582 | 139 |
val t = Thm.prop_of th' |
55437 | 140 |
val frees = Term.add_frees t [] |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
141 |
val freenames = Term.add_free_names t [] |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
142 |
val nctxt = Name.make_context freenames |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
143 |
fun mk_tuple_rewrites (x, T) nctxt = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
144 |
let |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
145 |
val Ts = HOLogic.flatten_tupleT T |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
146 |
val (xTs, nctxt') = declare_names x Ts nctxt |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
147 |
val paths = HOLogic.flat_tupleT_paths T |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
148 |
in ((Free (x, T), HOLogic.mk_ptuple paths T (map Free xTs)), nctxt') end |
55437 | 149 |
val (rewr, _) = fold_map mk_tuple_rewrites frees nctxt |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
150 |
val t' = Pattern.rewrite_term thy rewr [] t |
51552
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
50056
diff
changeset
|
151 |
val th'' = |
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
50056
diff
changeset
|
152 |
Goal.prove ctxt (Term.add_free_names t' []) [] t' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59205
diff
changeset
|
153 |
(fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt)) |
63170 | 154 |
val th''' = Local_Defs.unfold0 ctxt [@{thm split_conv}, @{thm fst_conv}, @{thm snd_conv}] th'' |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
155 |
in |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
156 |
th''' |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
157 |
end; |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
158 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
159 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
160 |
fun inline_equations thy th = |
33404 | 161 |
let |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
162 |
val ctxt = Proof_Context.init_global thy |
69593 | 163 |
val inline_defs = Named_Theorems.get ctxt \<^named_theorems>\<open>code_pred_inline\<close> |
57962 | 164 |
val th' = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps inline_defs) th |
55437 | 165 |
(*val _ = print_step options |
33404 | 166 |
("Inlining " ^ (Syntax.string_of_term_global thy (prop_of th)) |
167 |
^ "with " ^ (commas (map ((Syntax.string_of_term_global thy) o prop_of) inline_defs)) |
|
168 |
^" to " ^ (Syntax.string_of_term_global thy (prop_of th')))*) |
|
169 |
in |
|
170 |
th' |
|
171 |
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
|
172 |
|
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
|
173 |
fun normalize_equation thy th = |
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
|
174 |
mk_meta_equation th |
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
|
175 |
|> full_fun_cong_expand |
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
|
176 |
|> split_all_pairs 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
|
177 |
|> tap check_equation_format |
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
|
178 |
|> inline_equations thy |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
179 |
|
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
|
180 |
fun normalize_intros thy th = |
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
|
181 |
split_all_pairs thy th |
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
|
182 |
|> inline_equations 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
|
183 |
|
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
|
184 |
fun normalize thy th = |
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
|
185 |
if is_equationlike th then |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35267
diff
changeset
|
186 |
normalize_equation thy th |
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
|
187 |
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
|
188 |
normalize_intros thy th |
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
|
189 |
|
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
|
190 |
fun get_specification options thy t = |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
191 |
let |
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
|
192 |
(*val (c, T) = dest_Const t |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51552
diff
changeset
|
193 |
val t = Const (Axclass.unoverload_const thy (c, T), T)*) |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35267
diff
changeset
|
194 |
val _ = if show_steps options then |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
35267
diff
changeset
|
195 |
tracing ("getting specification of " ^ Syntax.string_of_term_global thy t ^ |
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
|
196 |
" with type " ^ Syntax.string_of_typ_global thy (fastype_of t)) |
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
|
197 |
else () |
42361 | 198 |
val ctxt = Proof_Context.init_global thy |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
199 |
fun filtering th = |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
200 |
if is_equationlike th andalso |
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
|
201 |
defining_const_of_equation (normalize_equation thy th) = fst (dest_Const t) then |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
202 |
SOME (normalize_equation thy th) |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
203 |
else |
40142
128f8a1611e6
relaxing the filtering condition for getting specifications from Spec_Rules
bulwahn
parents:
40053
diff
changeset
|
204 |
if is_introlike th andalso defining_const_of_introrule th = fst (dest_Const t) then |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
205 |
SOME th |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
206 |
else |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
207 |
NONE |
35758
c029f85d3879
adopting predicate compiler to changes in Spec_Rules; removed dependency to Nitpick_Intros
bulwahn
parents:
35624
diff
changeset
|
208 |
fun filter_defs ths = map_filter filtering (map (normalize thy o Thm.transfer thy) ths) |
55437 | 209 |
val spec = |
69593 | 210 |
(case filter_defs (Named_Theorems.get ctxt \<^named_theorems>\<open>code_pred_def\<close>) of |
55437 | 211 |
[] => |
212 |
(case Spec_Rules.retrieve ctxt t of |
|
213 |
[] => error ("No specification for " ^ Syntax.string_of_term_global thy t) |
|
71179 | 214 |
| ({rules = ths, ...} :: _) => filter_defs ths) |
57962 | 215 |
| ths => ths) |
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
|
216 |
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
|
217 |
if show_intermediate_results options then |
43277 | 218 |
tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^ |
61268 | 219 |
commas (map (Thm.string_of_thm_global thy) spec)) |
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
|
220 |
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
|
221 |
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
|
222 |
spec |
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
|
223 |
end |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
224 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
225 |
val logic_operator_names = |
69593 | 226 |
[\<^const_name>\<open>Pure.eq\<close>, |
227 |
\<^const_name>\<open>Pure.imp\<close>, |
|
228 |
\<^const_name>\<open>Trueprop\<close>, |
|
229 |
\<^const_name>\<open>Not\<close>, |
|
230 |
\<^const_name>\<open>HOL.eq\<close>, |
|
231 |
\<^const_name>\<open>HOL.implies\<close>, |
|
232 |
\<^const_name>\<open>All\<close>, |
|
233 |
\<^const_name>\<open>Ex\<close>, |
|
234 |
\<^const_name>\<open>HOL.conj\<close>, |
|
235 |
\<^const_name>\<open>HOL.disj\<close>] |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
236 |
|
55437 | 237 |
fun special_cases (c, _) = |
238 |
member (op =) |
|
69593 | 239 |
[\<^const_name>\<open>Product_Type.Unity\<close>, |
240 |
\<^const_name>\<open>False\<close>, |
|
241 |
\<^const_name>\<open>Suc\<close>, \<^const_name>\<open>Nat.zero_nat_inst.zero_nat\<close>, |
|
242 |
\<^const_name>\<open>Nat.one_nat_inst.one_nat\<close>, |
|
243 |
\<^const_name>\<open>Orderings.less\<close>, \<^const_name>\<open>Orderings.less_eq\<close>, |
|
244 |
\<^const_name>\<open>Groups.zero\<close>, |
|
245 |
\<^const_name>\<open>Groups.one\<close>, \<^const_name>\<open>Groups.plus\<close>, |
|
246 |
\<^const_name>\<open>Nat.ord_nat_inst.less_eq_nat\<close>, |
|
247 |
\<^const_name>\<open>Nat.ord_nat_inst.less_nat\<close>, |
|
55437 | 248 |
(* FIXME |
249 |
@{const_name number_nat_inst.number_of_nat}, |
|
250 |
*) |
|
69593 | 251 |
\<^const_name>\<open>Num.Bit0\<close>, |
252 |
\<^const_name>\<open>Num.Bit1\<close>, |
|
253 |
\<^const_name>\<open>Num.One\<close>, |
|
254 |
\<^const_name>\<open>Int.zero_int_inst.zero_int\<close>, |
|
255 |
\<^const_name>\<open>List.filter\<close>, |
|
256 |
\<^const_name>\<open>HOL.If\<close>, |
|
257 |
\<^const_name>\<open>Groups.minus\<close>] c |
|
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
|
258 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
259 |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
260 |
fun obtain_specification_graph options thy t = |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
261 |
let |
42361 | 262 |
val ctxt = Proof_Context.init_global thy |
50056 | 263 |
fun is_nondefining_const (c, _) = member (op =) logic_operator_names c |
264 |
fun has_code_pred_intros (c, _) = can (Core_Data.intros_of ctxt) c |
|
55399 | 265 |
fun case_consts (c, _) = is_some (Ctr_Sugar.ctr_sugar_of_case ctxt c) |
266 |
fun is_datatype_constructor (x as (_, T)) = |
|
267 |
(case body_type T of |
|
268 |
Type (Tcon, _) => can (Ctr_Sugar.dest_ctr ctxt Tcon) (Const x) |
|
269 |
| _ => false) |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
270 |
fun defiants_of specs = |
59582 | 271 |
fold (Term.add_consts o Thm.prop_of) specs [] |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
272 |
|> filter_out is_datatype_constructor |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
273 |
|> filter_out is_nondefining_const |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
274 |
|> filter_out has_code_pred_intros |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
275 |
|> filter_out case_consts |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
276 |
|> filter_out special_cases |
77723
b761c91c2447
performance tuning: prefer functor Set() over Table();
wenzelm
parents:
74561
diff
changeset
|
277 |
|> filter_out (fn (c, _) => Symset.member (#ignore_consts (Data.get thy)) c) |
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
|
278 |
|> map (fn (c, _) => (c, Sign.the_const_constraint thy c)) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
279 |
|> map Const |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
280 |
(* |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33756
diff
changeset
|
281 |
|> filter is_defining_constname*) |
35405 | 282 |
fun extend t gr = |
283 |
if can (Term_Graph.get_node gr) t then gr |
|
284 |
else |
|
285 |
let |
|
286 |
val specs = get_specification options thy t |
|
287 |
(*val _ = print_specification options thy constname specs*) |
|
288 |
val us = defiants_of specs |
|
289 |
in |
|
290 |
gr |
|
291 |
|> Term_Graph.new_node (t, specs) |
|
292 |
|> fold extend us |
|
293 |
|> fold (fn u => Term_Graph.add_edge (t, u)) us |
|
294 |
end |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
295 |
in |
35405 | 296 |
extend t Term_Graph.empty |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
297 |
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
|
298 |
|
55437 | 299 |
end |