| author | blanchet | 
| Fri, 14 Sep 2012 12:09:27 +0200 | |
| changeset 49371 | 9fa21648d0a1 | 
| parent 49323 | 6dff6b1f5417 | 
| child 50056 | 72efd6b4038d | 
| permissions | -rw-r--r-- | 
| 33265 | 1 | (* Title: HOL/Tools/Predicate_Compile/predicate_compile_fun.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 | Preprocessing functions to predicates. | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 5 | *) | 
| 
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_FUN = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 8 | sig | 
| 33630 
68e058d061f5
removed unnecessary oracle in the predicate compiler
 bulwahn parents: 
33522diff
changeset | 9 | val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 10 | val rewrite_intro : theory -> thm -> thm list | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 11 | val pred_of_function : theory -> string -> string option | 
| 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: 
35021diff
changeset | 12 | val add_function_predicate_translation : (term * term) -> theory -> theory | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 13 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 14 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 15 | structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 16 | struct | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 17 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33752diff
changeset | 18 | open Predicate_Compile_Aux; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 19 | |
| 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: 
35021diff
changeset | 20 | (* Table from function to inductive predicate *) | 
| 
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: 
35021diff
changeset | 21 | structure Fun_Pred = Theory_Data | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 22 | ( | 
| 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: 
35021diff
changeset | 23 | type T = (term * term) Item_Net.T; | 
| 38760 | 24 | val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst); | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 25 | val extend = I; | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 26 | val merge = Item_Net.merge; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 27 | ) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 28 | |
| 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: 
35021diff
changeset | 29 | fun lookup thy net t = | 
| 36051 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 30 | let | 
| 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 31 | val poss_preds = map_filter (fn (f, p) => | 
| 36039 | 32 | SOME (Envir.subst_term (Pattern.match thy (f, t) (Vartab.empty, Vartab.empty)) p) | 
| 36051 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 33 | handle Pattern.MATCH => NONE) (Item_Net.retrieve net t) | 
| 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 34 | in | 
| 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 35 | case poss_preds of | 
| 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 36 | [p] => SOME p | 
| 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 37 | | _ => NONE | 
| 
810357220388
improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
 bulwahn parents: 
36039diff
changeset | 38 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 39 | |
| 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: 
35021diff
changeset | 40 | fun pred_of_function thy name = | 
| 49323 | 41 | case Item_Net.retrieve (Fun_Pred.get thy) (Const (name, dummyT)) of | 
| 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: 
35021diff
changeset | 42 | [] => NONE | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 43 | | [(f, p)] => SOME (fst (dest_Const p)) | 
| 
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: 
35021diff
changeset | 44 |   | _ => error ("Multiple matches possible for lookup of constant " ^ name)
 | 
| 33250 
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: 
35021diff
changeset | 46 | fun defined_const thy name = is_some (pred_of_function thy name) | 
| 
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: 
35021diff
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: 
35021diff
changeset | 48 | fun add_function_predicate_translation (f, p) = | 
| 
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: 
35021diff
changeset | 49 | Fun_Pred.map (Item_Net.update (f, p)) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 50 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 51 | fun transform_ho_typ (T as Type ("fun", _)) =
 | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 52 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 53 | val (Ts, T') = strip_type T | 
| 38553 
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
 haftmann parents: 
37591diff
changeset | 54 | in if T' = HOLogic.boolT then T else (Ts @ [T']) ---> HOLogic.boolT end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 55 | | transform_ho_typ t = t | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 56 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 57 | fun transform_ho_arg arg = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 58 | case (fastype_of arg) of | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 59 |     (T as Type ("fun", _)) =>
 | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 60 | (case arg of | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 61 | Free (name, _) => Free (name, transform_ho_typ T) | 
| 35878 | 62 | | _ => raise Fail "A non-variable term at a higher-order position") | 
| 63 | | _ => arg | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 64 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 65 | fun pred_type T = | 
| 
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 (Ts, T') = strip_type T | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 68 | val Ts' = map transform_ho_typ Ts | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 69 | in | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 70 | (Ts' @ [T']) ---> HOLogic.boolT | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 71 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 72 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 73 | (* creates the list of premises for every intro rule *) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 74 | (* theory -> term -> (string list, term list list) *) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 75 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 76 | fun dest_code_eqn eqn = let | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35411diff
changeset | 77 | val (lhs, rhs) = Logic.dest_equals (Logic.unvarify_global (Thm.prop_of eqn)) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 78 | val (func, args) = strip_comb lhs | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 79 | in ((func, args), rhs) end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 80 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 81 | (* TODO: does not work with higher order functions yet *) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 82 | fun mk_rewr_eq (func, pred) = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 83 | let | 
| 40843 | 84 | val (argTs, resT) = strip_type (fastype_of func) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 85 | val nctxt = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 86 | Name.make_context (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) (func $ pred) []) | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
43324diff
changeset | 87 | val (argnames, nctxt') = fold_map Name.variant (replicate (length argTs) "a") nctxt | 
| 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
43324diff
changeset | 88 | val (resname, nctxt'') = Name.variant "r" nctxt' | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 89 | val args = map Free (argnames ~~ argTs) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 90 | val res = Free (resname, resT) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 91 | in Logic.mk_equals | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 92 | (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res])) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 93 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 94 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 95 | fun folds_map f xs y = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 96 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 97 | fun folds_map' acc [] y = [(rev acc, y)] | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 98 | | folds_map' acc (x :: xs) y = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 99 | maps (fn (x, y) => folds_map' (x :: acc) xs y) (f x y) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 100 | in | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 101 | folds_map' [] xs y | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 102 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 103 | |
| 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: 
35021diff
changeset | 104 | fun keep_functions 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: 
35021diff
changeset | 105 | case try dest_Const (fst (strip_comb t)) of | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 106 | SOME (c, _) => Predicate_Compile_Data.keep_function 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: 
35021diff
changeset | 107 | | _ => false | 
| 
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: 
35021diff
changeset | 108 | |
| 35875 | 109 | fun flatten thy lookup_pred t (names, prems) = | 
| 110 | let | |
| 111 | fun lift t (names, prems) = | |
| 112 | case lookup_pred (Envir.eta_contract t) of | |
| 113 | SOME pred => [(pred, (names, prems))] | |
| 114 | | NONE => | |
| 115 | let | |
| 116 | val (vars, body) = strip_abs t | |
| 42816 
ba14eafef077
eliminated obsolete "assert" function, including divergent (unused!?) clone in Predicate_Compile_Proof;
 wenzelm parents: 
42361diff
changeset | 117 |           val _ = @{assert} (fastype_of body = body_type (fastype_of body))
 | 
| 35875 | 118 | val absnames = Name.variant_list names (map fst vars) | 
| 119 | val frees = map2 (curry Free) absnames (map snd vars) | |
| 120 | val body' = subst_bounds (rev frees, body) | |
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42816diff
changeset | 121 | val resname = singleton (Name.variant_list (absnames @ names)) "res" | 
| 35875 | 122 | val resvar = Free (resname, fastype_of body) | 
| 123 | val t = flatten' body' ([], []) | |
| 124 | |> map (fn (res, (inner_names, inner_prems)) => | |
| 125 | let | |
| 126 | fun mk_exists (x, T) t = HOLogic.mk_exists (x, T, t) | |
| 127 | val vTs = | |
| 128 | fold Term.add_frees inner_prems [] | |
| 129 | |> filter (fn (x, T) => member (op =) inner_names x) | |
| 130 | val t = | |
| 131 | fold mk_exists vTs | |
| 132 | (foldr1 HOLogic.mk_conj (HOLogic.mk_eq (res, resvar) :: | |
| 133 | map HOLogic.dest_Trueprop inner_prems)) | |
| 134 | in | |
| 135 | t | |
| 136 | end) | |
| 137 | |> foldr1 HOLogic.mk_disj | |
| 138 | |> fold lambda (resvar :: rev frees) | |
| 139 | in | |
| 140 | [(t, (names, prems))] | |
| 141 | end | |
| 142 | and flatten_or_lift (t, T) (names, prems) = | |
| 143 | if fastype_of t = T then | |
| 144 | flatten' t (names, prems) | |
| 145 | else | |
| 146 | (* note pred_type might be to general! *) | |
| 147 | if (pred_type (fastype_of t) = T) then | |
| 148 | lift t (names, prems) | |
| 149 | else | |
| 150 |           error ("unexpected input for flatten or lift" ^ Syntax.string_of_term_global thy t ^
 | |
| 151 | ", " ^ Syntax.string_of_typ_global thy T) | |
| 152 | and flatten' (t as Const (name, T)) (names, prems) = [(t, (names, prems))] | |
| 153 | | flatten' (t as Free (f, T)) (names, prems) = [(t, (names, prems))] | |
| 154 | | flatten' (t as Abs _) (names, prems) = [(t, (names, prems))] | |
| 155 | | flatten' (t as _ $ _) (names, prems) = | |
| 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: 
35021diff
changeset | 156 | if Predicate_Compile_Aux.is_constrt thy t orelse keep_functions thy t then | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 157 | [(t, (names, prems))] | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 158 | else | 
| 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: 
35021diff
changeset | 159 | case (fst (strip_comb t)) of | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 160 |           Const (@{const_name "If"}, _) =>
 | 
| 
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: 
35021diff
changeset | 161 | (let | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 162 | val (_, [B, x, y]) = strip_comb 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: 
35021diff
changeset | 163 | in | 
| 35875 | 164 | flatten' B (names, prems) | 
| 165 | |> maps (fn (B', (names, prems)) => | |
| 166 | (flatten' x (names, prems) | |
| 167 | |> map (fn (res, (names, prems)) => (res, (names, (HOLogic.mk_Trueprop B') :: prems)))) | |
| 168 | @ (flatten' y (names, prems) | |
| 169 | |> map (fn (res, (names, prems)) => | |
| 170 | (* in general unsound! *) | |
| 171 | (res, (names, (HOLogic.mk_Trueprop (HOLogic.mk_not B')) :: prems))))) | |
| 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: 
35021diff
changeset | 172 | end) | 
| 
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: 
35021diff
changeset | 173 |         | Const (@{const_name "Let"}, _) => 
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 174 | (let | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 175 | val (_, [f, g]) = strip_comb 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: 
35021diff
changeset | 176 | in | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 177 | flatten' f (names, prems) | 
| 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: 
35021diff
changeset | 178 | |> maps (fn (res, (names, prems)) => | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 179 | flatten' (betapply (g, res)) (names, prems)) | 
| 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: 
35021diff
changeset | 180 | end) | 
| 
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: 
35021diff
changeset | 181 | | _ => | 
| 36029 
a790b94e090c
removing fishing for split thm in the predicate compiler
 bulwahn parents: 
35878diff
changeset | 182 | case find_split_thm thy (fst (strip_comb t)) of | 
| 
a790b94e090c
removing fishing for split thm in the predicate compiler
 bulwahn parents: 
35878diff
changeset | 183 | SOME raw_split_thm => | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 184 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 185 | val (f, args) = strip_comb t | 
| 42361 | 186 | val split_thm = prepare_split_thm (Proof_Context.init_global thy) raw_split_thm | 
| 36029 
a790b94e090c
removing fishing for split thm in the predicate compiler
 bulwahn parents: 
35878diff
changeset | 187 | val (assms, concl) = Logic.strip_horn (prop_of split_thm) | 
| 39802 
7cadad6a18cc
applying case beta reduction to case term before matching in predicate compile function flattening; moving case beta reduction function to Predicate_Compile_Aux
 bulwahn parents: 
39797diff
changeset | 188 | val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl) | 
| 
7cadad6a18cc
applying case beta reduction to case term before matching in predicate compile function flattening; moving case beta reduction function to Predicate_Compile_Aux
 bulwahn parents: 
39797diff
changeset | 189 | val t' = case_betapply thy t | 
| 
7cadad6a18cc
applying case beta reduction to case term before matching in predicate compile function flattening; moving case beta reduction function to Predicate_Compile_Aux
 bulwahn parents: 
39797diff
changeset | 190 | val subst = Pattern.match thy (split_t, t') (Vartab.empty, Vartab.empty) | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 191 | fun flatten_of_assm assm = | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 192 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 193 | val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm)) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 194 | val var_names = Name.variant_list names (map fst vTs) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 195 | val vars = map Free (var_names ~~ (map snd vTs)) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 196 | val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm')) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 197 | val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res) | 
| 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: 
35021diff
changeset | 198 | val (lhss : term list, rhss) = | 
| 
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: 
35021diff
changeset | 199 | split_list (map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems') | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 200 | in | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 201 | folds_map flatten' lhss (var_names @ names, prems) | 
| 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: 
35021diff
changeset | 202 | |> map (fn (ress, (names, prems)) => | 
| 
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: 
35021diff
changeset | 203 | let | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
35021diff
changeset | 204 | val prems' = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (ress ~~ rhss) | 
| 
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: 
35021diff
changeset | 205 | in (names, prems' @ prems) end) | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 206 | |> maps (flatten' inner_t) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 207 | end | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 208 | in | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 209 | maps flatten_of_assm assms | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 210 | end | 
| 36029 
a790b94e090c
removing fishing for split thm in the predicate compiler
 bulwahn parents: 
35878diff
changeset | 211 | | NONE => | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 212 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 213 | val (f, args) = strip_comb 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: 
35021diff
changeset | 214 | val args = map (Pattern.eta_long []) args | 
| 42816 
ba14eafef077
eliminated obsolete "assert" function, including divergent (unused!?) clone in Predicate_Compile_Proof;
 wenzelm parents: 
42361diff
changeset | 215 |             val _ = @{assert} (fastype_of t = body_type (fastype_of t))
 | 
| 35875 | 216 | val f' = lookup_pred f | 
| 217 | val Ts = case f' of | |
| 35873 | 218 | SOME pred => (fst (split_last (binder_types (fastype_of pred)))) | 
| 35875 | 219 | | NONE => binder_types (fastype_of f) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 220 | in | 
| 35875 | 221 | folds_map flatten_or_lift (args ~~ Ts) (names, prems) |> | 
| 222 | (case f' of | |
| 223 | NONE => | |
| 224 | map (fn (argvs, (names', prems')) => (list_comb (f, argvs), (names', prems'))) | |
| 225 | | SOME pred => | |
| 226 | map (fn (argvs, (names', prems')) => | |
| 227 | let | |
| 228 | fun lift_arg T t = | |
| 229 | if (fastype_of t) = T then t | |
| 230 | else | |
| 231 | let | |
| 42816 
ba14eafef077
eliminated obsolete "assert" function, including divergent (unused!?) clone in Predicate_Compile_Proof;
 wenzelm parents: 
42361diff
changeset | 232 |                           val _ = @{assert} (T =
 | 
| 35875 | 233 |                             (binder_types (fastype_of t) @ [@{typ bool}] ---> @{typ bool}))
 | 
| 234 | fun mk_if T (b, t, e) = | |
| 235 |                             Const (@{const_name If}, @{typ bool} --> T --> T --> T) $ b $ t $ e
 | |
| 236 | val Ts = binder_types (fastype_of t) | |
| 35878 | 237 | in | 
| 46219 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
43326diff
changeset | 238 |                           fold_rev Term.abs (map (pair "x") Ts @ [("b", @{typ bool})])
 | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
43326diff
changeset | 239 |                             (mk_if @{typ bool} (list_comb (t, map Bound (length Ts downto 1)),
 | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
43326diff
changeset | 240 |                               HOLogic.mk_eq (@{term True}, Bound 0),
 | 
| 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 wenzelm parents: 
43326diff
changeset | 241 |                               HOLogic.mk_eq (@{term False}, Bound 0)))
 | 
| 35875 | 242 | end | 
| 243 | val argvs' = map2 lift_arg Ts argvs | |
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42816diff
changeset | 244 | val resname = singleton (Name.variant_list names') "res" | 
| 35875 | 245 | val resvar = Free (resname, body_type (fastype_of t)) | 
| 246 | val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs' @ [resvar])) | |
| 247 | in (resvar, (resname :: names', prem :: prems')) end)) | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 248 | end | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 249 | in | 
| 35875 | 250 | map (apfst Envir.eta_contract) (flatten' (Pattern.eta_long [] t) (names, prems)) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 251 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 252 | |
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 253 | (* FIXME: create new predicate name -- does not avoid nameclashing *) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 254 | fun pred_of thy f = | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 255 | let | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 256 | val (name, T) = dest_Const f | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 257 | val base_name' = (Long_Name.base_name name ^ "P") | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 258 | val name' = Sign.full_bname thy base_name' | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 259 |     val T' = if (body_type T = @{typ bool}) then T else pred_type T
 | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 260 | in | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 261 | (name', Const (name', T')) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 262 | end | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 263 | |
| 35878 | 264 | (* assumption: mutual recursive predicates all have the same parameters. *) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 265 | fun define_predicates specs 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: 
35021diff
changeset | 266 | if forall (fn (const, _) => defined_const thy const) specs then | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 267 | ([], thy) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 268 | else | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 269 | let | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 270 | val consts = map fst specs | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 271 | val eqns = maps snd specs | 
| 35878 | 272 | (* create prednames *) | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 273 | val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 274 | val dst_funs = distinct (op =) funs | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 275 | val argss' = map (map transform_ho_arg) argss | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 276 | fun is_lifted (t1, t2) = (fastype_of t2 = pred_type (fastype_of t1)) | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 277 | (* FIXME: higher order arguments also occur in tuples! *) | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 278 | val lifted_args = distinct (op =) (filter is_lifted (flat argss ~~ flat argss')) | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 279 | val (prednames, preds) = split_list (map (pred_of thy) funs) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 280 | val dst_preds = distinct (op =) preds | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 281 | val dst_prednames = distinct (op =) prednames | 
| 35878 | 282 | (* mapping from term (Free or Const) to term *) | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 283 | val net = fold Item_Net.update | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 284 | ((dst_funs ~~ dst_preds) @ lifted_args) | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 285 | (Fun_Pred.get thy) | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 286 | fun lookup_pred t = lookup thy net t | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 287 | (* create intro rules *) | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 288 | fun mk_intros ((func, pred), (args, rhs)) = | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 289 |         if (body_type (fastype_of func) = @{typ bool}) then
 | 
| 35878 | 290 | (* TODO: preprocess predicate definition of rhs *) | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 291 | [Logic.list_implies ([HOLogic.mk_Trueprop rhs], HOLogic.mk_Trueprop (list_comb (pred, args)))] | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 292 | else | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 293 | let | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 294 | val names = Term.add_free_names rhs [] | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 295 | in flatten thy lookup_pred rhs (names, []) | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 296 | |> map (fn (resultt, (names', prems)) => | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 297 | Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt])))) | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 298 | end | 
| 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 299 |       fun mk_rewr_thm (func, pred) = @{thm refl}
 | 
| 35878 | 300 | val intr_ts = maps mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss)) | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 301 | val (intrs, thy') = thy | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 302 | |> Sign.add_consts_i | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 303 | (map (fn Const (name, T) => (Binding.name (Long_Name.base_name name), T, NoSyn)) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 304 | dst_preds) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 305 | |> fold_map Specification.axiom (map (pair (Binding.empty, [])) intr_ts) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 306 | val specs = map (fn predname => (predname, | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 307 | map Drule.export_without_context (filter (Predicate_Compile_Aux.is_intro predname) intrs))) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 308 | dst_prednames | 
| 35878 | 309 | val thy'' = Fun_Pred.map | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 310 | (fold Item_Net.update (map (pairself Logic.varify_global) | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 311 | (dst_funs ~~ dst_preds))) thy' | 
| 36035 
d82682936c52
adding registration of functions in the function flattening
 bulwahn parents: 
36029diff
changeset | 312 | fun functional_mode_of T = | 
| 
d82682936c52
adding registration of functions in the function flattening
 bulwahn parents: 
36029diff
changeset | 313 | list_fun_mode (replicate (length (binder_types T)) Input @ [Output]) | 
| 
d82682936c52
adding registration of functions in the function flattening
 bulwahn parents: 
36029diff
changeset | 314 | val thy''' = fold | 
| 40053 
3fa49ea76cbb
using a signature in core_data and moving some more functions to core_data
 bulwahn parents: 
39802diff
changeset | 315 | (fn (predname, Const (name, T)) => Core_Data.register_alternative_function | 
| 36035 
d82682936c52
adding registration of functions in the function flattening
 bulwahn parents: 
36029diff
changeset | 316 | predname (functional_mode_of T) name) | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 317 | (dst_prednames ~~ dst_funs) thy'' | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 318 | in | 
| 36035 
d82682936c52
adding registration of functions in the function flattening
 bulwahn parents: 
36029diff
changeset | 319 | (specs, thy''') | 
| 35876 
ac44e2312f0a
only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
 bulwahn parents: 
35875diff
changeset | 320 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 321 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 322 | fun rewrite_intro thy intro = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 323 | 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: 
35021diff
changeset | 324 | fun lookup_pred t = lookup thy (Fun_Pred.get thy) t | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 325 |     (*val _ = tracing ("Rewriting intro " ^ Display.string_of_thm_global thy intro)*)
 | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35411diff
changeset | 326 | val intro_t = Logic.unvarify_global (prop_of intro) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 327 | val (prems, concl) = Logic.strip_horn intro_t | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 328 | val frees = map fst (Term.add_frees intro_t []) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 329 | fun rewrite prem names = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 330 | 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: 
35021diff
changeset | 331 |         (*val _ = tracing ("Rewriting premise " ^ Syntax.string_of_term_global thy prem ^ "...")*)
 | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 332 | val t = HOLogic.dest_Trueprop prem | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 333 | val (lit, mk_lit) = case try HOLogic.dest_not t of | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 334 | SOME t => (t, HOLogic.mk_not) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 335 | | NONE => (t, I) | 
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 336 | val (P, args) = strip_comb lit | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 337 | in | 
| 35874 
bcfa6b4b21c6
renaming mk_prems to flatten in the function flattening
 bulwahn parents: 
35873diff
changeset | 338 | folds_map (flatten thy lookup_pred) args (names, []) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 339 | |> map (fn (resargs, (names', prems')) => | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 340 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 341 | val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs))) | 
| 38954 
80ce658600b0
changing order of premises generated when flattening functions in premises; adapting example for second attack for hotel key card system
 bulwahn parents: 
38760diff
changeset | 342 | in (prems' @ [prem'], names') end) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 343 | end | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 344 | val intro_ts' = folds_map rewrite prems frees | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 345 | |> maps (fn (prems', frees') => | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 346 | rewrite concl frees' | 
| 38956 | 347 | |> map (fn (conclprems, _) => | 
| 348 | let | |
| 349 | val (conclprems', concl') = split_last conclprems | |
| 350 | in | |
| 351 | Logic.list_implies ((flat prems') @ conclprems', concl') | |
| 352 | end)) | |
| 36249 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 353 |     (*val _ = tracing ("Rewritten intro to " ^
 | 
| 
24f2ab79b506
replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
 bulwahn parents: 
36051diff
changeset | 354 | commas (map (Syntax.string_of_term_global thy) intro_ts'))*) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 355 | in | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34954diff
changeset | 356 | map (Drule.export_without_context o Skip_Proof.make_thm thy) intro_ts' | 
| 33630 
68e058d061f5
removed unnecessary oracle in the predicate compiler
 bulwahn parents: 
33522diff
changeset | 357 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 358 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 359 | end; |