| author | wenzelm | 
| Thu, 29 Jun 2017 11:42:42 +0200 | |
| changeset 66212 | f41396c15bb1 | 
| parent 63170 | eae6549dbea2 | 
| child 69593 | 3dda49e08b9d | 
| 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: 
35267diff
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: 
35267diff
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: 
35267diff
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: 
35267diff
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: 
35267diff
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: 
35267diff
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: 
33756diff
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 = | 
| 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: 
35267diff
changeset | 30 |     {ignore_consts : unit Symtab.table,
 | 
| 
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: 
35267diff
changeset | 31 | keep_functions : unit Symtab.table, | 
| 
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: 
35267diff
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 = | 
| 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: 
35267diff
changeset | 34 |     {ignore_consts = Symtab.empty,
 | 
| 
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: 
35267diff
changeset | 35 | keep_functions = Symtab.empty, | 
| 
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: 
35267diff
changeset | 36 | processed_specs = Symtab.empty}; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 37 | val extend = I; | 
| 33522 | 38 | 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: 
35267diff
changeset | 39 |     ({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: 
35267diff
changeset | 40 |      {ignore_consts = c2, keep_functions = k2, processed_specs = s2}) =
 | 
| 
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: 
35267diff
changeset | 41 |      {ignore_consts = Symtab.merge (K true) (c1, c2),
 | 
| 
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: 
35267diff
changeset | 42 | keep_functions = Symtab.merge (K true) (k1, k2), | 
| 
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: 
35267diff
changeset | 43 | 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 | 44 | ); | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 45 | |
| 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: 
35267diff
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: 
35267diff
changeset | 47 | |
| 
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: 
35267diff
changeset | 48 | 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: 
35267diff
changeset | 49 | 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: 
35267diff
changeset | 50 | |
| 59057 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 wenzelm parents: 
57962diff
changeset | 51 | fun ignore_consts cs = | 
| 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 wenzelm parents: 
57962diff
changeset | 52 |   Data.map (map_data (@{apply 3(1)} (fold (fn c => Symtab.insert (op =) (c, ())) 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: 
35267diff
changeset | 53 | |
| 59057 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 wenzelm parents: 
57962diff
changeset | 54 | fun keep_functions cs = | 
| 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 wenzelm parents: 
57962diff
changeset | 55 |   Data.map (map_data (@{apply 3(2)} (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
 | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 56 | |
| 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: 
35267diff
changeset | 57 | fun keep_function thy = Symtab.defined (#keep_functions (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: 
35267diff
changeset | 58 | |
| 
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: 
35267diff
changeset | 59 | 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: 
35267diff
changeset | 60 | |
| 
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: 
35267diff
changeset | 61 | fun store_processed_specs (constname, specs) = | 
| 59057 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 wenzelm parents: 
57962diff
changeset | 62 |   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: 
35267diff
changeset | 63 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 64 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 65 | 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 | 66 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 67 | 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: 
33756diff
changeset | 68 | in fst (strip_comb u) end | 
| 55437 | 69 | (* | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 70 | in case pred of | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 71 | Const (c, T) => c | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 72 |     | _ => 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 | 73 | end | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 74 | *) | 
| 59582 | 75 | 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 | 76 | |
| 40142 
128f8a1611e6
relaxing the filtering condition for getting specifications from Spec_Rules
 bulwahn parents: 
40053diff
changeset | 77 | fun defining_const_of_introrule th = | 
| 55437 | 78 | (case defining_term_of_introrule th of | 
| 79 | Const (c, _) => c | |
| 59582 | 80 |   | _ => 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: 
40053diff
changeset | 81 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 82 | (*TODO*) | 
| 50056 | 83 | fun is_introlike_term _ = true | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 84 | |
| 59582 | 85 | 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 | 86 | |
| 56245 | 87 | fun check_equation_format_term (t as (Const (@{const_name Pure.eq}, _) $ u $ _)) =
 | 
| 55437 | 88 | (case strip_comb u of | 
| 89 | (Const (_, T), args) => | |
| 90 | if (length (binder_types T) = length args) then | |
| 91 | true | |
| 92 | else | |
| 93 |             raise TERM ("check_equation_format_term failed: Number of arguments mismatch", [t])
 | |
| 94 |       | _ => 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 | 95 | | check_equation_format_term t = | 
| 55437 | 96 |       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 | 97 | |
| 59582 | 98 | 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 | 99 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 100 | |
| 56245 | 101 | fun defining_term_of_equation_term (Const (@{const_name Pure.eq}, _) $ 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: 
33756diff
changeset | 102 | | defining_term_of_equation_term t = | 
| 55437 | 103 |       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 | 104 | |
| 59582 | 105 | 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: 
33756diff
changeset | 106 | |
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 107 | fun defining_const_of_equation th = | 
| 55437 | 108 | (case defining_term_of_equation th of | 
| 109 | Const (c, _) => c | |
| 59582 | 110 |   | _ => 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: 
33756diff
changeset | 111 | |
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 112 | |
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 113 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 114 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 115 | (* Normalizing equations *) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 116 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 117 | fun mk_meta_equation th = | 
| 59582 | 118 | (case Thm.prop_of th of | 
| 55437 | 119 |     Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _) =>
 | 
| 120 |       th RS @{thm eq_reflection}
 | |
| 121 | | _ => th) | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 122 | |
| 62958 
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
 wenzelm parents: 
61268diff
changeset | 123 | 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 | 124 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 125 | fun full_fun_cong_expand th = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 126 | let | 
| 59582 | 127 | 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 | 128 | 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 | 129 | 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 | 130 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 131 | fun declare_names s xs ctxt = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 132 | let | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43277diff
changeset | 133 | 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 | 134 | in (res, fold Name.declare (map fst res) ctxt) end | 
| 55437 | 135 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 136 | fun split_all_pairs thy th = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 137 | let | 
| 51552 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
50056diff
changeset | 138 | val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) | 
| 50056 | 139 | val ((_, [th']), _) = Variable.import true [th] ctxt | 
| 59582 | 140 | val t = Thm.prop_of th' | 
| 55437 | 141 | val frees = Term.add_frees t [] | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 142 | val freenames = Term.add_free_names t [] | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 143 | val nctxt = Name.make_context freenames | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 144 | fun mk_tuple_rewrites (x, T) nctxt = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 145 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 146 | val Ts = HOLogic.flatten_tupleT T | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 147 | 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 | 148 | val paths = HOLogic.flat_tupleT_paths T | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 149 | in ((Free (x, T), HOLogic.mk_ptuple paths T (map Free xTs)), nctxt') end | 
| 55437 | 150 | 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 | 151 | val t' = Pattern.rewrite_term thy rewr [] t | 
| 51552 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
50056diff
changeset | 152 | val th'' = | 
| 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
50056diff
changeset | 153 | 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: 
59205diff
changeset | 154 | (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt)) | 
| 63170 | 155 |     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 | 156 | in | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 157 | th''' | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 158 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 159 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 160 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 161 | fun inline_equations thy th = | 
| 33404 | 162 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51685diff
changeset | 163 | val ctxt = Proof_Context.init_global thy | 
| 57962 | 164 |     val inline_defs = Named_Theorems.get ctxt @{named_theorems code_pred_inline}
 | 
| 165 | val th' = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps inline_defs) th | |
| 55437 | 166 | (*val _ = print_step options | 
| 33404 | 167 |       ("Inlining " ^ (Syntax.string_of_term_global thy (prop_of th))
 | 
| 168 | ^ "with " ^ (commas (map ((Syntax.string_of_term_global thy) o prop_of) inline_defs)) | |
| 169 | ^" to " ^ (Syntax.string_of_term_global thy (prop_of th')))*) | |
| 170 | in | |
| 171 | th' | |
| 172 | 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: 
35267diff
changeset | 173 | |
| 
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: 
35267diff
changeset | 174 | 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: 
35267diff
changeset | 175 | 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: 
35267diff
changeset | 176 | |> 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: 
35267diff
changeset | 177 | |> 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: 
35267diff
changeset | 178 | |> 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: 
35267diff
changeset | 179 | |> inline_equations thy | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 180 | |
| 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: 
35267diff
changeset | 181 | 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: 
35267diff
changeset | 182 | 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: 
35267diff
changeset | 183 | |> 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: 
35267diff
changeset | 184 | |
| 
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: 
35267diff
changeset | 185 | 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: 
35267diff
changeset | 186 | 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: 
35267diff
changeset | 187 | 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: 
35267diff
changeset | 188 | 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: 
35267diff
changeset | 189 | 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: 
35267diff
changeset | 190 | |
| 
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: 
35267diff
changeset | 191 | fun get_specification options thy t = | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 192 | 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: 
35267diff
changeset | 193 | (*val (c, T) = dest_Const t | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51552diff
changeset | 194 | 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: 
35267diff
changeset | 195 | 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: 
35267diff
changeset | 196 |         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: 
35267diff
changeset | 197 | " 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: 
35267diff
changeset | 198 | else () | 
| 42361 | 199 | 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: 
33756diff
changeset | 200 | fun filtering th = | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 201 | 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: 
35267diff
changeset | 202 | 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: 
33756diff
changeset | 203 | 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: 
33756diff
changeset | 204 | else | 
| 40142 
128f8a1611e6
relaxing the filtering condition for getting specifications from Spec_Rules
 bulwahn parents: 
40053diff
changeset | 205 | 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: 
33756diff
changeset | 206 | SOME th | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 207 | else | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 208 | NONE | 
| 35758 
c029f85d3879
adopting predicate compiler to changes in Spec_Rules; removed dependency to Nitpick_Intros
 bulwahn parents: 
35624diff
changeset | 209 | fun filter_defs ths = map_filter filtering (map (normalize thy o Thm.transfer thy) ths) | 
| 55437 | 210 | val spec = | 
| 57962 | 211 |       (case filter_defs (Named_Theorems.get ctxt @{named_theorems code_pred_def}) of
 | 
| 55437 | 212 | [] => | 
| 213 | (case Spec_Rules.retrieve ctxt t of | |
| 214 |             [] => error ("No specification for " ^ Syntax.string_of_term_global thy t)
 | |
| 215 | | ((_, (_, ths)) :: _) => filter_defs ths) | |
| 57962 | 216 | | 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: 
35267diff
changeset | 217 | 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: 
35267diff
changeset | 218 | if show_intermediate_results options then | 
| 43277 | 219 |         tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^
 | 
| 61268 | 220 | 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: 
35267diff
changeset | 221 | 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: 
35267diff
changeset | 222 | 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: 
35267diff
changeset | 223 | 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: 
35267diff
changeset | 224 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 225 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 226 | val logic_operator_names = | 
| 56245 | 227 |   [@{const_name Pure.eq},
 | 
| 228 |    @{const_name Pure.imp},
 | |
| 38558 | 229 |    @{const_name Trueprop},
 | 
| 230 |    @{const_name Not},
 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 231 |    @{const_name HOL.eq},
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38731diff
changeset | 232 |    @{const_name HOL.implies},
 | 
| 38558 | 233 |    @{const_name All},
 | 
| 55437 | 234 |    @{const_name Ex},
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 235 |    @{const_name HOL.conj},
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 236 |    @{const_name HOL.disj}]
 | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 237 | |
| 55437 | 238 | fun special_cases (c, _) = | 
| 239 | member (op =) | |
| 240 |    [@{const_name Product_Type.Unity},
 | |
| 241 |     @{const_name False},
 | |
| 242 |     @{const_name Suc}, @{const_name Nat.zero_nat_inst.zero_nat},
 | |
| 243 |     @{const_name Nat.one_nat_inst.one_nat},
 | |
| 244 |     @{const_name Orderings.less}, @{const_name Orderings.less_eq},
 | |
| 245 |     @{const_name Groups.zero},
 | |
| 246 |     @{const_name Groups.one},  @{const_name Groups.plus},
 | |
| 247 |     @{const_name Nat.ord_nat_inst.less_eq_nat},
 | |
| 248 |     @{const_name Nat.ord_nat_inst.less_nat},
 | |
| 249 | (* FIXME | |
| 250 |     @{const_name number_nat_inst.number_of_nat},
 | |
| 251 | *) | |
| 252 |     @{const_name Num.Bit0},
 | |
| 253 |     @{const_name Num.Bit1},
 | |
| 254 |     @{const_name Num.One},
 | |
| 255 |     @{const_name Int.zero_int_inst.zero_int},
 | |
| 256 |     @{const_name List.filter},
 | |
| 257 |     @{const_name HOL.If},
 | |
| 258 |     @{const_name Groups.minus}] 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: 
35267diff
changeset | 259 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 260 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 261 | 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 | 262 | let | 
| 42361 | 263 | val ctxt = Proof_Context.init_global thy | 
| 50056 | 264 | fun is_nondefining_const (c, _) = member (op =) logic_operator_names c | 
| 265 | fun has_code_pred_intros (c, _) = can (Core_Data.intros_of ctxt) c | |
| 55399 | 266 | fun case_consts (c, _) = is_some (Ctr_Sugar.ctr_sugar_of_case ctxt c) | 
| 267 | fun is_datatype_constructor (x as (_, T)) = | |
| 268 | (case body_type T of | |
| 269 | Type (Tcon, _) => can (Ctr_Sugar.dest_ctr ctxt Tcon) (Const x) | |
| 270 | | _ => false) | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 271 | fun defiants_of specs = | 
| 59582 | 272 | 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: 
33756diff
changeset | 273 | |> 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: 
33756diff
changeset | 274 | |> filter_out is_nondefining_const | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 275 | |> 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: 
33756diff
changeset | 276 | |> filter_out case_consts | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 277 | |> filter_out special_cases | 
| 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: 
35267diff
changeset | 278 | |> filter_out (fn (c, _) => Symtab.defined (#ignore_consts (Data.get thy)) c) | 
| 
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: 
35267diff
changeset | 279 | |> 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: 
33756diff
changeset | 280 | |> map Const | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 281 | (* | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33756diff
changeset | 282 | |> filter is_defining_constname*) | 
| 35405 | 283 | fun extend t gr = | 
| 284 | if can (Term_Graph.get_node gr) t then gr | |
| 285 | else | |
| 286 | let | |
| 287 | val specs = get_specification options thy t | |
| 288 | (*val _ = print_specification options thy constname specs*) | |
| 289 | val us = defiants_of specs | |
| 290 | in | |
| 291 | gr | |
| 292 | |> Term_Graph.new_node (t, specs) | |
| 293 | |> fold extend us | |
| 294 | |> fold (fn u => Term_Graph.add_edge (t, u)) us | |
| 295 | end | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 296 | in | 
| 35405 | 297 | extend t Term_Graph.empty | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 298 | 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: 
35267diff
changeset | 299 | |
| 55437 | 300 | end |