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