src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML
author bulwahn
Tue, 31 Aug 2010 10:51:49 +0200
changeset 38957 2eb265efa1b0
parent 38956 2e5bf3bc7361
child 39789 533dd8cda12c
permissions -rw-r--r--
exporting mode analysis for use in prolog generation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33265
01c9c6dbd890 proper headers;
wenzelm
parents: 33250
diff changeset
     1
(*  Title:      HOL/Tools/Predicate_Compile/predicate_compile_fun.ML
01c9c6dbd890 proper headers;
wenzelm
parents: 33250
diff changeset
     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
01c9c6dbd890 proper headers;
wenzelm
parents: 33250
diff changeset
     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: 33522
diff 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: 35021
diff 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: 33752
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff changeset
    23
  type T = (term * term) Item_Net.T;
38760
wenzelm
parents: 38553
diff changeset
    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: 35021
diff 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: 35021
diff 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: 36039
diff changeset
    30
  let
810357220388 improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
bulwahn
parents: 36039
diff changeset
    31
    val poss_preds = map_filter (fn (f, p) =>
36039
bulwahn
parents: 36035
diff changeset
    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: 36039
diff 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: 36039
diff changeset
    34
  in
810357220388 improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
bulwahn
parents: 36039
diff 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: 36039
diff 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: 36039
diff changeset
    37
    | _ => NONE
810357220388 improving lookup function to handle overloaded constants more correctly in the function flattening of the predicate compiler
bulwahn
parents: 36039
diff 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: 35021
diff changeset
    40
fun 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: 35021
diff changeset
    41
  case Item_Net.retrieve (Fun_Pred.get thy) (Const (name, Term.dummyT)) 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 37591
diff 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
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
    62
      | _ => raise Fail "A non-variable term at a higher-order position")
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
    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: 35411
diff 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
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    84
    val (argTs, resT) = (strip_type (fastype_of func))
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) [])
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    87
    val (argnames, nctxt') = Name.variants (replicate (length argTs) "a") nctxt
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    88
    val ([resname], nctxt'') = Name.variants ["r"] nctxt'
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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff changeset
   108
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   109
fun flatten thy lookup_pred t (names, prems) =
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   110
  let
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   111
    fun lift t (names, prems) =
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   112
      case lookup_pred (Envir.eta_contract t) of
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   113
        SOME pred => [(pred, (names, prems))]
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   114
      | NONE =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   115
        let
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   116
          val (vars, body) = strip_abs t
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   117
          val _ = assert (fastype_of body = body_type (fastype_of body))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   118
          val absnames = Name.variant_list names (map fst vars)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   119
          val frees = map2 (curry Free) absnames (map snd vars)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   120
          val body' = subst_bounds (rev frees, body)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   121
          val resname = Name.variant (absnames @ names) "res"
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   122
          val resvar = Free (resname, fastype_of body)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   123
          val t = flatten' body' ([], [])
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   124
            |> map (fn (res, (inner_names, inner_prems)) =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   125
              let
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   126
                fun mk_exists (x, T) t = HOLogic.mk_exists (x, T, t)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   127
                val vTs = 
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   128
                  fold Term.add_frees inner_prems []
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   129
                  |> filter (fn (x, T) => member (op =) inner_names x)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   130
                val t = 
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   131
                  fold mk_exists vTs
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   132
                  (foldr1 HOLogic.mk_conj (HOLogic.mk_eq (res, resvar) ::
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   133
                    map HOLogic.dest_Trueprop inner_prems))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   134
              in
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   135
                t
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   136
              end)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   137
              |> foldr1 HOLogic.mk_disj
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   138
              |> fold lambda (resvar :: rev frees)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   139
        in
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   140
          [(t, (names, prems))]
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   141
        end
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   142
    and flatten_or_lift (t, T) (names, prems) =
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   143
      if fastype_of t = T then
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   144
        flatten' t (names, prems)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   145
      else
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   146
        (* note pred_type might be to general! *)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   147
        if (pred_type (fastype_of t) = T) then
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   148
          lift t (names, prems)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   149
        else
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   150
          error ("unexpected input for flatten or lift" ^ Syntax.string_of_term_global thy t ^
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   151
          ", " ^  Syntax.string_of_typ_global thy T)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   152
    and flatten' (t as Const (name, T)) (names, prems) = [(t, (names, prems))]
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   153
      | flatten' (t as Free (f, T)) (names, prems) = [(t, (names, prems))]
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   154
      | flatten' (t as Abs _) (names, prems) = [(t, (names, prems))]
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff changeset
   163
            in
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   164
              flatten' B (names, prems)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   165
              |> maps (fn (B', (names, prems)) =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   166
                (flatten' x (names, prems)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   167
                |> map (fn (res, (names, prems)) => (res, (names, (HOLogic.mk_Trueprop B') :: prems))))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   168
                @ (flatten' y (names, prems)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   169
                |> map (fn (res, (names, prems)) =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   170
                  (* in general unsound! *)
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff 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: 35021
diff changeset
   176
            in
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff 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: 35021
diff changeset
   178
              |> maps (fn (res, (names, prems)) =>
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff 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: 35021
diff changeset
   180
            end)
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37135
diff changeset
   181
        | Const (@{const_name prod_case}, _) => 
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: 35021
diff changeset
   182
            (let
35877
295e1af6c8dc generalized split transformation in the function flattening
bulwahn
parents: 35876
diff changeset
   183
              val (_, g :: res :: 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: 35021
diff changeset
   184
              val [res1, res2] = Name.variant_list names ["res1", "res2"]
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35021
diff changeset
   185
              val (T1, T2) = HOLogic.dest_prodT (fastype_of res)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35021
diff changeset
   186
              val (resv1, resv2) = (Free (res1, T1), Free (res2, T2))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35021
diff changeset
   187
            in
35877
295e1af6c8dc generalized split transformation in the function flattening
bulwahn
parents: 35876
diff changeset
   188
              flatten' (betapplys (g, (resv1 :: resv2 :: args)))
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: 35021
diff changeset
   189
              (res1 :: res2 :: names,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35021
diff changeset
   190
              HOLogic.mk_Trueprop (HOLogic.mk_eq (res, HOLogic.mk_prod (resv1, resv2))) :: 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: 35021
diff changeset
   191
            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: 35021
diff changeset
   192
        | _ =>
36029
a790b94e090c removing fishing for split thm in the predicate compiler
bulwahn
parents: 35878
diff changeset
   193
        case find_split_thm thy (fst (strip_comb t)) of
a790b94e090c removing fishing for split thm in the predicate compiler
bulwahn
parents: 35878
diff changeset
   194
          SOME raw_split_thm =>
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   195
          let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   196
            val (f, args) = strip_comb t
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36249
diff changeset
   197
            val split_thm = prepare_split_thm (ProofContext.init_global thy) raw_split_thm
36029
a790b94e090c removing fishing for split thm in the predicate compiler
bulwahn
parents: 35878
diff changeset
   198
            val (assms, concl) = Logic.strip_horn (prop_of split_thm)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   199
            val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl) 
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   200
            val subst = Pattern.match thy (split_t, t) (Vartab.empty, Vartab.empty)
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   201
            val (_, split_args) = strip_comb split_t
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   202
            val match = split_args ~~ args
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff changeset
   203
            fun flatten_of_assm assm =
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   204
              let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   205
                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
   206
                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
   207
                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
   208
                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
   209
                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: 35021
diff changeset
   210
                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: 35021
diff changeset
   211
                  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
   212
              in
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff changeset
   213
                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: 35021
diff changeset
   214
                |> 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: 35021
diff changeset
   215
                  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: 35021
diff changeset
   216
                    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: 35021
diff changeset
   217
                  in (names, prems' @ prems) end)
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff changeset
   218
                |> maps (flatten' inner_t)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   219
              end
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   220
          in
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff changeset
   221
            maps flatten_of_assm assms
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   222
          end
36029
a790b94e090c removing fishing for split thm in the predicate compiler
bulwahn
parents: 35878
diff changeset
   223
      | NONE =>
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   224
          let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   225
            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: 35021
diff changeset
   226
            val args = map (Pattern.eta_long []) args
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35021
diff changeset
   227
            val _ = assert (fastype_of t = body_type (fastype_of t))
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   228
            val f' = lookup_pred f
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   229
            val Ts = case f' of
35873
d09a58c890e3 simplifying function flattening
bulwahn
parents: 35845
diff changeset
   230
              SOME pred => (fst (split_last (binder_types (fastype_of pred))))
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   231
            | NONE => binder_types (fastype_of f)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   232
          in
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   233
            folds_map flatten_or_lift (args ~~ Ts) (names, prems) |>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   234
            (case f' of
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   235
              NONE =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   236
                map (fn (argvs, (names', prems')) => (list_comb (f, argvs), (names', prems')))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   237
            | SOME pred =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   238
                map (fn (argvs, (names', prems')) =>
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   239
                  let
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   240
                    fun lift_arg T t =
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   241
                      if (fastype_of t) = T then t
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   242
                      else
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   243
                        let
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   244
                          val _ = assert (T =
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   245
                            (binder_types (fastype_of t) @ [@{typ bool}] ---> @{typ bool}))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   246
                          fun mk_if T (b, t, e) =
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   247
                            Const (@{const_name If}, @{typ bool} --> T --> T --> T) $ b $ t $ e
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   248
                          val Ts = binder_types (fastype_of t)
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   249
                        in
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   250
                          list_abs (map (pair "x") Ts @ [("b", @{typ bool})],
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   251
                            mk_if @{typ bool} (list_comb (t, map Bound (length Ts downto 1)),
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   252
                            HOLogic.mk_eq (@{term True}, Bound 0),
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   253
                            HOLogic.mk_eq (@{term False}, Bound 0)))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   254
                        end
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   255
                    val argvs' = map2 lift_arg Ts argvs
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   256
                    val resname = Name.variant names' "res"
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   257
                    val resvar = Free (resname, body_type (fastype_of t))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   258
                    val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs' @ [resvar]))
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   259
                  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
   260
          end
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   261
  in
35875
b0d24a74b06b restructuring function flattening
bulwahn
parents: 35874
diff changeset
   262
    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
   263
  end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   264
36249
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   265
(* 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: 36051
diff changeset
   266
fun pred_of thy f =
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   267
  let
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   268
    val (name, T) = dest_Const f
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   269
    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: 36051
diff changeset
   270
    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: 36051
diff changeset
   271
    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: 36051
diff changeset
   272
  in
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   273
    (name', Const (name', T'))
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   274
  end
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   275
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   276
(* 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
   277
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: 35021
diff changeset
   278
  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
   279
    ([], thy)
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   280
  else
35876
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   281
    let
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   282
      val consts = map fst specs
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   283
      val eqns = maps snd specs
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   284
      (* create prednames *)
35876
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   285
      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: 36051
diff changeset
   286
      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: 35875
diff changeset
   287
      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: 35875
diff changeset
   288
      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: 36051
diff changeset
   289
      (* 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: 35875
diff changeset
   290
      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: 36051
diff changeset
   291
      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: 36051
diff changeset
   292
      val dst_preds = distinct (op =) preds
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   293
      val dst_prednames = distinct (op =) prednames
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   294
      (* 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: 35875
diff changeset
   295
      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: 36051
diff changeset
   296
        ((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: 35875
diff changeset
   297
          (Fun_Pred.get thy)
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   298
      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: 35875
diff changeset
   299
      (* create intro rules *)
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   300
      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: 35875
diff changeset
   301
        if (body_type (fastype_of func) = @{typ bool}) then
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   302
         (* 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: 35875
diff changeset
   303
          [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
   304
        else
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   305
          let
35876
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   306
            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: 35875
diff changeset
   307
          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: 35875
diff changeset
   308
            |> map (fn (resultt, (names', prems)) =>
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   309
              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: 35875
diff changeset
   310
          end
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   311
      fun mk_rewr_thm (func, pred) = @{thm refl}
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   312
      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: 36051
diff changeset
   313
      val (intrs, thy') = thy
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   314
        |> Sign.add_consts_i
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   315
          (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: 36051
diff changeset
   316
           dst_preds)
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   317
        |> 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: 36051
diff changeset
   318
      val specs = map (fn predname => (predname,
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   319
          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: 36051
diff changeset
   320
        dst_prednames
35878
74a74828d682 cleaning the function flattening
bulwahn
parents: 35877
diff changeset
   321
      val thy'' = Fun_Pred.map
36249
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   322
        (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: 36051
diff changeset
   323
          (dst_funs ~~ dst_preds))) thy'
36035
d82682936c52 adding registration of functions in the function flattening
bulwahn
parents: 36029
diff changeset
   324
      fun functional_mode_of T =
d82682936c52 adding registration of functions in the function flattening
bulwahn
parents: 36029
diff changeset
   325
        list_fun_mode (replicate (length (binder_types T)) Input @ [Output])
d82682936c52 adding registration of functions in the function flattening
bulwahn
parents: 36029
diff changeset
   326
      val thy''' = fold
d82682936c52 adding registration of functions in the function flattening
bulwahn
parents: 36029
diff changeset
   327
        (fn (predname, Const (name, T)) => Predicate_Compile_Core.register_alternative_function
d82682936c52 adding registration of functions in the function flattening
bulwahn
parents: 36029
diff changeset
   328
          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: 36051
diff changeset
   329
      (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: 35875
diff changeset
   330
    in
36035
d82682936c52 adding registration of functions in the function flattening
bulwahn
parents: 36029
diff changeset
   331
      (specs, thy''')
35876
ac44e2312f0a only adding lifted arguments to item net in the function flattening; correcting indentation; removing dead code
bulwahn
parents: 35875
diff changeset
   332
    end
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   333
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   334
fun rewrite_intro thy intro =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   335
  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: 35021
diff changeset
   336
    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: 36051
diff changeset
   337
    (*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: 35411
diff changeset
   338
    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
   339
    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
   340
    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
   341
    fun rewrite prem names =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   342
      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: 35021
diff changeset
   343
        (*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: 36051
diff changeset
   344
        val t = HOLogic.dest_Trueprop prem
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   345
        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
   346
            SOME t => (t, HOLogic.mk_not)
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   347
          | NONE => (t, I)
36249
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   348
        val (P, args) = strip_comb lit
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   349
      in
35874
bcfa6b4b21c6 renaming mk_prems to flatten in the function flattening
bulwahn
parents: 35873
diff changeset
   350
        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
   351
        |> map (fn (resargs, (names', prems')) =>
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   352
          let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   353
            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: 38760
diff changeset
   354
          in (prems' @ [prem'], names') end)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   355
      end
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   356
    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
   357
      |> maps (fn (prems', frees') =>
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   358
        rewrite concl frees'
38956
2e5bf3bc7361 renaming
bulwahn
parents: 38954
diff changeset
   359
        |> map (fn (conclprems, _) =>
2e5bf3bc7361 renaming
bulwahn
parents: 38954
diff changeset
   360
          let
2e5bf3bc7361 renaming
bulwahn
parents: 38954
diff changeset
   361
            val (conclprems', concl') = split_last conclprems
2e5bf3bc7361 renaming
bulwahn
parents: 38954
diff changeset
   362
          in
2e5bf3bc7361 renaming
bulwahn
parents: 38954
diff changeset
   363
            Logic.list_implies ((flat prems') @ conclprems', concl')
2e5bf3bc7361 renaming
bulwahn
parents: 38954
diff changeset
   364
          end))
36249
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   365
    (*val _ = tracing ("Rewritten intro to " ^
24f2ab79b506 replaced call to inductive package by axiomatization in the function flattening of the predicate compiler
bulwahn
parents: 36051
diff changeset
   366
      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
   367
  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: 34954
diff changeset
   368
    map (Drule.export_without_context o Skip_Proof.make_thm thy) intro_ts'
33630
68e058d061f5 removed unnecessary oracle in the predicate compiler
bulwahn
parents: 33522
diff changeset
   369
  end
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   370
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   371
end;