src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML
author bulwahn
Thu, 12 Nov 2009 09:11:26 +0100
changeset 33626 42f69386943a
parent 33623 4ec42d38224f
child 33628 ed2111a5c3ed
permissions -rw-r--r--
new names for predicate functions in the predicate compiler * * * adopting examples of the predicate compiler
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33265
01c9c6dbd890 proper headers;
wenzelm
parents: 33256
diff changeset
     1
(*  Title:      HOL/Tools/Predicate_Compile/predicate_compile_core.ML
01c9c6dbd890 proper headers;
wenzelm
parents: 33256
diff changeset
     2
    Author:     Lukas Bulwahn, TU Muenchen
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
     3
33265
01c9c6dbd890 proper headers;
wenzelm
parents: 33256
diff changeset
     4
A compiler from predicates specified by intro/elim rules to equations.
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
     5
*)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
     6
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
     7
signature PREDICATE_COMPILE_CORE =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
     8
sig
33478
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
     9
  val setup : theory -> theory
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    10
  val code_pred : Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    11
  val code_pred_cmd : Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
    12
  val values_cmd : string list -> Predicate_Compile_Aux.mode' option list option
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
    13
    -> int option * (bool * bool) -> int -> string -> Toplevel.state -> unit
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
    14
  val register_predicate : (string * thm list * thm * int) -> theory -> theory
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
    15
  val register_intros : string * thm list -> theory -> theory
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    16
  val is_registered : theory -> string -> bool
33328
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    17
  val predfun_intro_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    18
  val predfun_elim_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    19
  val predfun_name_of: theory -> string -> Predicate_Compile_Aux.mode -> string
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    20
  val all_preds_of : theory -> string list
33328
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    21
  val modes_of: theory -> string -> Predicate_Compile_Aux.mode list
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    22
  val depth_limited_modes_of: theory -> string -> Predicate_Compile_Aux.mode list
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    23
  val depth_limited_function_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
    24
  val random_modes_of: theory -> string -> Predicate_Compile_Aux.mode list
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
    25
  val random_function_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string
33328
1d93dd8a02c9 moved datatype mode and string functions to the auxillary structure
bulwahn
parents: 33327
diff changeset
    26
  val all_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
    27
  val all_random_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list
33478
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    28
  val intros_of : theory -> string -> thm list
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    29
  val nparams_of : theory -> string -> int
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    30
  val add_intro : thm -> theory -> theory
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    31
  val set_elim : thm -> theory -> theory
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
    32
  val set_nparams : string -> int -> theory -> theory
33478
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    33
  val print_stored_rules : theory -> unit
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    34
  val print_all_modes : theory -> unit
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
    35
  val mk_casesrule : Proof.context -> term -> int -> thm list -> term
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
    36
  val eval_ref : (unit -> term Predicate.pred) option Unsynchronized.ref
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
    37
  val random_eval_ref : (unit -> int * int -> term Predicate.pred * (int * int))
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
    38
    option Unsynchronized.ref
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    39
  val code_pred_intros_attrib : attribute
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    40
  (* used by Quickcheck_Generator *) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    41
  (* temporary for testing of the compilation *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    42
  datatype compilation_funs = CompilationFuns of {
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    43
    mk_predT : typ -> typ,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    44
    dest_predT : typ -> typ,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    45
    mk_bot : typ -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    46
    mk_single : term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    47
    mk_bind : term * term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    48
    mk_sup : term * term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    49
    mk_if : term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    50
    mk_not : term -> term,
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
    51
    mk_map : typ -> typ -> term -> term -> term
33140
83951822bfd0 cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
bulwahn
parents: 33139
diff changeset
    52
  };
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
    53
  val pred_compfuns : compilation_funs
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
    54
  val randompred_compfuns : compilation_funs
33256
b350516cb1f9 adding a prototype of a counter-example generator based on the predicate compiler to HOL/ex
bulwahn
parents: 33251
diff changeset
    55
  val add_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
    56
  val add_quickcheck_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
    57
  val add_depth_limited_equations : Predicate_Compile_Aux.options
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
    58
    -> string list -> theory -> theory
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
    59
  val mk_tracing : string -> term -> term
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    60
end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    61
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    62
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    63
struct
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    64
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
    65
open Predicate_Compile_Aux;
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
    66
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    67
(** auxiliary **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    68
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    69
(* debug stuff *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    70
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    71
fun print_tac s = Seq.single;
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
    72
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
    73
fun print_tac' options s = 
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
    74
  if show_proof_trace options then Tactical.print_tac s else Seq.single;
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
    75
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    76
fun debug_tac msg = Seq.single; (* (fn st => (Output.tracing msg; Seq.single st)); *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    77
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    78
datatype assertion = Max_number_of_subgoals of int
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    79
fun assert_tac (Max_number_of_subgoals i) st =
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    80
  if (nprems_of st <= i) then Seq.single st
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    81
  else error ("assert_tac: Numbers of subgoals mismatch at goal state :"
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    82
    ^ "\n" ^ Pretty.string_of (Pretty.chunks
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    83
      (Goal_Display.pretty_goals_without_context (! Goal_Display.goals_limit) st)));
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
    84
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    85
(* reference to preprocessing of InductiveSet package *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    86
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
    87
val ind_set_codegen_preproc = (fn thy => I) (*Inductive_Set.codegen_preproc;*)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    88
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    89
(** fundamentals **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    90
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    91
(* syntactic operations *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    92
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    93
fun mk_eq (x, xs) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    94
  let fun mk_eqs _ [] = []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    95
        | mk_eqs a (b::cs) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    96
            HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    97
  in mk_eqs x xs end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    98
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    99
fun mk_scomp (t, u) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   100
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   101
    val T = fastype_of t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   102
    val U = fastype_of u
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   103
    val [A] = binder_types T
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   104
    val D = body_type U 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   105
  in 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   106
    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   107
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   108
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   109
fun dest_funT (Type ("fun",[S, T])) = (S, T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   110
  | dest_funT T = raise TYPE ("dest_funT", [T], [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   111
 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   112
fun mk_fun_comp (t, u) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   113
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   114
    val (_, B) = dest_funT (fastype_of t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   115
    val (C, A) = dest_funT (fastype_of u)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   116
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   117
    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   118
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   119
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   120
fun dest_randomT (Type ("fun", [@{typ Random.seed},
32674
b629fbcc5313 merged; adopted to changes from Code_Evaluation in the predicate compiler
bulwahn
parents: 32673
diff changeset
   121
  Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   122
  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   123
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   124
fun mk_tracing s t =
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   125
  Const(@{const_name Code_Evaluation.tracing},
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   126
    @{typ String.literal} --> (fastype_of t) --> (fastype_of t)) $ (HOLogic.mk_literal s) $ t
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   127
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   128
(* destruction of intro rules *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   129
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   130
(* FIXME: look for other place where this functionality was used before *)
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   131
fun strip_intro_concl nparams intro =
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   132
  let
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   133
    val _ $ u = Logic.strip_imp_concl intro
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   134
    val (pred, all_args) = strip_comb u
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   135
    val (params, args) = chop nparams all_args
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   136
  in (pred, (params, args)) end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   137
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   138
(** data structures **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   139
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   140
fun gen_split_smode (mk_tuple, strip_tuple) smode ts =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   141
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   142
    fun split_tuple' _ _ [] = ([], [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   143
    | split_tuple' is i (t::ts) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   144
      (if i mem is then apfst else apsnd) (cons t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   145
        (split_tuple' is (i+1) ts)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   146
    fun split_tuple is t = split_tuple' is 1 (strip_tuple t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   147
    fun split_smode' _ _ [] = ([], [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   148
      | split_smode' smode i (t::ts) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   149
        (if i mem (map fst smode) then
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   150
          case (the (AList.lookup (op =) smode i)) of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   151
            NONE => apfst (cons t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   152
            | SOME is =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   153
              let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   154
                val (ts1, ts2) = split_tuple is t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   155
                fun cons_tuple ts = if null ts then I else cons (mk_tuple ts)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   156
                in (apfst (cons_tuple ts1)) o (apsnd (cons_tuple ts2)) end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   157
          else apsnd (cons t))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   158
        (split_smode' smode (i+1) ts)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   159
  in split_smode' smode 1 ts end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   160
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
   161
fun split_smode smode ts = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple) smode ts
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
   162
fun split_smodeT smode ts = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT) smode ts
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   163
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   164
fun gen_split_mode split_smode (iss, is) ts =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   165
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   166
    val (t1, t2) = chop (length iss) ts 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   167
  in (t1, split_smode is t2) end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   168
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
   169
fun split_mode (iss, is) ts = gen_split_mode split_smode (iss, is) ts
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
   170
fun split_modeT (iss, is) ts = gen_split_mode split_smodeT (iss, is) ts
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   171
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   172
datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   173
  | Generator of (string * typ);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   174
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   175
type moded_clause = term list * (indprem * tmode) list
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   176
type 'a pred_mode_table = (string * (mode * 'a) list) list
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   177
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   178
datatype predfun_data = PredfunData of {
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   179
  name : string,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   180
  definition : thm,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   181
  intro : thm,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   182
  elim : thm
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   183
};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   184
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   185
fun rep_predfun_data (PredfunData data) = data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   186
fun mk_predfun_data (name, definition, intro, elim) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   187
  PredfunData {name = name, definition = definition, intro = intro, elim = elim}
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   188
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   189
datatype function_data = FunctionData of {
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   190
  name : string,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   191
  equation : thm option (* is not used at all? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   192
};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   193
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   194
fun rep_function_data (FunctionData data) = data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   195
fun mk_function_data (name, equation) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   196
  FunctionData {name = name, equation = equation}
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   197
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   198
datatype pred_data = PredData of {
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   199
  intros : thm list,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   200
  elim : thm option,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   201
  nparams : int,
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   202
  functions : bool * (mode * predfun_data) list,
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   203
  random_functions : bool * (mode * function_data) list,
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   204
  depth_limited_functions : bool * (mode * function_data) list,
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   205
  annotated_functions : bool * (mode * function_data) list
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   206
};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   207
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   208
fun rep_pred_data (PredData data) = data;
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   209
fun mk_pred_data ((intros, elim, nparams),
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   210
  (functions, random_functions, depth_limited_functions, annotated_functions)) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   211
  PredData {intros = intros, elim = elim, nparams = nparams,
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   212
    functions = functions, random_functions = random_functions,
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   213
    depth_limited_functions = depth_limited_functions, annotated_functions = annotated_functions}
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   214
fun map_pred_data f (PredData {intros, elim, nparams,
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   215
  functions, random_functions, depth_limited_functions, annotated_functions}) =
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   216
  mk_pred_data (f ((intros, elim, nparams), (functions, random_functions,
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   217
    depth_limited_functions, annotated_functions)))
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   218
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   219
fun eq_option eq (NONE, NONE) = true
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   220
  | eq_option eq (SOME x, SOME y) = eq (x, y)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   221
  | eq_option eq _ = false
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   222
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   223
fun eq_pred_data (PredData d1, PredData d2) = 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   224
  eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   225
  eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   226
  #nparams d1 = #nparams d2
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   227
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   228
structure PredData = Theory_Data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   229
(
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   230
  type T = pred_data Graph.T;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   231
  val empty = Graph.empty;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   232
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   233
  val merge = Graph.merge eq_pred_data;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   234
);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   235
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   236
(* queries *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   237
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   238
fun lookup_pred_data thy name =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   239
  Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   240
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   241
fun the_pred_data thy name = case lookup_pred_data thy name
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   242
 of NONE => error ("No such predicate " ^ quote name)  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   243
  | SOME data => data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   244
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   245
val is_registered = is_some oo lookup_pred_data 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   246
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   247
val all_preds_of = Graph.keys o PredData.get
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   248
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   249
fun intros_of thy = map (Thm.transfer thy) o #intros o the_pred_data thy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   250
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   251
fun the_elim_of thy name = case #elim (the_pred_data thy name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   252
 of NONE => error ("No elimination rule for predicate " ^ quote name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   253
  | SOME thm => Thm.transfer thy thm 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   254
  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   255
val has_elim = is_some o #elim oo the_pred_data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   256
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   257
val nparams_of = #nparams oo the_pred_data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   258
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   259
val modes_of = (map fst) o snd o #functions oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   260
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   261
fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   262
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   263
val defined_functions = fst o #functions oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   264
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   265
fun lookup_predfun_data thy name mode =
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   266
  Option.map rep_predfun_data
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   267
    (AList.lookup (op =) (snd (#functions (the_pred_data thy name))) mode)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   268
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   269
fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   270
  of NONE => error ("No function defined for mode " ^ string_of_mode thy name mode ^
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   271
    " of predicate " ^ name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   272
   | SOME data => data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   273
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   274
val predfun_name_of = #name ooo the_predfun_data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   275
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   276
val predfun_definition_of = #definition ooo the_predfun_data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   277
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   278
val predfun_intro_of = #intro ooo the_predfun_data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   279
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   280
val predfun_elim_of = #elim ooo the_predfun_data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   281
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   282
fun lookup_random_function_data thy name mode =
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   283
  Option.map rep_function_data
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   284
  (AList.lookup (op =) (snd (#random_functions (the_pred_data thy name))) mode)
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   285
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   286
fun the_random_function_data thy name mode = case lookup_random_function_data thy name mode of
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   287
     NONE => error ("No random function defined for mode " ^ string_of_mode thy name mode ^
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   288
       " of predicate " ^ name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   289
   | SOME data => data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   290
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   291
val random_function_name_of = #name ooo the_random_function_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   292
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   293
val random_modes_of = (map fst) o snd o #random_functions oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   294
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   295
val defined_random_functions = fst o #random_functions oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   296
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   297
fun all_random_modes_of thy =
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   298
  map (fn name => (name, random_modes_of thy name)) (all_preds_of thy) 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   299
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
   300
fun lookup_depth_limited_function_data thy name mode =
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   301
  Option.map rep_function_data
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   302
    (AList.lookup (op =) (snd (#depth_limited_functions (the_pred_data thy name))) mode)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   303
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   304
fun the_depth_limited_function_data thy name mode =
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   305
  case lookup_depth_limited_function_data thy name mode of
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   306
    NONE => error ("No depth-limited function defined for mode " ^ string_of_mode thy name mode
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   307
      ^ " of predicate " ^ name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   308
   | SOME data => data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   309
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
   310
val depth_limited_function_name_of = #name ooo the_depth_limited_function_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   311
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   312
val depth_limited_modes_of = (map fst) o snd o #depth_limited_functions oo the_pred_data
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   313
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   314
val defined_depth_limited_functions = fst o #depth_limited_functions oo the_pred_data
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   315
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   316
fun lookup_annotated_function_data thy name mode =
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   317
  Option.map rep_function_data
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   318
    (AList.lookup (op =) (snd (#annotated_functions (the_pred_data thy name))) mode)
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   319
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   320
fun the_annotated_function_data thy name mode = case lookup_annotated_function_data thy name mode
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   321
  of NONE => error ("No annotated function defined for mode " ^ string_of_mode thy name mode
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   322
    ^ " of predicate " ^ name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   323
   | SOME data => data
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   324
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   325
val annotated_function_name_of = #name ooo the_annotated_function_data
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   326
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   327
val annotated_modes_of = (map fst) o snd o #annotated_functions oo the_pred_data
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   328
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   329
val defined_annotated_functions = fst o #annotated_functions oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   330
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   331
(* diagnostic display functions *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   332
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   333
fun print_modes options thy modes =
33251
4b13ab778b78 added option show_modes to predicate compiler
bulwahn
parents: 33250
diff changeset
   334
  if show_modes options then
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
   335
    tracing ("Inferred modes:\n" ^
33251
4b13ab778b78 added option show_modes to predicate compiler
bulwahn
parents: 33250
diff changeset
   336
      cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   337
        (string_of_mode thy s) ms)) modes))
33251
4b13ab778b78 added option show_modes to predicate compiler
bulwahn
parents: 33250
diff changeset
   338
  else ()
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   339
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   340
fun print_pred_mode_table string_of_entry thy pred_mode_table =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   341
  let
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   342
    fun print_mode pred (mode, entry) =  "mode : " ^ string_of_mode thy pred mode
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   343
      ^ string_of_entry pred mode entry
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   344
    fun print_pred (pred, modes) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   345
      "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
   346
    val _ = tracing (cat_lines (map print_pred pred_mode_table))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   347
  in () end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   348
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   349
fun string_of_prem thy (Prem (ts, p)) =
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   350
    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^ "(premise)"
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   351
  | string_of_prem thy (Negprem (ts, p)) =
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   352
    (Syntax.string_of_term_global thy (HOLogic.mk_not (list_comb (p, ts)))) ^ "(negative premise)"
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   353
  | string_of_prem thy (Sidecond t) =
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   354
    (Syntax.string_of_term_global thy t) ^ "(sidecondition)"
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   355
  | string_of_prem thy _ = error "string_of_prem: unexpected input"
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   356
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   357
fun string_of_moded_prem thy (Prem (ts, p), tmode) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   358
    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   359
    "(" ^ (string_of_tmode tmode) ^ ")"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   360
  | string_of_moded_prem thy (Generator (v, T), _) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   361
    "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   362
  | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   363
    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   364
    "(negative mode: " ^ string_of_smode is ^ ")"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   365
  | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   366
    (Syntax.string_of_term_global thy t) ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   367
    "(sidecond mode: " ^ string_of_smode is ^ ")"    
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   368
  | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   369
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   370
fun print_moded_clauses thy =
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   371
  let
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   372
    fun string_of_clause pred mode clauses =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   373
      cat_lines (map (fn (ts, prems) => (space_implode " --> "
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   374
        (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   375
        ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   376
  in print_pred_mode_table string_of_clause thy end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   377
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   378
fun string_of_clause thy pred (ts, prems) =
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   379
  (space_implode " --> "
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   380
  (map (string_of_prem thy) prems)) ^ " --> " ^ pred ^ " "
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   381
   ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   382
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   383
fun print_compiled_terms options thy =
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   384
  if show_compilation options then
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   385
    print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   386
  else K ()
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   387
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   388
fun print_stored_rules thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   389
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   390
    val preds = (Graph.keys o PredData.get) thy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   391
    fun print pred () = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   392
      val _ = writeln ("predicate: " ^ pred)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   393
      val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   394
      val _ = writeln ("introrules: ")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   395
      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   396
        (rev (intros_of thy pred)) ()
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   397
    in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   398
      if (has_elim thy pred) then
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   399
        writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   400
      else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   401
        writeln ("no elimrule defined")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   402
    end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   403
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   404
    fold print preds ()
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   405
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   406
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   407
fun print_all_modes thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   408
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   409
    val _ = writeln ("Inferred modes:")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   410
    fun print (pred, modes) u =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   411
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   412
        val _ = writeln ("predicate: " ^ pred)
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   413
        val _ = writeln ("modes: " ^ (commas (map (string_of_mode thy pred) modes)))
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   414
      in u end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   415
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   416
    fold print (all_modes_of thy) ()
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   417
  end
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   418
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
   419
(* validity checks *)
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
   420
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   421
fun check_expected_modes preds (options : Predicate_Compile_Aux.options) modes =
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   422
      case expected_modes options of
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   423
      SOME (s, ms) => (case AList.lookup (op =) modes s of
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   424
        SOME modes =>
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   425
          let
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   426
            val modes' = map (translate_mode (the (AList.lookup (op =) preds s))) modes
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   427
          in
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   428
            if not (eq_set eq_mode' (ms, modes')) then
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   429
              error ("expected modes were not inferred:\n"
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   430
              ^ "  inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode' modes')  ^ "\n"
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   431
              ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode' ms))
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   432
            else ()
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   433
          end
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   434
        | NONE => ())
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
   435
    | NONE => ()
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
   436
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   437
(* importing introduction rules *)
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   438
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   439
fun unify_consts thy cs intr_ts =
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   440
  (let
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   441
     val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   442
     fun varify (t, (i, ts)) =
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   443
       let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify [] t))
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   444
       in (maxidx_of_term t', t'::ts) end;
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
   445
     val (i, cs') = List.foldr varify (~1, []) cs;
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
   446
     val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   447
     val rec_consts = fold add_term_consts_2 cs' [];
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   448
     val intr_consts = fold add_term_consts_2 intr_ts' [];
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   449
     fun unify (cname, cT) =
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33268
diff changeset
   450
       let val consts = map snd (filter (fn c => fst c = cname) intr_consts)
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   451
       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   452
     val (env, _) = fold unify rec_consts (Vartab.empty, i');
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   453
     val subst = map_types (Envir.norm_type env)
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   454
   in (map subst cs', map subst intr_ts')
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   455
   end) handle Type.TUNIFY =>
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   456
     (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   457
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   458
fun import_intros inp_pred nparams [] ctxt =
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   459
  let
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   460
    val ([outp_pred], ctxt') = Variable.import_terms false [inp_pred] ctxt
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   461
    val (paramTs, _) = chop nparams (binder_types (fastype_of outp_pred))
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   462
    val (param_names, ctxt'') = Variable.variant_fixes (map (fn i => "p" ^ (string_of_int i))
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   463
      (1 upto nparams)) ctxt'
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   464
    val params = map Free (param_names ~~ paramTs)
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   465
    in (((outp_pred, params), []), ctxt') end
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   466
  | import_intros inp_pred nparams (th :: ths) ctxt =
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   467
    let
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   468
      val ((_, [th']), ctxt') = Variable.import false [th] ctxt
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   469
      val thy = ProofContext.theory_of ctxt'
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   470
      val (pred, (params, args)) = strip_intro_concl nparams (prop_of th')
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   471
      val ho_args = filter (is_predT o fastype_of) args
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   472
      fun subst_of (pred', pred) =
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   473
        let
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   474
          val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   475
        in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   476
      fun instantiate_typ th =
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   477
        let
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   478
          val (pred', _) = strip_intro_concl 0 (prop_of th)
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   479
          val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   480
            error "Trying to instantiate another predicate" else ()
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   481
        in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   482
      fun instantiate_ho_args th =
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   483
        let
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   484
          val (_, (params', args')) = strip_intro_concl nparams (prop_of th)
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   485
          val ho_args' = map dest_Var (filter (is_predT o fastype_of) args')
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   486
        in Thm.certify_instantiate ([], map dest_Var params' ~~ params) th end
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   487
      val outp_pred =
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   488
        Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   489
      val ((_, ths'), ctxt1) =
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   490
        Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   491
    in
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   492
      (((outp_pred, params), th' :: ths'), ctxt1)
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   493
    end
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   494
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   495
(* generation of case rules from user-given introduction rules *)
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   496
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   497
fun mk_casesrule ctxt pred nparams introrules =
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   498
  let
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   499
    val (((pred, params), intros_th), ctxt1) = import_intros pred nparams introrules ctxt
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   500
    val intros = map prop_of intros_th
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   501
    val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   502
    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   503
    val (_, argsT) = chop nparams (binder_types (fastype_of pred))
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   504
    val (argnames, ctxt3) = Variable.variant_fixes
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   505
      (map (fn i => "a" ^ string_of_int i) (1 upto length argsT)) ctxt2
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   506
    val argvs = map2 (curry Free) argnames argsT
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   507
    fun mk_case intro =
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   508
      let
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   509
        val (_, (_, args)) = strip_intro_concl nparams intro
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   510
        val prems = Logic.strip_imp_prems intro
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   511
        val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args)
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   512
        val frees = (fold o fold_aterms)
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   513
          (fn t as Free _ =>
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   514
              if member (op aconv) params t then I else insert (op aconv) t
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   515
           | _ => I) (args @ prems) []
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   516
      in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   517
    val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   518
    val cases = map mk_case intros
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   519
  in Logic.list_implies (assm :: cases, prop) end;
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   520
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   521
(** preprocessing rules **)  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   522
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   523
fun imp_prems_conv cv ct =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   524
  case Thm.term_of ct of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   525
    Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   526
  | _ => Conv.all_conv ct
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   527
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   528
fun Trueprop_conv cv ct =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   529
  case Thm.term_of ct of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   530
    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   531
  | _ => error "Trueprop_conv"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   532
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   533
fun preprocess_intro thy rule =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   534
  Conv.fconv_rule
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   535
    (imp_prems_conv
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   536
      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   537
    (Thm.transfer thy rule)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   538
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   539
fun preprocess_elim thy nparams elimrule =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   540
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   541
    fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   542
       HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   543
     | replace_eqs t = t
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
   544
    val ctxt = ProofContext.init thy
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
   545
    val ((_, [elimrule]), ctxt') = Variable.import false [elimrule] ctxt
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
   546
    val prems = Thm.prems_of elimrule
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   547
    val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   548
    fun preprocess_case t =
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
   549
      let
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   550
       val params = Logic.strip_params t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   551
       val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   552
       val assums_hyp' = assums1 @ (map replace_eqs assums2)
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
   553
      in
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   554
       list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
   555
      end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   556
    val cases' = map preprocess_case (tl prems)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   557
    val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   558
    val bigeq = (Thm.symmetric (Conv.implies_concl_conv
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   559
      (MetaSimplifier.rewrite true [@{thm Predicate.eq_is_eq}])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   560
        (cterm_of thy elimrule')))
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
   561
    val tac = (fn _ => Skip_Proof.cheat_tac thy)    
33109
7025bc7a5054 changed elimination preprocessing due to an error with a JinjaThread predicate
bulwahn
parents: 33108
diff changeset
   562
    val eq = Goal.prove ctxt' [] [] (Logic.mk_equals ((Thm.prop_of elimrule), elimrule')) tac
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   563
  in
33109
7025bc7a5054 changed elimination preprocessing due to an error with a JinjaThread predicate
bulwahn
parents: 33108
diff changeset
   564
    Thm.equal_elim eq elimrule |> singleton (Variable.export ctxt' ctxt)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   565
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   566
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   567
(* special case: predicate with no introduction rule *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   568
fun noclause thy predname elim = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   569
  val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   570
  val Ts = binder_types T
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   571
  val names = Name.variant_list []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   572
        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   573
  val vs = map2 (curry Free) names Ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   574
  val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   575
  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   576
  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   577
  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   578
  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
33441
99a5f22a967d eliminated funny record patterns and made SML/NJ happy;
wenzelm
parents: 33377
diff changeset
   579
        (fn _ => etac @{thm FalseE} 1)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   580
  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
33441
99a5f22a967d eliminated funny record patterns and made SML/NJ happy;
wenzelm
parents: 33377
diff changeset
   581
        (fn _ => etac elim 1) 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   582
in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   583
  ([intro], elim)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   584
end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   585
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   586
fun expand_tuples_elim th = th
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   587
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   588
(* updaters *)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   589
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   590
fun apfst4 f (x1, x2, x3, x4) = (f x1, x2, x3, x4)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   591
fun apsnd4 f (x1, x2, x3, x4) = (x1, f x2, x3, x4)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   592
fun aptrd4 f (x1, x2, x3, x4) = (x1, x2, f x3, x4)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   593
fun apfourth4 f (x1, x2, x3, x4) = (x1, x2, x3, f x4)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   594
fun appair f g (x, y) = (f x, g x)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   595
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   596
val no_compilation = ((false, []), (false, []), (false, []), (false, []))
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   597
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   598
fun fetch_pred_data thy name =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   599
  case try (Inductive.the_inductive (ProofContext.init thy)) name of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   600
    SOME (info as (_, result)) => 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   601
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   602
        fun is_intro_of intro =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   603
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   604
            val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   605
          in (fst (dest_Const const) = name) end;      
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   606
        val intros = ind_set_codegen_preproc thy
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   607
          (map (expand_tuples thy #> preprocess_intro thy) (filter is_intro_of (#intrs result)))
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   608
        val index = find_index (fn s => s = name) (#names (fst info))
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   609
        val pre_elim = nth (#elims result) index
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   610
        val pred = nth (#preds result) index
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   611
        val nparams = length (Inductive.params_of (#raw_induct result))
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   612
        (*val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams 
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   613
          (expand_tuples_elim pre_elim))*)
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   614
        val elim =
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
   615
          (Drule.standard o Skip_Proof.make_thm thy)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   616
          (mk_casesrule (ProofContext.init thy) pred nparams intros)
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   617
        val (intros, elim) = (*if null intros then noclause thy name elim else*) (intros, elim)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   618
      in
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   619
        mk_pred_data ((intros, SOME elim, nparams), no_compilation)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   620
      end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   621
  | NONE => error ("No such predicate: " ^ quote name)
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   622
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   623
fun add_predfun name mode data =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   624
  let
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   625
    val add = (apsnd o apfst4) (fn (x, y) => (true, cons (mode, mk_predfun_data data) y))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   626
  in PredData.map (Graph.map_node name (map_pred_data add)) end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   627
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   628
fun is_inductive_predicate thy name =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   629
  is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   630
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   631
fun depending_preds_of thy (key, value) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   632
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   633
    val intros = (#intros o rep_pred_data) value
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   634
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   635
    fold Term.add_const_names (map Thm.prop_of intros) []
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   636
      |> filter (fn c => (not (c = key)) andalso
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   637
        (is_inductive_predicate thy c orelse is_registered thy c))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   638
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   639
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   640
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   641
(* code dependency graph *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   642
(*
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   643
fun dependencies_of thy name =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   644
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   645
    val (intros, elim, nparams) = fetch_pred_data thy name 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   646
    val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   647
    val keys = depending_preds_of thy intros
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   648
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   649
    (data, keys)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   650
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   651
*)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   652
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   653
fun add_intro thm thy = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   654
   val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   655
   fun cons_intro gr =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   656
     case try (Graph.get_node gr) name of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   657
       SOME pred_data => Graph.map_node name (map_pred_data
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
   658
         (apfst (fn (intros, elim, nparams) => (intros @ [thm], elim, nparams)))) gr
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   659
     | NONE =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   660
       let
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   661
         val nparams = the_default (guess_nparams T)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   662
           (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   663
       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), no_compilation)) gr end;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   664
  in PredData.map cons_intro thy end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   665
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   666
fun set_elim thm = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   667
    val (name, _) = dest_Const (fst 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   668
      (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   669
    fun set (intros, _, nparams) = (intros, SOME thm, nparams)  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   670
  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   671
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   672
fun set_nparams name nparams = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   673
    fun set (intros, elim, _ ) = (intros, elim, nparams) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   674
  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   675
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   676
fun register_predicate (constname, pre_intros, pre_elim, nparams) thy =
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   677
  let
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   678
    (* preprocessing *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   679
    val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   680
    val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   681
  in
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   682
    if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   683
      PredData.map
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   684
        (Graph.new_node (constname,
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   685
          mk_pred_data ((intros, SOME elim, nparams), no_compilation))) thy
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   686
    else thy
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   687
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   688
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   689
fun register_intros (constname, pre_intros) thy =
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   690
  let
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   691
    val T = Sign.the_const_type thy constname
33120
ca77d8c34ce2 cleaned up
bulwahn
parents: 33118
diff changeset
   692
    fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl 0 (prop_of intr))))
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   693
    val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   694
      error ("register_intros: Introduction rules of different constants are used\n" ^
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   695
        "expected rules for " ^ constname ^ ", but received rules for " ^
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   696
          commas (map constname_of_intro pre_intros))
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   697
      else ()
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   698
    val pred = Const (constname, T)
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   699
    val nparams = guess_nparams T
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   700
    val pre_elim = 
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
   701
      (Drule.standard o Skip_Proof.make_thm thy)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   702
      (mk_casesrule (ProofContext.init thy) pred nparams pre_intros)
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   703
  in register_predicate (constname, pre_intros, pre_elim, nparams) thy end
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   704
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
   705
fun set_random_function_name pred mode name = 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   706
  let
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   707
    val set = (apsnd o apsnd4) (fn (x, y) => (true, cons (mode, mk_function_data (name, NONE)) y))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   708
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   709
    PredData.map (Graph.map_node pred (map_pred_data set))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   710
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   711
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
   712
fun set_depth_limited_function_name pred mode name = 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   713
  let
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   714
    val set = (apsnd o aptrd4) (fn (x, y) => (true, cons (mode, mk_function_data (name, NONE)) y))
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   715
  in
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   716
    PredData.map (Graph.map_node pred (map_pred_data set))
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   717
  end
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   718
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   719
fun set_annotated_function_name pred mode name =
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   720
  let
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   721
    val set = (apsnd o apfourth4)
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
   722
      (fn (x, y) => (true, cons (mode, mk_function_data (name, NONE)) y))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   723
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   724
    PredData.map (Graph.map_node pred (map_pred_data set))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   725
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   726
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   727
datatype compilation_funs = CompilationFuns of {
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   728
  mk_predT : typ -> typ,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   729
  dest_predT : typ -> typ,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   730
  mk_bot : typ -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   731
  mk_single : term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   732
  mk_bind : term * term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   733
  mk_sup : term * term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   734
  mk_if : term -> term,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   735
  mk_not : term -> term,
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   736
  mk_map : typ -> typ -> term -> term -> term
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   737
};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   738
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   739
fun mk_predT (CompilationFuns funs) = #mk_predT funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   740
fun dest_predT (CompilationFuns funs) = #dest_predT funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   741
fun mk_bot (CompilationFuns funs) = #mk_bot funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   742
fun mk_single (CompilationFuns funs) = #mk_single funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   743
fun mk_bind (CompilationFuns funs) = #mk_bind funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   744
fun mk_sup (CompilationFuns funs) = #mk_sup funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   745
fun mk_if (CompilationFuns funs) = #mk_if funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   746
fun mk_not (CompilationFuns funs) = #mk_not funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   747
fun mk_map (CompilationFuns funs) = #mk_map funs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   748
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   749
structure PredicateCompFuns =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   750
struct
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   751
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   752
fun mk_predT T = Type (@{type_name Predicate.pred}, [T])
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   753
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   754
fun dest_predT (Type (@{type_name Predicate.pred}, [T])) = T
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   755
  | dest_predT T = raise TYPE ("dest_predT", [T], []);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   756
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   757
fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   758
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   759
fun mk_single t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   760
  let val T = fastype_of t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   761
  in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   762
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   763
fun mk_bind (x, f) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   764
  let val T as Type ("fun", [_, U]) = fastype_of f
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   765
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   766
    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   767
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   768
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   769
val mk_sup = HOLogic.mk_binop @{const_name sup};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   770
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   771
fun mk_if cond = Const (@{const_name Predicate.if_pred},
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   772
  HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   773
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   774
fun mk_not t = let val T = mk_predT HOLogic.unitT
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   775
  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   776
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   777
fun mk_Enum f =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   778
  let val T as Type ("fun", [T', _]) = fastype_of f
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   779
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   780
    Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   781
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   782
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   783
fun mk_Eval (f, x) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   784
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   785
    val T = fastype_of x
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   786
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   787
    Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   788
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   789
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   790
fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   791
  (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   792
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   793
val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   794
  mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   795
  mk_map = mk_map};
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   796
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   797
end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   798
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   799
structure RandomPredCompFuns =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   800
struct
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   801
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   802
fun mk_randompredT T =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   803
  @{typ Random.seed} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ Random.seed})
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   804
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   805
fun dest_randompredT (Type ("fun", [@{typ Random.seed}, Type (@{type_name "*"},
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   806
  [Type (@{type_name "Predicate.pred"}, [T]), @{typ Random.seed}])])) = T
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   807
  | dest_randompredT T = raise TYPE ("dest_randompredT", [T], []);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   808
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   809
fun mk_bot T = Const(@{const_name Quickcheck.empty}, mk_randompredT T)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   810
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   811
fun mk_single t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   812
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   813
    val T = fastype_of t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   814
  in
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   815
    Const (@{const_name Quickcheck.single}, T --> mk_randompredT T) $ t
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   816
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   817
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   818
fun mk_bind (x, f) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   819
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   820
    val T as (Type ("fun", [_, U])) = fastype_of f
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   821
  in
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   822
    Const (@{const_name Quickcheck.bind}, fastype_of x --> T --> U) $ x $ f
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   823
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   824
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   825
val mk_sup = HOLogic.mk_binop @{const_name Quickcheck.union}
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   826
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   827
fun mk_if cond = Const (@{const_name Quickcheck.if_randompred},
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   828
  HOLogic.boolT --> mk_randompredT HOLogic.unitT) $ cond;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   829
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   830
fun mk_not t = let val T = mk_randompredT HOLogic.unitT
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   831
  in Const (@{const_name Quickcheck.not_randompred}, T --> T) $ t end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   832
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   833
fun mk_map T1 T2 tf tp = Const (@{const_name Quickcheck.map},
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   834
  (T1 --> T2) --> mk_randompredT T1 --> mk_randompredT T2) $ tf $ tp
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   835
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   836
val compfuns = CompilationFuns {mk_predT = mk_randompredT, dest_predT = dest_randompredT,
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   837
    mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if,
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   838
    mk_not = mk_not, mk_map = mk_map};
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   839
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   840
end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   841
(* for external use with interactive mode *)
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   842
val pred_compfuns = PredicateCompFuns.compfuns
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   843
val randompred_compfuns = RandomPredCompFuns.compfuns;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   844
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   845
fun lift_random random =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   846
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   847
    val T = dest_randomT (fastype_of random)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   848
  in
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   849
    Const (@{const_name Quickcheck.Random}, (@{typ Random.seed} -->
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   850
      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
   851
      RandomPredCompFuns.mk_randompredT T) $ random
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   852
  end;
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   853
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   854
(* function types and names of different compilations *)
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   855
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   856
fun funT_of compfuns (iss, is) T =
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   857
  let
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   858
    val Ts = binder_types T
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   859
    val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   860
    val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   861
  in
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   862
    (paramTs' @ inargTs) ---> (mk_predT compfuns (HOLogic.mk_tupleT outargTs))
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   863
  end;
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   864
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
   865
fun depth_limited_funT_of compfuns (iss, is) T =
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   866
  let
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   867
    val Ts = binder_types T
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   868
    val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   869
    val paramTs' =
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   870
      map2 (fn SOME is => depth_limited_funT_of compfuns ([], is) | NONE => I) iss paramTs
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   871
  in
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
   872
    (paramTs' @ inargTs @ [@{typ bool}, @{typ "code_numeral"}])
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
   873
      ---> (mk_predT compfuns (HOLogic.mk_tupleT outargTs))
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   874
  end;
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   875
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   876
fun random_function_funT_of (iss, is) T =
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   877
  let
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   878
    val Ts = binder_types T
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   879
    val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   880
    val paramTs' = map2 (fn SOME is => random_function_funT_of ([], is) | NONE => I) iss paramTs
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   881
  in
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   882
    (paramTs' @ inargTs @ [@{typ code_numeral}]) --->
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   883
      (mk_predT RandomPredCompFuns.compfuns (HOLogic.mk_tupleT outargTs))
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
   884
  end
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
   885
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   886
(* Mode analysis *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   887
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   888
(*** check if a term contains only constructor functions ***)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   889
fun is_constrt thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   890
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   891
    val cnstrs = flat (maps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   892
      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   893
      (Symtab.dest (Datatype.get_all thy)));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   894
    fun check t = (case strip_comb t of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   895
        (Free _, []) => true
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   896
      | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   897
            (SOME (i, Tname), Type (Tname', _)) =>
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   898
              length ts = i andalso Tname = Tname' andalso forall check ts
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   899
          | _ => false)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   900
      | _ => false)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   901
  in check end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   902
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   903
(*** check if a type is an equality type (i.e. doesn't contain fun)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   904
  FIXME this is only an approximation ***)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   905
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   906
  | is_eqT _ = true;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   907
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   908
fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   909
val terms_vs = distinct (op =) o maps term_vs;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   910
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   911
(** collect all Frees in a term (with duplicates!) **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   912
fun term_vTs tm =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   913
  fold_aterms (fn Free xT => cons xT | _ => I) tm [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   914
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   915
fun subsets i j =
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   916
  if i <= j then
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   917
    let
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   918
      fun merge xs [] = xs
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   919
        | merge [] ys = ys
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   920
        | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   921
            else y::merge (x::xs) ys;
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   922
      val is = subsets (i+1) j
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   923
    in merge (map (fn ks => i::ks) is) is end
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   924
  else [[]];
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   925
     
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   926
(* FIXME: should be in library - cprod = map_prod I *)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   927
fun cprod ([], ys) = []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   928
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   929
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32970
diff changeset
   930
fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   931
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   932
fun cprods_subset [] = [[]]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   933
  | cprods_subset (xs :: xss) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   934
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   935
    val yss = (cprods_subset xss)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   936
  in maps (fn ys => map (fn x => cons x ys) xs) yss @ yss end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   937
  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   938
fun modes_of_term modes t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   939
  let
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   940
    val ks = map_index (fn (i, T) => (i + 1, NONE)) (binder_types (fastype_of t));
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   941
    val default = [Mode (([], ks), ks, [])];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   942
    fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   943
        let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   944
          val (args1, args2) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   945
            if length args < length iss then
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   946
              error ("Too few arguments for inductive predicate " ^ name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   947
            else chop (length iss) args;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   948
          val k = length args2;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   949
          val prfx = map (rpair NONE) (1 upto k)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   950
        in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   951
          if not (is_prefix op = prfx is) then [] else
33116
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33115
diff changeset
   952
          let val is' = map (fn (i, t) => (i - k, t)) (List.drop (is, k))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   953
          in map (fn x => Mode (m, is', x)) (cprods (map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   954
            (fn (NONE, _) => [NONE]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   955
              | (SOME js, arg) => map SOME (filter
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   956
                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   957
                    (iss ~~ args1)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   958
          end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   959
        end)) (AList.lookup op = modes name)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   960
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   961
    case strip_comb (Envir.eta_contract t) of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   962
      (Const (name, _), args) => the_default default (mk_modes name args)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   963
    | (Var ((name, _), _), args) => the (mk_modes name args)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   964
    | (Free (name, _), args) => the (mk_modes name args)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   965
    | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   966
    | _ => default
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   967
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   968
  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   969
fun select_mode_prem thy modes vs ps =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   970
  find_first (is_some o snd) (ps ~~ map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   971
    (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   972
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   973
            val (in_ts, out_ts) = split_smode is us;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   974
            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   975
            val vTs = maps term_vTs out_ts';
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   976
            val dupTs = map snd (duplicates (op =) vTs) @
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
   977
              map_filter (AList.lookup (op =) vTs) vs;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   978
          in
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   979
            subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   980
            forall (is_eqT o fastype_of) in_ts' andalso
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   981
            subset (op =) (term_vs t, vs) andalso
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   982
            forall is_eqT dupTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   983
          end)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   984
            (modes_of_term modes t handle Option =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   985
               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   986
      | Negprem (us, t) => find_first (fn Mode (_, is, _) =>
33149
2c8f1c450b47 merge -- imported from bulwahn d759e2728188;
wenzelm
parents: 33043 33148
diff changeset
   987
            is = map (rpair NONE) (1 upto length us) andalso
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   988
            subset (op =) (terms_vs us, vs) andalso
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   989
            subset (op =) (term_vs t, vs))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   990
            (modes_of_term modes t handle Option =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   991
               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   992
      | Sidecond t => if subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], []))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   993
          else NONE
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   994
      ) ps);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   995
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   996
fun fold_prem f (Prem (args, _)) = fold f args
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   997
  | fold_prem f (Negprem (args, _)) = fold f args
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   998
  | fold_prem f (Sidecond t) = f t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   999
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1000
fun all_subsets [] = [[]]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1001
  | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1002
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1003
fun generator vTs v = 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1004
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1005
    val T = the (AList.lookup (op =) vTs v)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1006
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1007
    (Generator (v, T), Mode (([], []), [], []))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1008
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1009
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1010
fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1011
  let
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  1012
    val modes' = modes @ map_filter
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1013
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1014
        (param_vs ~~ iss);
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  1015
    val gen_modes' = gen_modes @ map_filter
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1016
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1017
        (param_vs ~~ iss);  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1018
    val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1019
    val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1020
    fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1021
      | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1022
          NONE =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1023
            (if with_generator then
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1024
              (case select_mode_prem thy gen_modes' vs ps of
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
  1025
                SOME (p as Prem _, SOME mode) => check_mode_prems ((p, mode) :: acc_ps) 
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
  1026
                  (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1027
                  (filter_out (equal p) ps)
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  1028
              | _ =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1029
                  let 
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1030
                    val all_generator_vs = all_subsets (subtract (op =) vs prem_vs)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1031
                      |> sort (int_ord o (pairself length))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1032
                  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1033
                    case (find_first (fn generator_vs => is_some
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1034
                      (select_mode_prem thy modes' (union (op =) vs generator_vs) ps))
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1035
                        all_generator_vs) of
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1036
                      SOME generator_vs => check_mode_prems
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1037
                        ((map (generator vTs) generator_vs) @ acc_ps)
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
  1038
                        (union (op =) vs generator_vs) ps
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
  1039
                    | NONE => NONE
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1040
                  end)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1041
            else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1042
              NONE)
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
  1043
        | SOME (p, SOME mode) => check_mode_prems ((p, mode) :: acc_ps) 
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
  1044
            (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1045
            (filter_out (equal p) ps))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1046
    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1047
    val in_vs = terms_vs in_ts;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1048
    val concl_vs = terms_vs ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1049
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1050
    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1051
    forall (is_eqT o fastype_of) in_ts' then
33042
ddf1f03a9ad9 curried union as canonical list operation
haftmann
parents: 33039
diff changeset
  1052
      case check_mode_prems [] (union (op =) param_vs in_vs) ps of
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1053
         NONE => NONE
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1054
       | SOME (acc_ps, vs) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1055
         if with_generator then
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33039
diff changeset
  1056
           SOME (ts, (rev acc_ps) @ (map (generator vTs) (subtract (op =) vs concl_vs)))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1057
         else
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
  1058
           if subset (op =) (concl_vs, vs) then SOME (ts, rev acc_ps) else NONE
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1059
    else NONE
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1060
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1061
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1062
fun print_failed_mode options thy modes p m rs i =
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1063
  if show_mode_inference options then
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1064
    let
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
  1065
      val _ = tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
  1066
      p ^ " violates mode " ^ string_of_mode thy p m)
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
  1067
      val _ = tracing (string_of_clause thy p (nth rs i))
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1068
    in () end
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1069
  else ()
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1070
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1071
fun check_modes_pred options with_generator thy param_vs clauses modes gen_modes (p, ms) =
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1072
  let
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1073
    val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33268
diff changeset
  1074
  in (p, filter (fn m => case find_index
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1075
    (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1076
      ~1 => true
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1077
    | i => (print_failed_mode options thy modes p m rs i; false)) ms)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1078
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1079
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1080
fun get_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1081
  let
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1082
    val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1083
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1084
    (p, map (fn m =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1085
      (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1086
  end;
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  1087
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1088
fun fixp f (x : (string * mode list) list) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1089
  let val y = f x
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1090
  in if x = y then x else fixp f y end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1091
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1092
fun infer_modes options thy extra_modes all_modes param_vs clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1093
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1094
    val modes =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1095
      fixp (fn modes =>
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1096
        map (check_modes_pred options false thy param_vs clauses (modes @ extra_modes) []) modes)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1097
          all_modes
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1098
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1099
    map (get_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1100
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1101
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1102
fun remove_from rem [] = []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1103
  | remove_from rem ((k, vs) :: xs) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1104
    (case AList.lookup (op =) rem k of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1105
      NONE => (k, vs)
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33039
diff changeset
  1106
    | SOME vs' => (k, subtract (op =) vs' vs))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1107
    :: remove_from rem xs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1108
    
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1109
fun infer_modes_with_generator options thy extra_modes all_modes param_vs clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1110
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1111
    val prednames = map fst clauses
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1112
    val extra_modes' = all_modes_of thy
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
  1113
    val gen_modes = all_random_modes_of thy
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1114
      |> filter_out (fn (name, _) => member (op =) prednames name)
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1115
    val starting_modes = remove_from extra_modes' all_modes
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1116
    fun eq_mode (m1, m2) = (m1 = m2)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1117
    val modes =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1118
      fixp (fn modes =>
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1119
        map (check_modes_pred options true thy param_vs clauses extra_modes'
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1120
          (gen_modes @ modes)) modes) starting_modes
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1121
  in
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1122
    AList.join (op =)
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1123
    (fn _ => fn ((mps1, mps2)) =>
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1124
      merge (fn ((m1, _), (m2, _)) => eq_mode (m1, m2)) (mps1, mps2))
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1125
    (infer_modes options thy extra_modes all_modes param_vs clauses,
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1126
    map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1127
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1128
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1129
(* term construction *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1130
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1131
fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1132
      NONE => (Free (s, T), (names, (s, [])::vs))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1133
    | SOME xs =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1134
        let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1135
          val s' = Name.variant names s;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1136
          val v = Free (s', T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1137
        in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1138
          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1139
        end);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1140
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1141
fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1142
  | distinct_v (t $ u) nvs =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1143
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1144
        val (t', nvs') = distinct_v t nvs;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1145
        val (u', nvs'') = distinct_v u nvs';
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1146
      in (t' $ u', nvs'') end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1147
  | distinct_v x nvs = (x, nvs);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1148
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1149
(** specific rpred functions -- move them to the correct place in this file *)
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1150
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1151
fun mk_Eval_of additional_arguments ((x, T), NONE) names = (x, names)
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1152
  | mk_Eval_of additional_arguments ((x, T), SOME mode) names =
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1153
  let
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1154
    val Ts = binder_types T
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1155
    (*val argnames = Name.variant_list names
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1156
        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1157
    val args = map Free (argnames ~~ Ts)
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1158
    val (inargs, outargs) = split_smode mode args*)
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1159
    fun mk_split_lambda [] t = lambda (Free (Name.variant names "x", HOLogic.unitT)) t
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1160
      | mk_split_lambda [x] t = lambda x t
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1161
      | mk_split_lambda xs t =
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1162
      let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1163
        fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1164
          | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1165
      in
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1166
        mk_split_lambda' xs t
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1167
      end;
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1168
    fun mk_arg (i, T) =
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1169
      let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1170
        val vname = Name.variant names ("x" ^ string_of_int i)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1171
        val default = Free (vname, T)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1172
      in 
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1173
        case AList.lookup (op =) mode i of
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1174
          NONE => (([], [default]), [default])
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1175
        | SOME NONE => (([default], []), [default])
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1176
        | SOME (SOME pis) =>
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1177
          case HOLogic.strip_tupleT T of
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1178
            [] => error "pair mode but unit tuple" (*(([default], []), [default])*)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1179
          | [_] => error "pair mode but not a tuple" (*(([default], []), [default])*)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1180
          | Ts =>
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1181
            let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1182
              val vnames = Name.variant_list names
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1183
                (map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1184
                  (1 upto length Ts))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1185
              val args = map Free (vnames ~~ Ts)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1186
              fun split_args (i, arg) (ins, outs) =
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1187
                if member (op =) pis i then
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1188
                  (arg::ins, outs)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1189
                else
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1190
                  (ins, arg::outs)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1191
              val (inargs, outargs) = fold_rev split_args ((1 upto length Ts) ~~ args) ([], [])
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1192
              fun tuple args = if null args then [] else [HOLogic.mk_tuple args]
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1193
            in ((tuple inargs, tuple outargs), args) end
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1194
      end
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1195
    val (inoutargs, args) = split_list (map mk_arg (1 upto (length Ts) ~~ Ts))
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1196
    val (inargs, outargs) = pairself flat (split_list inoutargs)
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1197
    val r = PredicateCompFuns.mk_Eval 
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1198
      (list_comb (x, inargs @ additional_arguments), HOLogic.mk_tuple outargs)
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1199
    val t = fold_rev mk_split_lambda args r
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1200
  in
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1201
    (t, names)
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1202
  end;
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1203
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1204
structure Comp_Mod =
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1205
struct
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1206
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1207
datatype comp_modifiers = Comp_Modifiers of
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1208
{
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1209
  function_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string,
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1210
  set_function_name : string -> Predicate_Compile_Aux.mode -> string -> theory -> theory,
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1211
  function_name_prefix : string,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1212
  funT_of : compilation_funs -> mode -> typ -> typ,
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1213
  additional_arguments : string list -> term list,
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1214
  wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term,
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1215
  transform_additional_arguments : indprem -> term list -> term list
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1216
}
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1217
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1218
fun dest_comp_modifiers (Comp_Modifiers c) = c
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1219
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1220
val function_name_of = #function_name_of o dest_comp_modifiers
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1221
val set_function_name = #set_function_name o dest_comp_modifiers
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1222
val function_name_prefix = #function_name_prefix o dest_comp_modifiers
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1223
val funT_of = #funT_of o dest_comp_modifiers
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1224
val additional_arguments = #additional_arguments o dest_comp_modifiers
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1225
val wrap_compilation = #wrap_compilation o dest_comp_modifiers
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1226
val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1227
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1228
end;
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1229
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1230
fun compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss arg = 
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1231
  let
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1232
    fun map_params (t as Free (f, T)) =
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1233
      if member (op =) param_vs f then
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1234
        case (the (AList.lookup (op =) (param_vs ~~ iss) f)) of
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1235
          SOME is =>
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1236
            let
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1237
              val T' = Comp_Mod.funT_of compilation_modifiers compfuns ([], is) T
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1238
            in fst (mk_Eval_of additional_arguments ((Free (f, T'), T), SOME is) []) end
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1239
        | NONE => t
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1240
      else t
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1241
      | map_params t = t
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1242
    in map_aterms map_params arg end
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1243
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1244
fun compile_match compilation_modifiers compfuns additional_arguments
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1245
  param_vs iss thy eqs eqs' out_ts success_t =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1246
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1247
    val eqs'' = maps mk_eq eqs @ eqs'
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1248
    val eqs'' =
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1249
      map (compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss) eqs''
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1250
    val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1251
    val name = Name.variant names "x";
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1252
    val name' = Name.variant (name :: names) "y";
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1253
    val T = HOLogic.mk_tupleT (map fastype_of out_ts);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1254
    val U = fastype_of success_t;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1255
    val U' = dest_predT compfuns U;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1256
    val v = Free (name, T);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1257
    val v' = Free (name', T);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1258
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1259
    lambda v (fst (Datatype.make_case
32671
fbd224850767 adapted configuration for DatatypeCase.make_case
bulwahn
parents: 32669
diff changeset
  1260
      (ProofContext.init thy) DatatypeCase.Quiet [] v
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1261
      [(HOLogic.mk_tuple out_ts,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1262
        if null eqs'' then success_t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1263
        else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1264
          foldr1 HOLogic.mk_conj eqs'' $ success_t $
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1265
            mk_bot compfuns U'),
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1266
       (v', mk_bot compfuns U')]))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1267
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1268
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1269
(*FIXME function can be removed*)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1270
fun mk_funcomp f t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1271
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1272
    val names = Term.add_free_names t [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1273
    val Ts = binder_types (fastype_of t);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1274
    val vs = map Free
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1275
      (Name.variant_list names (replicate (length Ts) "x") ~~ Ts)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1276
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1277
    fold_rev lambda vs (f (list_comb (t, vs)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1278
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1279
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1280
fun compile_param compilation_modifiers compfuns thy (NONE, t) = t
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1281
  | compile_param compilation_modifiers compfuns thy (m as SOME (Mode (mode, _, ms)), t) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1282
   let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1283
     val (f, args) = strip_comb (Envir.eta_contract t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1284
     val (params, args') = chop (length ms) args
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1285
     val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1286
     val f' =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1287
       case f of
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1288
         Const (name, T) => Const (Comp_Mod.function_name_of compilation_modifiers thy name mode,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1289
           Comp_Mod.funT_of compilation_modifiers compfuns mode T)
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1290
       | Free (name, T) => Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1291
       | _ => error ("PredicateCompiler: illegal parameter term")
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  1292
   in
33133
2eb7dfcf3bc3 simplified and improved compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33132
diff changeset
  1293
     list_comb (f', params' @ args')
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  1294
   end
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  1295
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1296
fun compile_expr compilation_modifiers compfuns thy ((Mode (mode, _, ms)), t)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1297
  inargs additional_arguments =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1298
  case strip_comb t of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1299
    (Const (name, T), params) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1300
       let
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1301
         val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1302
         val name' = Comp_Mod.function_name_of compilation_modifiers thy name mode
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1303
         val T' = Comp_Mod.funT_of compilation_modifiers compfuns mode T
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1304
       in
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1305
         (list_comb (Const (name', T'), params' @ inargs @ additional_arguments))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1306
       end
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  1307
  | (Free (name, T), params) =>
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1308
    list_comb (Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T),
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1309
      params @ inargs @ additional_arguments)
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33144
diff changeset
  1310
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1311
fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1312
  (iss, is) inp (ts, moded_ps) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1313
  let
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1314
    val compile_match = compile_match compilation_modifiers compfuns
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1315
      additional_arguments param_vs iss thy
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1316
    fun check_constrt t (names, eqs) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1317
      if is_constrt thy t then (t, (names, eqs)) else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1318
        let
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
  1319
          val s = Name.variant names "x"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1320
          val v = Free (s, fastype_of t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1321
        in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1322
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1323
    val (in_ts, out_ts) = split_smode is ts;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1324
    val (in_ts', (all_vs', eqs)) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1325
      fold_map check_constrt in_ts (all_vs, []);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1326
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1327
    fun compile_prems out_ts' vs names [] =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1328
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1329
            val (out_ts'', (names', eqs')) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1330
              fold_map check_constrt out_ts' (names, []);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1331
            val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1332
              out_ts'' (names', map (rpair []) vs);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1333
          in
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1334
            compile_match constr_vs (eqs @ eqs') out_ts'''
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1335
              (mk_single compfuns (HOLogic.mk_tuple out_ts))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1336
          end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1337
      | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1338
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1339
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1340
            val (out_ts', (names', eqs)) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1341
              fold_map check_constrt out_ts (names, [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1342
            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1343
              out_ts' ((names', map (rpair []) vs))
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1344
            val additional_arguments' =
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1345
              Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1346
            val (compiled_clause, rest) = case p of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1347
               Prem (us, t) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1348
                 let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1349
                   val (in_ts, out_ts''') = split_smode is us;
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1350
                   val in_ts = map (compile_arg compilation_modifiers compfuns
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1351
                     additional_arguments thy param_vs iss) in_ts
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1352
                   val u =
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1353
                     compile_expr compilation_modifiers compfuns thy
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1354
                       (mode, t) in_ts additional_arguments'
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1355
                   val rest = compile_prems out_ts''' vs' names'' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1356
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1357
                   (u, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1358
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1359
             | Negprem (us, t) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1360
                 let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1361
                   val (in_ts, out_ts''') = split_smode is us
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1362
                   val in_ts = map (compile_arg compilation_modifiers compfuns
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1363
                     additional_arguments thy param_vs iss) in_ts
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1364
                   val u = mk_not compfuns
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1365
                     (compile_expr compilation_modifiers compfuns thy
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1366
                       (mode, t) in_ts additional_arguments')
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1367
                   val rest = compile_prems out_ts''' vs' names'' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1368
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1369
                   (u, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1370
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1371
             | Sidecond t =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1372
                 let
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1373
                   val t = compile_arg compilation_modifiers compfuns additional_arguments
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1374
                     thy param_vs iss t
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1375
                   val rest = compile_prems [] vs' names'' ps;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1376
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1377
                   (mk_if compfuns t, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1378
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1379
             | Generator (v, T) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1380
                 let
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1381
                   val [size] = additional_arguments
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1382
                   val u = lift_random (HOLogic.mk_random T size)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1383
                   val rest = compile_prems [Free (v, T)]  vs' names'' ps;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1384
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1385
                   (u, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1386
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1387
          in
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1388
            compile_match constr_vs' eqs out_ts''
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1389
              (mk_bind compfuns (compiled_clause, rest))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1390
          end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1391
    val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1392
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1393
    mk_bind compfuns (mk_single compfuns inp, prem_t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1394
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1395
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1396
fun compile_pred compilation_modifiers compfuns thy all_vs param_vs s T mode moded_cls =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1397
  let
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1398
    val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1399
    val (Us1, Us2) = split_smodeT (snd mode) Ts2
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1400
    val Ts1' =
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1401
      map2 (fn NONE => I | SOME is => Comp_Mod.funT_of compilation_modifiers compfuns ([], is))
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1402
        (fst mode) Ts1
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1403
    fun mk_input_term (i, NONE) =
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1404
        [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1405
      | mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1406
               [] => error "strange unit input"
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1407
             | [T] => [Free (Name.variant (all_vs @ param_vs)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1408
               ("x" ^ string_of_int i), nth Ts2 (i - 1))]
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1409
             | Ts => let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1410
               val vnames = Name.variant_list (all_vs @ param_vs)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1411
                (map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1412
                  pis)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1413
             in if null pis then []
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1414
               else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1415
    val in_ts = maps mk_input_term (snd mode)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1416
    val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1417
    val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1418
      (all_vs @ param_vs)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1419
    val cl_ts =
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1420
      map (compile_clause compilation_modifiers compfuns
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1421
        thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts)) moded_cls;
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1422
    val compilation = Comp_Mod.wrap_compilation compilation_modifiers compfuns
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1423
      s T mode additional_arguments
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1424
      (if null cl_ts then
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1425
        mk_bot compfuns (HOLogic.mk_tupleT Us2)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1426
      else foldr1 (mk_sup compfuns) cl_ts)
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1427
    val fun_const =
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1428
      Const (Comp_Mod.function_name_of compilation_modifiers thy s mode,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1429
        Comp_Mod.funT_of compilation_modifiers compfuns mode T)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1430
  in
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1431
    HOLogic.mk_Trueprop
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1432
      (HOLogic.mk_eq (list_comb (fun_const, params @ in_ts @ additional_arguments), compilation))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1433
  end;
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1434
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1435
(* special setup for simpset *)                  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1436
val HOL_basic_ss' = HOL_basic_ss addsimps (@{thms "HOL.simp_thms"} @ [@{thm Pair_eq}])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1437
  setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1438
  setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1439
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1440
(* Definition of executable functions and their intro and elim rules *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1441
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1442
fun print_arities arities = tracing ("Arities:\n" ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1443
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1444
    space_implode " -> " (map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1445
      (fn NONE => "X" | SOME k' => string_of_int k')
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1446
        (ks @ [SOME k]))) arities));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1447
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1448
fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1449
let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1450
  val Ts = binder_types (fastype_of pred)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1451
  val funtrm = Const (mode_id, funT)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1452
  val (Ts1, Ts2) = chop (length iss) Ts;
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1453
  val Ts1' =
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1454
    map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1455
  val param_names = Name.variant_list []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1456
    (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1)));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1457
  val params = map Free (param_names ~~ Ts1')
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1458
  fun mk_args (i, T) argnames =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1459
    let
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1460
      val vname = Name.variant (param_names @ argnames) ("x" ^ string_of_int (length Ts1' + i))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1461
      val default = (Free (vname, T), vname :: argnames)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1462
    in
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1463
      case AList.lookup (op =) is i of
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1464
             NONE => default
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1465
           | SOME NONE => default
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1466
           | SOME (SOME pis) =>
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1467
             case HOLogic.strip_tupleT T of
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1468
               [] => default
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1469
             | [_] => default
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1470
             | Ts => 
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1471
            let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1472
              val vnames = Name.variant_list (param_names @ argnames)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1473
                (map (fn j => "x" ^ string_of_int (length Ts1' + i) ^ "p" ^ string_of_int j)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1474
                  (1 upto (length Ts)))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1475
             in (HOLogic.mk_tuple (map Free (vnames ~~ Ts)), vnames  @ argnames) end
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1476
    end
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1477
  val (args, argnames) = fold_map mk_args (1 upto (length Ts2) ~~ Ts2) []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1478
  val (inargs, outargs) = split_smode is args
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1479
  val param_names' = Name.variant_list (param_names @ argnames)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1480
    (map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1481
  val param_vs = map Free (param_names' ~~ Ts1)
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1482
  val (params', names) = fold_map (mk_Eval_of []) ((params ~~ Ts1) ~~ iss) []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1483
  val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ args))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1484
  val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ args))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1485
  val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1486
  val funargs = params @ inargs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1487
  val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1488
                  if null outargs then Free("y", HOLogic.unitT) else HOLogic.mk_tuple outargs))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1489
  val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1490
                   HOLogic.mk_tuple outargs))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1491
  val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1492
  val simprules = [defthm, @{thm eval_pred},
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1493
    @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1494
  val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1495
  val introthm = Goal.prove (ProofContext.init thy)
33487
bulwahn
parents: 33485 33441
diff changeset
  1496
    (argnames @ param_names @ param_names' @ ["y"]) [] introtrm (fn _ => unfolddef_tac)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1497
  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1498
  val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1499
  val elimthm = Goal.prove (ProofContext.init thy)
33487
bulwahn
parents: 33485 33441
diff changeset
  1500
    (argnames @ param_names @ param_names' @ ["y", "P"]) [] elimtrm (fn _ => unfolddef_tac)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1501
in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1502
  (introthm, elimthm)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1503
end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1504
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  1505
fun create_constname_of_mode options thy prefix name T mode = 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1506
  let
33626
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33623
diff changeset
  1507
    val system_proposal = prefix ^ (Long_Name.base_name name)
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33623
diff changeset
  1508
      ^ "_" ^ ascii_string_of_mode' (translate_mode T mode)
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  1509
    val name = the_default system_proposal (proposed_names options name (translate_mode T mode))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1510
  in
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  1511
    Sign.full_bname thy name
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1512
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1513
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1514
fun split_tupleT is T =
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1515
  let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1516
    fun split_tuple' _ _ [] = ([], [])
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1517
      | split_tuple' is i (T::Ts) =
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1518
      (if i mem is then apfst else apsnd) (cons T)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1519
        (split_tuple' is (i+1) Ts)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1520
  in
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1521
    split_tuple' is 1 (HOLogic.strip_tupleT T)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1522
  end
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1523
  
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1524
fun mk_arg xin xout pis T =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1525
  let
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1526
    val n = length (HOLogic.strip_tupleT T)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1527
    val ni = length pis
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1528
    fun mk_proj i j t =
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1529
      (if i = j then I else HOLogic.mk_fst)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1530
        (funpow (i - 1) HOLogic.mk_snd t)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1531
    fun mk_arg' i (si, so) = if i mem pis then
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1532
        (mk_proj si ni xin, (si+1, so))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1533
      else
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1534
        (mk_proj so (n - ni) xout, (si, so+1))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1535
    val (args, _) = fold_map mk_arg' (1 upto n) (1, 1)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1536
  in
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1537
    HOLogic.mk_tuple args
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1538
  end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1539
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  1540
fun create_definitions options preds (name, modes) thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1541
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1542
    val compfuns = PredicateCompFuns.compfuns
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1543
    val T = AList.lookup (op =) preds name |> the
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1544
    fun create_definition (mode as (iss, is)) thy = let
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  1545
      val mode_cname = create_constname_of_mode options thy "" name T mode
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1546
      val mode_cbasename = Long_Name.base_name mode_cname
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1547
      val Ts = binder_types T
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1548
      val (Ts1, Ts2) = chop (length iss) Ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1549
      val (Us1, Us2) =  split_smodeT is Ts2
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1550
      val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1551
      val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (HOLogic.mk_tupleT Us2))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1552
      val names = Name.variant_list []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1553
        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1554
      (* old *)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1555
      (*
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1556
      val xs = map Free (names ~~ (Ts1' @ Ts2))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1557
      val (xparams, xargs) = chop (length iss) xs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1558
      val (xins, xouts) = split_smode is xargs
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1559
      *)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1560
      (* new *)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1561
      val param_names = Name.variant_list []
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1562
        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1')))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1563
      val xparams = map Free (param_names ~~ Ts1')
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1564
      fun mk_vars (i, T) names =
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1565
        let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1566
          val vname = Name.variant names ("x" ^ string_of_int (length Ts1' + i))
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1567
        in
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1568
          case AList.lookup (op =) is i of
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1569
             NONE => ((([], [Free (vname, T)]), Free (vname, T)), vname :: names)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1570
           | SOME NONE => ((([Free (vname, T)], []), Free (vname, T)), vname :: names)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1571
           | SOME (SOME pis) =>
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1572
             let
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1573
               val (Tins, Touts) = split_tupleT pis T
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1574
               val name_in = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "in")
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1575
               val name_out = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "out")
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1576
               val xin = Free (name_in, HOLogic.mk_tupleT Tins)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1577
               val xout = Free (name_out, HOLogic.mk_tupleT Touts)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1578
               val xarg = mk_arg xin xout pis T
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1579
             in
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1580
               (((if null Tins then [] else [xin],
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1581
               if null Touts then [] else [xout]), xarg), name_in :: name_out :: names) end
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1582
             end
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1583
      val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1584
      val (xinout, xargs) = split_list xinoutargs
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1585
      val (xins, xouts) = pairself flat (split_list xinout)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1586
      val (xparams', names') = fold_map (mk_Eval_of []) ((xparams ~~ Ts1) ~~ iss) names
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1587
      fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1588
        | mk_split_lambda [x] t = lambda x t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1589
        | mk_split_lambda xs t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1590
        let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1591
          fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1592
            | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1593
        in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1594
          mk_split_lambda' xs t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1595
        end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1596
      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1597
        (list_comb (Const (name, T), xparams' @ xargs)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1598
      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1599
      val def = Logic.mk_equals (lhs, predterm)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1600
      val ([definition], thy') = thy |>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1601
        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1602
        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1603
      val (intro, elim) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1604
        create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1605
      in thy'
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1606
        |> add_predfun name mode (mode_cname, definition, intro, elim)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1607
        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1608
        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1609
        |> Theory.checkpoint
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1610
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1611
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1612
    fold create_definition modes thy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1613
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1614
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  1615
fun define_functions comp_modifiers compfuns options preds (name, modes) thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1616
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1617
    val T = AList.lookup (op =) preds name |> the
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1618
    fun create_definition mode thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1619
      let
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1620
        val function_name_prefix = Comp_Mod.function_name_prefix comp_modifiers
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  1621
        val mode_cname = create_constname_of_mode options thy function_name_prefix name T mode
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1622
        val funT = Comp_Mod.funT_of comp_modifiers compfuns mode T
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1623
      in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1624
        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  1625
        |> Comp_Mod.set_function_name comp_modifiers name mode mode_cname
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1626
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1627
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1628
    fold create_definition modes thy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1629
  end;
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  1630
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1631
(* Proving equivalence of term *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1632
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1633
fun is_Type (Type _) = true
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1634
  | is_Type _ = false
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1635
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1636
(* returns true if t is an application of an datatype constructor *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1637
(* which then consequently would be splitted *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1638
(* else false *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1639
fun is_constructor thy t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1640
  if (is_Type (fastype_of t)) then
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1641
    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1642
      NONE => false
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1643
    | SOME info => (let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1644
      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1645
      val (c, _) = strip_comb t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1646
      in (case c of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1647
        Const (name, _) => name mem_string constr_consts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1648
        | _ => false) end))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1649
  else false
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1650
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1651
(* MAJOR FIXME:  prove_params should be simple
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1652
 - different form of introrule for parameters ? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1653
fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1654
  | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1655
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1656
    val  (f, args) = strip_comb (Envir.eta_contract t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1657
    val (params, _) = chop (length ms) args
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1658
    val f_tac = case f of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1659
      Const (name, T) => simp_tac (HOL_basic_ss addsimps 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1660
         ([@{thm eval_pred}, (predfun_definition_of thy name mode),
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1661
         @{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1662
         @{thm "snd_conv"}, @{thm pair_collapse}, @{thm "Product_Type.split_conv"}])) 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1663
    | Free _ => TRY (rtac @{thm refl} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1664
    | Abs _ => error "prove_param: No valid parameter term"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1665
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1666
    REPEAT_DETERM (etac @{thm thin_rl} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1667
    THEN REPEAT_DETERM (rtac @{thm ext} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1668
    THEN print_tac "prove_param"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1669
    THEN f_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1670
    THEN print_tac "after simplification in prove_args"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1671
    THEN (EVERY (map (prove_param thy) (ms ~~ params)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1672
    THEN (REPEAT_DETERM (atac 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1673
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1674
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1675
fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1676
  case strip_comb t of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1677
    (Const (name, T), args) =>  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1678
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1679
        val introrule = predfun_intro_of thy name mode
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1680
        val (args1, args2) = chop (length ms) args
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1681
      in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1682
        rtac @{thm bindI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1683
        THEN print_tac "before intro rule:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1684
        (* for the right assumption in first position *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1685
        THEN rotate_tac premposition 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1686
        THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1687
        THEN rtac introrule 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1688
        THEN print_tac "after intro rule"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1689
        (* work with parameter arguments *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1690
        THEN (atac 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1691
        THEN (print_tac "parameter goal")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1692
        THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1693
        THEN (REPEAT_DETERM (atac 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1694
      end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1695
  | _ => rtac @{thm bindI} 1
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1696
    THEN asm_full_simp_tac
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1697
      (HOL_basic_ss' addsimps [@{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1698
         @{thm "snd_conv"}, @{thm pair_collapse}]) 1
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1699
    THEN (atac 1)
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1700
    THEN print_tac "after prove parameter call"
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1701
    
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1702
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1703
fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1704
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1705
fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1706
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1707
fun prove_match thy (out_ts : term list) = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1708
  fun get_case_rewrite t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1709
    if (is_constructor thy t) then let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1710
      val case_rewrites = (#case_rewrites (Datatype.the_info thy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1711
        ((fst o dest_Type o fastype_of) t)))
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  1712
      in case_rewrites @ maps get_case_rewrite (snd (strip_comb t)) end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1713
    else []
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  1714
  val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: maps get_case_rewrite out_ts
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1715
(* replace TRY by determining if it necessary - are there equations when calling compile match? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1716
in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1717
   (* make this simpset better! *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1718
  asm_full_simp_tac (HOL_basic_ss' addsimps simprules) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1719
  THEN print_tac "after prove_match:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1720
  THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1721
         THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1722
         THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1723
  THEN print_tac "after if simplification"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1724
end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1725
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1726
(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1727
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1728
fun prove_sidecond thy modes t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1729
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1730
    fun preds_of t nameTs = case strip_comb t of 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1731
      (f as Const (name, T), args) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1732
        if AList.defined (op =) modes name then (name, T) :: nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1733
          else fold preds_of args nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1734
      | _ => nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1735
    val preds = preds_of t []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1736
    val defs = map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1737
      (fn (pred, T) => predfun_definition_of thy pred
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1738
        ([], map (rpair NONE) (1 upto (length (binder_types T)))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1739
        preds
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1740
  in 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1741
    (* remove not_False_eq_True when simpset in prove_match is better *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1742
    simp_tac (HOL_basic_ss addsimps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1743
      (@{thms "HOL.simp_thms"} @ (@{thm not_False_eq_True} :: @{thm eval_pred} :: defs))) 1 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1744
    (* need better control here! *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1745
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1746
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1747
fun prove_clause options thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1748
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1749
    val (in_ts, clause_out_ts) = split_smode is ts;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1750
    fun prove_prems out_ts [] =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1751
      (prove_match thy out_ts)
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1752
      THEN print_tac "before simplifying assumptions"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1753
      THEN asm_full_simp_tac HOL_basic_ss' 1
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1754
      THEN print_tac "before single intro rule"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1755
      THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1756
    | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1757
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1758
        val premposition = (find_index (equal p) clauses) + nargs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1759
        val rest_tac = (case p of Prem (us, t) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1760
            let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1761
              val (_, out_ts''') = split_smode is us
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1762
              val rec_tac = prove_prems out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1763
            in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1764
              print_tac "before clause:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1765
              THEN asm_simp_tac HOL_basic_ss 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1766
              THEN print_tac "before prove_expr:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1767
              THEN prove_expr thy (mode, t, us) premposition
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1768
              THEN print_tac "after prove_expr:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1769
              THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1770
            end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1771
          | Negprem (us, t) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1772
            let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1773
              val (_, out_ts''') = split_smode is us
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1774
              val rec_tac = prove_prems out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1775
              val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1776
              val (_, params) = strip_comb t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1777
            in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1778
              rtac @{thm bindI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1779
              THEN (if (is_some name) then
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1780
                  simp_tac (HOL_basic_ss addsimps
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1781
                    [predfun_definition_of thy (the name) (iss, is)]) 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1782
                  THEN rtac @{thm not_predI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1783
                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1784
                  THEN (REPEAT_DETERM (atac 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1785
                  (* FIXME: work with parameter arguments *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1786
                  THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1787
                else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1788
                  rtac @{thm not_predI'} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1789
                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1790
              THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1791
            end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1792
          | Sidecond t =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1793
           rtac @{thm bindI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1794
           THEN rtac @{thm if_predI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1795
           THEN print_tac "before sidecond:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1796
           THEN prove_sidecond thy modes t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1797
           THEN print_tac "after sidecond:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1798
           THEN prove_prems [] ps)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1799
      in (prove_match thy out_ts)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1800
          THEN rest_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1801
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1802
    val prems_tac = prove_prems in_ts moded_ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1803
  in
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1804
    print_tac' options "Proving clause..."
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1805
    THEN rtac @{thm bindI} 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1806
    THEN rtac @{thm singleI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1807
    THEN prems_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1808
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1809
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1810
fun select_sup 1 1 = []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1811
  | select_sup _ 1 = [rtac @{thm supI1}]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1812
  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1813
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
  1814
fun prove_one_direction options thy clauses preds modes pred mode moded_clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1815
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1816
    val T = the (AList.lookup (op =) preds pred)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1817
    val nargs = length (binder_types T) - nparams_of thy pred
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1818
    val pred_case_rule = the_elim_of thy pred
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1819
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1820
    REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  1821
    THEN print_tac' options "before applying elim rule"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1822
    THEN etac (predfun_elim_of thy pred mode) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1823
    THEN etac pred_case_rule 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1824
    THEN (EVERY (map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1825
           (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1826
             (1 upto (length moded_clauses))))
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1827
    THEN (EVERY (map2 (prove_clause options thy nargs modes mode) clauses moded_clauses))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1828
    THEN print_tac "proved one direction"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1829
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1830
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1831
(** Proof in the other direction **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1832
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1833
fun prove_match2 thy out_ts = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1834
  fun split_term_tac (Free _) = all_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1835
    | split_term_tac t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1836
      if (is_constructor thy t) then let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1837
        val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1838
        val num_of_constrs = length (#case_rewrites info)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1839
        (* special treatment of pairs -- because of fishing *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1840
        val split_rules = case (fst o dest_Type o fastype_of) t of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1841
          "*" => [@{thm prod.split_asm}] 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1842
          | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1843
        val (_, ts) = strip_comb t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1844
      in
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1845
        (print_tac ("Term " ^ (Syntax.string_of_term_global thy t) ^ 
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1846
          "splitting with rules \n" ^
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1847
        commas (map (Display.string_of_thm_global thy) split_rules)))
33115
f765c3234059 changed proof method to handle widen predicate in JinjaThreads
bulwahn
parents: 33114
diff changeset
  1848
        THEN TRY ((Splitter.split_asm_tac split_rules 1)
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1849
        THEN (print_tac "after splitting with split_asm rules")
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1850
        (* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1851
          THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1852
          THEN (REPEAT_DETERM_N (num_of_constrs - 1)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1853
            (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
  1854
        THEN (assert_tac (Max_number_of_subgoals 2))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1855
        THEN (EVERY (map split_term_tac ts))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1856
      end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1857
    else all_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1858
  in
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1859
    split_term_tac (HOLogic.mk_tuple out_ts)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1860
    THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1861
    THEN (etac @{thm botE} 2))))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1862
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1863
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1864
(* VERY LARGE SIMILIRATIY to function prove_param 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1865
-- join both functions
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1866
*)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1867
(* TODO: remove function *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1868
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1869
fun prove_param2 thy (NONE, t) = all_tac 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1870
  | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1871
    val  (f, args) = strip_comb (Envir.eta_contract t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1872
    val (params, _) = chop (length ms) args
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1873
    val f_tac = case f of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1874
        Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1875
           (@{thm eval_pred}::(predfun_definition_of thy name mode)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1876
           :: @{thm "Product_Type.split_conv"}::[])) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1877
      | Free _ => all_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1878
      | _ => error "prove_param2: illegal parameter term"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1879
  in  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1880
    print_tac "before simplification in prove_args:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1881
    THEN f_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1882
    THEN print_tac "after simplification in prove_args"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1883
    THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1884
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1885
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1886
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1887
fun prove_expr2 thy (Mode (mode, is, ms), t) = 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1888
  (case strip_comb t of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1889
    (Const (name, T), args) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1890
      etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1891
      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1892
      THEN print_tac "prove_expr2-before"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1893
      THEN (debug_tac (Syntax.string_of_term_global thy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1894
        (prop_of (predfun_elim_of thy name mode))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1895
      THEN (etac (predfun_elim_of thy name mode) 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1896
      THEN print_tac "prove_expr2"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1897
      THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1898
      THEN print_tac "finished prove_expr2"      
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1899
    | _ => etac @{thm bindE} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1900
    
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1901
(* FIXME: what is this for? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1902
(* replace defined by has_mode thy pred *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1903
(* TODO: rewrite function *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1904
fun prove_sidecond2 thy modes t = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1905
  fun preds_of t nameTs = case strip_comb t of 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1906
    (f as Const (name, T), args) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1907
      if AList.defined (op =) modes name then (name, T) :: nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1908
        else fold preds_of args nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1909
    | _ => nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1910
  val preds = preds_of t []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1911
  val defs = map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1912
    (fn (pred, T) => predfun_definition_of thy pred 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1913
      ([], map (rpair NONE) (1 upto (length (binder_types T)))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1914
      preds
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1915
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1916
   (* only simplify the one assumption *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1917
   full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1918
   (* need better control here! *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1919
   THEN print_tac "after sidecond2 simplification"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1920
   end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1921
  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1922
fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1923
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1924
    val pred_intro_rule = nth (intros_of thy pred) (i - 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1925
    val (in_ts, clause_out_ts) = split_smode is ts;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1926
    fun prove_prems2 out_ts [] =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1927
      print_tac "before prove_match2 - last call:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1928
      THEN prove_match2 thy out_ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1929
      THEN print_tac "after prove_match2 - last call:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1930
      THEN (etac @{thm singleE} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1931
      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1932
      THEN (asm_full_simp_tac HOL_basic_ss' 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1933
      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1934
      THEN (asm_full_simp_tac HOL_basic_ss' 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1935
      THEN SOLVED (print_tac "state before applying intro rule:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1936
      THEN (rtac pred_intro_rule 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1937
      (* How to handle equality correctly? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1938
      THEN (print_tac "state before assumption matching")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1939
      THEN (REPEAT (atac 1 ORELSE 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1940
         (CHANGED (asm_full_simp_tac (HOL_basic_ss' addsimps
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1941
           [@{thm split_eta}, @{thm "split_beta"}, @{thm "fst_conv"},
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1942
             @{thm "snd_conv"}, @{thm pair_collapse}]) 1)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1943
          THEN print_tac "state after simp_tac:"))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1944
    | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1945
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1946
        val rest_tac = (case p of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1947
          Prem (us, t) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1948
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1949
            val (_, out_ts''') = split_smode is us
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1950
            val rec_tac = prove_prems2 out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1951
          in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1952
            (prove_expr2 thy (mode, t)) THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1953
          end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1954
        | Negprem (us, t) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1955
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1956
            val (_, out_ts''') = split_smode is us
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1957
            val rec_tac = prove_prems2 out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1958
            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1959
            val (_, params) = strip_comb t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1960
          in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1961
            print_tac "before neg prem 2"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1962
            THEN etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1963
            THEN (if is_some name then
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1964
                full_simp_tac (HOL_basic_ss addsimps
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1965
                  [predfun_definition_of thy (the name) (iss, is)]) 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1966
                THEN etac @{thm not_predE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1967
                THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1968
                THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1969
              else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1970
                etac @{thm not_predE'} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1971
            THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1972
          end 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1973
        | Sidecond t =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1974
          etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1975
          THEN etac @{thm if_predE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1976
          THEN prove_sidecond2 thy modes t 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1977
          THEN prove_prems2 [] ps)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1978
      in print_tac "before prove_match2:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1979
         THEN prove_match2 thy out_ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1980
         THEN print_tac "after prove_match2:"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1981
         THEN rest_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1982
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1983
    val prems_tac = prove_prems2 in_ts ps 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1984
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1985
    print_tac "starting prove_clause2"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1986
    THEN etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1987
    THEN (etac @{thm singleE'} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1988
    THEN (TRY (etac @{thm Pair_inject} 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1989
    THEN print_tac "after singleE':"
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1990
    THEN prems_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1991
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1992
 
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1993
fun prove_other_direction options thy modes pred mode moded_clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1994
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1995
    fun prove_clause clause i =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1996
      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1997
      THEN (prove_clause2 thy modes pred mode clause i)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1998
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1999
    (DETERM (TRY (rtac @{thm unit.induct} 1)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2000
     THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2001
     THEN (rtac (predfun_intro_of thy pred mode) 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2002
     THEN (REPEAT_DETERM (rtac @{thm refl} 2))
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2003
     THEN (if null moded_clauses then
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2004
         etac @{thm botE} 1
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2005
       else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2006
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2007
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2008
(** proof procedure **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2009
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
  2010
fun prove_pred options thy clauses preds modes pred mode (moded_clauses, compiled_term) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2011
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2012
    val ctxt = ProofContext.init thy
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2013
    val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2014
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2015
    Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2016
      (if not (skip_proof options) then
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2017
        (fn _ =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2018
        rtac @{thm pred_iffI} 1
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  2019
        THEN print_tac' options "after pred_iffI"
33128
1f990689349f further cleaning up
bulwahn
parents: 33127
diff changeset
  2020
        THEN prove_one_direction options thy clauses preds modes pred mode moded_clauses
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
  2021
        THEN print_tac' options "proved one direction"
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2022
        THEN prove_other_direction options thy modes pred mode moded_clauses
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
  2023
        THEN print_tac' options "proved other direction")
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
  2024
      else (fn _ => Skip_Proof.cheat_tac thy))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2025
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2026
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2027
(* composition of mode inference, definition, compilation and proof *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2028
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2029
(** auxillary combinators for table of preds and modes **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2030
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2031
fun map_preds_modes f preds_modes_table =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2032
  map (fn (pred, modes) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2033
    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2034
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2035
fun join_preds_modes table1 table2 =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2036
  map_preds_modes (fn pred => fn mode => fn value =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2037
    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2038
    
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2039
fun maps_modes preds_modes_table =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2040
  map (fn (pred, modes) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2041
    (pred, map (fn (mode, value) => value) modes)) preds_modes_table  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2042
    
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2043
fun compile_preds comp_modifiers compfuns thy all_vs param_vs preds moded_clauses =
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2044
  map_preds_modes (fn pred => compile_pred comp_modifiers compfuns thy all_vs param_vs pred
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2045
      (the (AList.lookup (op =) preds pred))) moded_clauses
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2046
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
  2047
fun prove options thy clauses preds modes moded_clauses compiled_terms =
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
  2048
  map_preds_modes (prove_pred options thy clauses preds modes)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2049
    (join_preds_modes moded_clauses compiled_terms)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2050
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
  2051
fun prove_by_skip options thy _ _ _ _ compiled_terms =
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
  2052
  map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2053
    compiled_terms
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2054
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2055
(* preparation of introduction rules into special datastructures *)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2056
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2057
fun dest_prem thy params t =
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2058
  (case strip_comb t of
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2059
    (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2060
  | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem thy params t of
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2061
      Prem (ts, t) => Negprem (ts, t)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2062
    | Negprem _ => error ("Double negation not allowed in premise: " ^
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2063
        Syntax.string_of_term_global thy (c $ t)) 
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2064
    | Sidecond t => Sidecond (c $ t))
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2065
  | (c as Const (s, _), ts) =>
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2066
    if is_registered thy s then
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2067
      let val (ts1, ts2) = chop (nparams_of thy s) ts
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2068
      in Prem (ts2, list_comb (c, ts1)) end
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2069
    else Sidecond t
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2070
  | _ => Sidecond t)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2071
    
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2072
fun prepare_intrs options thy prednames intros =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2073
  let
33126
bb8806eb5da7 importing of polymorphic introduction rules with different schematic variable names
bulwahn
parents: 33124
diff changeset
  2074
    val intrs = map prop_of intros
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2075
    val nparams = nparams_of thy (hd prednames)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2076
    val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2077
    val (preds, intrs) = unify_consts thy preds intrs
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2078
    val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs]
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2079
      (ProofContext.init thy)
33126
bb8806eb5da7 importing of polymorphic introduction rules with different schematic variable names
bulwahn
parents: 33124
diff changeset
  2080
    val preds = map dest_Const preds
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2081
    val extra_modes = all_modes_of thy
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2082
      |> filter_out (fn (name, _) => member (op =) prednames name)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2083
    val params = case intrs of
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2084
        [] =>
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2085
          let
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2086
            val (paramTs, _) = chop nparams (binder_types (snd (hd preds)))
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2087
            val param_names = Name.variant_list [] (map (fn i => "p" ^ string_of_int i)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2088
              (1 upto length paramTs))
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2089
          in map Free (param_names ~~ paramTs) end
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2090
      | intr :: _ => fst (chop nparams
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2091
        (snd (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)))))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2092
    val param_vs = maps term_vs params
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2093
    val all_vs = terms_vs intrs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2094
    fun add_clause intr (clauses, arities) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2095
    let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2096
      val _ $ t = Logic.strip_imp_concl intr;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2097
      val (Const (name, T), ts) = strip_comb t;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2098
      val (ts1, ts2) = chop nparams ts;
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2099
      val prems = map (dest_prem thy params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2100
      val (Ts, Us) = chop nparams (binder_types T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2101
    in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2102
      (AList.update op = (name, these (AList.lookup op = clauses name) @
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2103
        [(ts2, prems)]) clauses,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2104
       AList.update op = (name, (map (fn U => (case strip_type U of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2105
                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2106
               | _ => NONE)) Ts,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2107
             length Us)) arities)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2108
    end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2109
    val (clauses, arities) = fold add_clause intrs ([], []);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2110
    fun modes_of_arities arities =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2111
      (map (fn (s, (ks, k)) => (s, cprod (cprods (map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2112
            (fn NONE => [NONE]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2113
              | SOME k' => map SOME (map (map (rpair NONE)) (subsets 1 k'))) ks),
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2114
       map (map (rpair NONE)) (subsets 1 k)))) arities)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2115
    fun modes_of_typ T =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2116
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2117
        val (Ts, Us) = chop nparams (binder_types T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2118
        fun all_smodes_of_typs Ts = cprods_subset (
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2119
          map_index (fn (i, U) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2120
            case HOLogic.strip_tupleT U of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2121
              [] => [(i + 1, NONE)]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2122
            | [U] => [(i + 1, NONE)]
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2123
            | Us =>  (i + 1, NONE) ::
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2124
              (map (pair (i + 1) o SOME)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2125
                (subtract (op =) [[], 1 upto (length Us)] (subsets 1 (length Us)))))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2126
          Ts)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2127
      in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2128
        cprod (cprods (map (fn T => case strip_type T of
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2129
          (Rs as _ :: _, Type ("bool", [])) =>
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2130
            map SOME (all_smodes_of_typs Rs) | _ => [NONE]) Ts), all_smodes_of_typs Us)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2131
      end
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2132
    val all_modes = map (fn (s, T) =>
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2133
      case proposed_modes options s of
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2134
        NONE => (s, modes_of_typ T)
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2135
      | SOME modes' => (s, map (translate_mode' nparams) modes'))
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2136
        preds
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2137
  in (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2138
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2139
(* sanity check of introduction rules *)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2140
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2141
fun check_format_of_intro_rule thy intro =
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2142
  let
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2143
    val concl = Logic.strip_imp_concl (prop_of intro)
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2144
    val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2145
    val params = List.take (args, nparams_of thy (fst (dest_Const p)))
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2146
    fun check_arg arg = case HOLogic.strip_tupleT (fastype_of arg) of
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2147
      (Ts as _ :: _ :: _) =>
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2148
        if (length (HOLogic.strip_tuple arg) = length Ts) then true
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
  2149
        else
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
  2150
        error ("Format of introduction rule is invalid: tuples must be expanded:"
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2151
        ^ (Syntax.string_of_term_global thy arg) ^ " in " ^
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2152
        (Display.string_of_thm_global thy intro)) 
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2153
      | _ => true
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2154
    val prems = Logic.strip_imp_prems (prop_of intro)
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2155
    fun check_prem (Prem (args, _)) = forall check_arg args
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2156
      | check_prem (Negprem (args, _)) = forall check_arg args
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2157
      | check_prem _ = true
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2158
  in
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2159
    forall check_arg args andalso
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2160
    forall (check_prem o dest_prem thy params o HOLogic.dest_Trueprop) prems
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2161
  end
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2162
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2163
(*
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2164
fun check_intros_elim_match thy prednames =
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2165
  let
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2166
    fun check predname =
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2167
      let
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2168
        val intros = intros_of thy predname
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2169
        val elim = the_elim_of thy predname
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2170
        val nparams = nparams_of thy predname
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2171
        val elim' =
33150
75eddea4abef further changes due to the previous merge in the predicate compiler
bulwahn
parents: 33149
diff changeset
  2172
          (Drule.standard o (Skip_Proof.make_thm thy))
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2173
          (mk_casesrule (ProofContext.init thy) nparams intros)
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2174
      in
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2175
        if not (Thm.equiv_thm (elim, elim')) then
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2176
          error "Introduction and elimination rules do not match!"
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2177
        else true
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2178
      end
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2179
  in forall check prednames end
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2180
*)
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
  2181
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2182
(* create code equation *)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2183
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2184
fun add_code_equations thy nparams preds result_thmss =
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2185
  let
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2186
    fun add_code_equation ((predname, T), (pred, result_thms)) =
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2187
      let
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2188
        val full_mode = (replicate nparams NONE,
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2189
          map (rpair NONE) (1 upto (length (binder_types T) - nparams)))
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2190
      in
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2191
        if member (op =) (modes_of thy predname) full_mode then
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2192
          let
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2193
            val Ts = binder_types T
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2194
            val arg_names = Name.variant_list []
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2195
              (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2196
            val args = map Free (arg_names ~~ Ts)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2197
            val predfun = Const (predfun_name_of thy predname full_mode,
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2198
              Ts ---> PredicateCompFuns.mk_predT @{typ unit})
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2199
            val rhs = PredicateCompFuns.mk_Eval (list_comb (predfun, args), @{term "Unity"})
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2200
            val eq_term = HOLogic.mk_Trueprop
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2201
              (HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs))
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2202
            val def = predfun_definition_of thy predname full_mode
33441
99a5f22a967d eliminated funny record patterns and made SML/NJ happy;
wenzelm
parents: 33377
diff changeset
  2203
            val tac = fn _ => Simplifier.simp_tac
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2204
              (HOL_basic_ss addsimps [def, @{thm eval_pred}]) 1
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2205
            val eq = Goal.prove (ProofContext.init thy) arg_names [] eq_term tac
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2206
          in
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2207
            (pred, result_thms @ [eq])
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2208
          end
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2209
        else
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2210
          (pred, result_thms)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2211
      end
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2212
  in
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2213
    map add_code_equation (preds ~~ result_thmss)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2214
  end
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2215
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2216
(** main function of predicate compiler **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2217
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2218
datatype steps = Steps of
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2219
  {
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2220
  compile_preds : theory -> string list -> string list -> (string * typ) list
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2221
    -> (moded_clause list) pred_mode_table -> term pred_mode_table,
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2222
  define_functions : options -> (string * typ) list -> string * mode list -> theory -> theory,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2223
  infer_modes : options -> theory -> (string * mode list) list -> (string * mode list) list
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2224
    -> string list -> (string * (term list * indprem list) list) list
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2225
    -> moded_clause list pred_mode_table,
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2226
  prove : options -> theory -> (string * (term list * indprem list) list) list
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2227
    -> (string * typ) list -> (string * mode list) list
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2228
    -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2229
  add_code_equations : theory -> int -> (string * typ) list
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2230
    -> (string * thm list) list -> (string * thm list) list,
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2231
  defined : theory -> string -> bool,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2232
  qname : bstring
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2233
  }
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2234
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2235
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2236
fun add_equations_of steps options prednames thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2237
  let
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2238
    fun dest_steps (Steps s) = s
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2239
    val _ = print_step options
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2240
      ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2241
      (*val _ = check_intros_elim_match thy prednames*)
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
  2242
      (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2243
    val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) =
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2244
      prepare_intrs options thy prednames (maps (intros_of thy) prednames)
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2245
    val _ = print_step options "Infering modes..."
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2246
    val moded_clauses =
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2247
      #infer_modes (dest_steps steps) options thy extra_modes all_modes param_vs clauses
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2248
    val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
  2249
    val _ = check_expected_modes preds options modes
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
  2250
    val _ = print_modes options thy modes
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
  2251
      (*val _ = print_moded_clauses thy moded_clauses*)
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2252
    val _ = print_step options "Defining executable functions..."
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2253
    val thy' = fold (#define_functions (dest_steps steps) options preds) modes thy
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2254
      |> Theory.checkpoint
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2255
    val _ = print_step options "Compiling equations..."
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2256
    val compiled_terms =
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2257
      #compile_preds (dest_steps steps) thy' all_vs param_vs preds moded_clauses
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
  2258
    val _ = print_compiled_terms options thy' compiled_terms
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2259
    val _ = print_step options "Proving equations..."
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2260
    val result_thms = #prove (dest_steps steps) options thy' clauses preds (extra_modes @ modes)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2261
      moded_clauses compiled_terms
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2262
    val result_thms' = #add_code_equations (dest_steps steps) thy' nparams preds
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2263
      (maps_modes result_thms)
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2264
    val qname = #qname (dest_steps steps)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2265
    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2266
      (fn thm => Context.mapping (Code.add_eqn thm) I))))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2267
    val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2268
      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2269
        [attrib thy ])] thy))
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2270
      result_thms' thy' |> Theory.checkpoint
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2271
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2272
    thy''
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2273
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2274
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2275
fun extend' value_of edges_of key (G, visited) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2276
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2277
    val (G', v) = case try (Graph.get_node G) key of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2278
        SOME v => (G, v)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2279
      | NONE => (Graph.new_node (key, value_of key) G, value_of key)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2280
    val (G'', visited') = fold (extend' value_of edges_of)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2281
      (subtract (op =) visited (edges_of (key, v)))
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
  2282
      (G', key :: visited)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2283
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2284
    (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2285
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2286
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2287
fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2288
  
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2289
fun gen_add_equations steps options names thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2290
  let
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2291
    fun dest_steps (Steps s) = s
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2292
    val thy' = thy
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2293
      |> PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2294
      |> Theory.checkpoint;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2295
    fun strong_conn_of gr keys =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2296
      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2297
    val scc = strong_conn_of (PredData.get thy') names
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2298
    val thy'' = fold_rev
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2299
      (fn preds => fn thy =>
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2300
        if not (forall (#defined (dest_steps steps) thy) preds) then
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2301
          add_equations_of steps options preds thy
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2302
        else thy)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2303
      scc thy' |> Theory.checkpoint
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2304
  in thy'' end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2305
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2306
(* different instantiantions of the predicate compiler *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2307
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2308
val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2309
  {function_name_of = predfun_name_of : (theory -> string -> mode -> string),
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2310
  set_function_name = (fn _ => fn _ => fn _ => I),
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2311
  function_name_prefix = "",
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2312
  funT_of = funT_of : (compilation_funs -> mode -> typ -> typ),
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2313
  additional_arguments = K [],
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2314
  wrap_compilation = K (K (K (K (K I))))
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2315
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2316
  transform_additional_arguments = K I : (indprem -> term list -> term list)
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2317
  }
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2318
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2319
val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2320
  {function_name_of = depth_limited_function_name_of,
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2321
  set_function_name = set_depth_limited_function_name,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2322
  funT_of = depth_limited_funT_of : (compilation_funs -> mode -> typ -> typ),
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2323
  function_name_prefix = "depth_limited_",
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2324
  additional_arguments = fn names =>
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2325
    let
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2326
      val [depth_name, polarity_name] = Name.variant_list names ["depth", "polarity"]
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2327
    in [Free (polarity_name, @{typ "bool"}), Free (depth_name, @{typ "code_numeral"})] end,
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2328
  wrap_compilation =
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2329
    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2330
    let
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2331
      val [polarity, depth] = additional_arguments
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2332
      val (_, Ts2) = chop (length (fst mode)) (binder_types T)
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2333
      val (_, Us2) = split_smodeT (snd mode) Ts2
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  2334
      val T' = mk_predT compfuns (HOLogic.mk_tupleT Us2)
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2335
      val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2336
      val full_mode = null Us2
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2337
    in
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2338
      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2339
        $ (if_const $ polarity $ mk_bot compfuns (dest_predT compfuns T')
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2340
          $ (if full_mode then mk_single compfuns HOLogic.unit else
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2341
            Const (@{const_name undefined}, T')))
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2342
        $ compilation
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2343
    end,
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2344
  transform_additional_arguments =
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2345
    fn prem => fn additional_arguments =>
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2346
    let
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2347
      val [polarity, depth] = additional_arguments
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
  2348
      val polarity' = (case prem of Prem _ => I | Negprem _ => HOLogic.mk_not | _ => I) polarity
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2349
      val depth' =
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2350
        Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2351
          $ depth $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"})
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2352
    in [polarity', depth'] end
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2353
  }
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2354
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33333
diff changeset
  2355
val random_comp_modifiers = Comp_Mod.Comp_Modifiers
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2356
  {function_name_of = random_function_name_of,
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2357
  set_function_name = set_random_function_name,
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2358
  function_name_prefix = "random_",
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
  2359
  funT_of = K random_function_funT_of : (compilation_funs -> mode -> typ -> typ),
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2360
  additional_arguments = fn names => [Free (Name.variant names "size", @{typ code_numeral})],
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2361
  wrap_compilation = K (K (K (K (K I))))
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2362
    : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2363
  transform_additional_arguments = K I : (indprem -> term list -> term list)
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2364
  }
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2365
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2366
val annotated_comp_modifiers = Comp_Mod.Comp_Modifiers
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2367
  {function_name_of = annotated_function_name_of,
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2368
  set_function_name = set_annotated_function_name,
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2369
  function_name_prefix = "annotated_",
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2370
  funT_of = funT_of : (compilation_funs -> mode -> typ -> typ),
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2371
  additional_arguments = K [],
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2372
  wrap_compilation =
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2373
    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
33619
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
  2374
      mk_tracing ("calling predicate " ^ s ^
d93a3cb55068 first steps towards a new mode datastructure; new syntax for mode annotations and new output of modes
bulwahn
parents: 33526
diff changeset
  2375
        " with mode " ^ string_of_mode' (translate_mode T mode)) compilation,
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2376
  transform_additional_arguments = K I : (indprem -> term list -> term list)
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2377
  }
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2378
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2379
val add_equations = gen_add_equations
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2380
  (Steps {infer_modes = infer_modes,
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2381
  define_functions = create_definitions,
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2382
  compile_preds = compile_preds predicate_comp_modifiers PredicateCompFuns.compfuns,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2383
  prove = prove,
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2384
  add_code_equations = add_code_equations,
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2385
  defined = defined_functions,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2386
  qname = "equation"})
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2387
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
  2388
val add_depth_limited_equations = gen_add_equations
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2389
  (Steps {infer_modes = infer_modes,
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2390
  define_functions = define_functions depth_limited_comp_modifiers PredicateCompFuns.compfuns,
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2391
  compile_preds = compile_preds depth_limited_comp_modifiers PredicateCompFuns.compfuns,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2392
  prove = prove_by_skip,
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2393
  add_code_equations = K (K (K I)),
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2394
  defined = defined_depth_limited_functions,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2395
  qname = "depth_limited_equation"})
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2396
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2397
val add_annotated_equations = gen_add_equations
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2398
  (Steps {infer_modes = infer_modes,
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2399
  define_functions = define_functions annotated_comp_modifiers PredicateCompFuns.compfuns,
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2400
  compile_preds = compile_preds annotated_comp_modifiers PredicateCompFuns.compfuns,
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2401
  prove = prove_by_skip,
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2402
  add_code_equations = K (K (K I)),
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2403
  defined = defined_annotated_functions,
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2404
  qname = "annotated_equation"})
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2405
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2406
val add_quickcheck_equations = gen_add_equations
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2407
  (Steps {infer_modes = infer_modes_with_generator,
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2408
  define_functions = define_functions random_comp_modifiers RandomPredCompFuns.compfuns,
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33333
diff changeset
  2409
  compile_preds = compile_preds random_comp_modifiers RandomPredCompFuns.compfuns,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2410
  prove = prove_by_skip,
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2411
  add_code_equations = K (K (K I)),
33483
a15a2f7f7560 improved handling of already defined functions in the predicate compiler; could cause trouble before when no modes for a predicate were infered
bulwahn
parents: 33482
diff changeset
  2412
  defined = defined_random_functions,
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33333
diff changeset
  2413
  qname = "random_equation"})
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2414
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2415
(** user interface **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2416
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2417
(* code_pred_intro attribute *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2418
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2419
fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2420
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2421
val code_pred_intros_attrib = attrib add_intro;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2422
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2423
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2424
(*FIXME
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2425
- Naming of auxiliary rules necessary?
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2426
- add default code equations P x y z = P_i_i_i x y z
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2427
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2428
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2429
val setup = PredData.put (Graph.empty) #>
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2430
  Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro))
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2431
    "adding alternative introduction rules for code generation of inductive predicates"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2432
  (*FIXME name discrepancy in attribs and ML code*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2433
  (*FIXME intros should be better named intro*)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2434
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
  2435
(* TODO: make Theory_Data to Generic_Data & remove duplication of local theory and theory *)
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2436
fun generic_code_pred prep_const options raw_const lthy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2437
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2438
    val thy = ProofContext.theory_of lthy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2439
    val const = prep_const thy raw_const
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2440
    val lthy' = LocalTheory.theory (PredData.map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2441
        (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2442
      |> LocalTheory.checkpoint
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2443
    val thy' = ProofContext.theory_of lthy'
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33268
diff changeset
  2444
    val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim thy')
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2445
    fun mk_cases const =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2446
      let
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2447
        val T = Sign.the_const_type thy const
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2448
        val pred = Const (const, T)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2449
        val nparams = nparams_of thy' const
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2450
        val intros = intros_of thy' const
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2451
      in mk_casesrule lthy' pred nparams intros end  
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2452
    val cases_rules = map mk_cases preds
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2453
    val cases =
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33333
diff changeset
  2454
      map (fn case_rule => Rule_Cases.Case {fixes = [],
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2455
        assumes = [("", Logic.strip_imp_prems case_rule)],
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2456
        binds = [], cases = []}) cases_rules
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2457
    val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2458
    val lthy'' = lthy'
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2459
      |> fold Variable.auto_fixes cases_rules 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2460
      |> ProofContext.add_cases true case_env
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2461
    fun after_qed thms goal_ctxt =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2462
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2463
        val global_thms = ProofContext.export goal_ctxt
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2464
          (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2465
      in
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2466
        goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #>
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33333
diff changeset
  2467
          (if is_random options then
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2468
            (add_equations options [const] #>
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2469
            add_quickcheck_equations options [const])
33134
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
  2470
           else if is_depth_limited options then
88c9c3460fe7 renamed functions from sizelim to more natural name depth_limited for compilation of depth-limited search in the predicate compiler
bulwahn
parents: 33133
diff changeset
  2471
             add_depth_limited_equations options [const]
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2472
           else if is_annotated options then
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2473
             add_annotated_equations options [const]
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2474
           else
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2475
             add_equations options [const]))
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
  2476
      end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2477
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2478
    Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2479
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2480
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2481
val code_pred = generic_code_pred (K I);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2482
val code_pred_cmd = generic_code_pred Code.read_const
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2483
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2484
(* transformation for code generation *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2485
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32674
diff changeset
  2486
val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option);
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2487
val random_eval_ref =
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2488
  Unsynchronized.ref (NONE : (unit -> int * int -> term Predicate.pred * (int * int)) option);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2489
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2490
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2491
(* TODO: make analyze_compr generic with respect to the compilation modifiers*)
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2492
fun analyze_compr thy compfuns param_user_modes (depth_limit, (random, annotated)) t_compr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2493
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2494
    val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2495
      | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2496
    val (body, Ts, fp) = HOLogic.strip_psplits split;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2497
    val (pred as Const (name, T), all_args) = strip_comb body;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2498
    val (params, args) = chop (nparams_of thy name) all_args;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2499
    val user_mode = map_filter I (map_index
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2500
      (fn (i, t) => case t of Bound j => if j < length Ts then NONE
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2501
        else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2502
    val user_mode' = map (rpair NONE) user_mode
33484
e96a0c52d74b renamed generator to random_function in the predicate compiler
bulwahn
parents: 33483
diff changeset
  2503
    val all_modes_of = if random then all_random_modes_of else all_modes_of
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2504
    fun fits_to is NONE = true
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2505
      | fits_to is (SOME pm) = (is = (snd (translate_mode' 0 pm)))
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2506
    fun valid ((SOME (Mode (_, is, ms))) :: ms') (pm :: pms) =
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2507
        fits_to is pm andalso valid (ms @ ms') pms
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2508
      | valid (NONE :: ms') pms = valid ms' pms
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2509
      | valid [] [] = true
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2510
      | valid [] _ = error "Too many mode annotations"
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2511
      | valid (SOME _ :: _) [] = error "Not enough mode annotations"
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2512
    val modes = filter (fn Mode (_, is, ms) => is = user_mode'
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2513
        andalso (the_default true (Option.map (valid ms) param_user_modes)))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2514
      (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2515
    val m = case modes
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2516
     of [] => error ("No mode possible for comprehension "
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2517
                ^ Syntax.string_of_term_global thy t_compr)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2518
      | [m] => m
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2519
      | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2520
                ^ Syntax.string_of_term_global thy t_compr); m);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2521
    val (inargs, outargs) = split_smode user_mode' args;
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2522
    val additional_arguments =
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2523
      case depth_limit of
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2524
        NONE => (if random then [@{term "5 :: code_numeral"}] else [])
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2525
      | SOME d => [@{term "True"}, HOLogic.mk_number @{typ "code_numeral"} d]
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2526
    val comp_modifiers =
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2527
      case depth_limit of
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2528
        NONE =>
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2529
          (if random then random_comp_modifiers else
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2530
           if annotated then annotated_comp_modifiers else predicate_comp_modifiers)
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2531
      | SOME _ => depth_limited_comp_modifiers
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2532
    val t_pred = compile_expr comp_modifiers compfuns thy
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2533
      (m, list_comb (pred, params)) inargs additional_arguments;
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2534
    val t_eval = if null outargs then t_pred else
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  2535
      let
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2536
        val outargs_bounds = map (fn Bound i => i) outargs;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2537
        val outargsTs = map (nth Ts) outargs_bounds;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2538
        val T_pred = HOLogic.mk_tupleT outargsTs;
33525
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2539
        val T_compr = HOLogic.mk_ptupleT fp (rev Ts);
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2540
        val k = length outargs - 1;
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2541
        val arrange_bounds = map_index (fn (i, j) => (k-i, k-j)) outargs_bounds
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2542
          |> sort (prod_ord (K EQUAL) int_ord)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2543
          |> map fst;
33525
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2544
        val (outargsTs', outargsT) = split_last outargsTs;
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2545
        val (arrange, _) = fold_rev (fn U => fn (t, T) =>
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2546
            (HOLogic.split_const (U, T, T_compr) $ Abs ("", U, t),
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2547
             HOLogic.mk_prodT (U, T)))
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2548
          outargsTs' (Abs ("", outargsT,
05c384cb1181 Repaired handling of comprehensions in "values" command.
berghofe
parents: 33487
diff changeset
  2549
            HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds)), outargsT)
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2550
      in mk_map compfuns T_pred T_compr arrange t_pred end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2551
  in t_eval end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2552
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2553
fun eval thy param_user_modes (options as (depth_limit, (random, annotated))) t_compr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2554
  let
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33150
diff changeset
  2555
    val compfuns = if random then RandomPredCompFuns.compfuns else PredicateCompFuns.compfuns
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2556
    val t = analyze_compr thy compfuns param_user_modes options t_compr;
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2557
    val T = dest_predT compfuns (fastype_of t);
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2558
    val t' = mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2559
    val eval =
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2560
      if random then
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2561
        Code_ML.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2562
            (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2563
          |> Random_Engine.run
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2564
      else
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2565
        Code_ML.eval NONE ("Predicate_Compile_Core.eval_ref", eval_ref) Predicate.map thy t' []
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2566
  in (T, eval) end;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2567
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2568
fun values ctxt param_user_modes options k t_compr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2569
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2570
    val thy = ProofContext.theory_of ctxt;
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2571
    val (T, ts) = eval thy param_user_modes options t_compr;
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  2572
    val (ts, _) = Predicate.yieldn k ts;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2573
    val setT = HOLogic.mk_setT T;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2574
    val elemsT = HOLogic.mk_set T ts;
33476
27cca5416a88 Adopted output of values command
bulwahn
parents: 33474
diff changeset
  2575
    val cont = Free ("...", setT)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2576
  in if k = ~1 orelse length ts < k then elemsT
33476
27cca5416a88 Adopted output of values command
bulwahn
parents: 33474
diff changeset
  2577
    else Const (@{const_name Set.union}, setT --> setT --> setT) $ elemsT $ cont
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2578
  end;
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  2579
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2580
fun values_cmd print_modes param_user_modes options k raw_t state =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2581
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2582
    val ctxt = Toplevel.context_of state;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2583
    val t = Syntax.read_term ctxt raw_t;
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  2584
    val t' = values ctxt param_user_modes options k t;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2585
    val ty' = Term.type_of t';
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2586
    val ctxt' = Variable.auto_fixes t' ctxt;
33478
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
  2587
    val p = PrintMode.with_modes print_modes (fn () =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2588
      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2589
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2590
  in Pretty.writeln p end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2591
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2592
end;