src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML
author bulwahn
Wed, 29 Sep 2010 10:33:15 +0200
changeset 39785 05c4e9ecf5f6
parent 39783 a8c52b982ff2
child 40049 75d9f57123d6
permissions -rw-r--r--
improving the compilation to handle predicate arguments in higher-order argument positions
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
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
     9
  type mode = Predicate_Compile_Aux.mode
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    10
  type options = Predicate_Compile_Aux.options
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    11
  type compilation = Predicate_Compile_Aux.compilation
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    12
  type compilation_funs = Predicate_Compile_Aux.compilation_funs
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    13
  
33478
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    14
  val setup : theory -> theory
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    15
  val code_pred : options -> string -> Proof.context -> Proof.state
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    16
  val code_pred_cmd : options -> string -> Proof.context -> Proof.state
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    17
  val values_cmd : string list -> mode option list option
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    18
    -> ((string option * bool) * (compilation * int list)) -> int -> string -> Toplevel.state -> 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
    19
  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
    20
  val register_intros : string * thm list -> theory -> theory
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    21
  val is_registered : Proof.context -> string -> bool
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    22
  val function_name_of : compilation -> Proof.context -> string -> mode -> string
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    23
  val predfun_intro_of: Proof.context -> string -> mode -> thm
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    24
  val predfun_elim_of: Proof.context -> string -> mode -> thm
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    25
  val all_preds_of : Proof.context -> string list
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    26
  val modes_of: compilation -> Proof.context -> string -> mode list
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    27
  val all_modes_of : compilation -> Proof.context -> (string * mode list) list
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    28
  val all_random_modes_of : Proof.context -> (string * mode list) list
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
    29
  val intros_of : Proof.context -> string -> thm list
38072
7b8c295af291 exporting retrieval function for graph of introduction rules in the predicate compiler core
bulwahn
parents: 37678
diff changeset
    30
  val intros_graph_of : Proof.context -> thm list Graph.T
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
    31
  val add_intro : string option * thm -> theory -> theory
33478
b70fe083694d moved values command from core to predicate compile
bulwahn
parents: 33476
diff changeset
    32
  val set_elim : thm -> theory -> theory
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    33
  val register_alternative_function : string -> mode -> string -> theory -> theory
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    34
  val alternative_compilation_of_global : theory -> string -> mode ->
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    35
    (compilation_funs -> typ -> term) option
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    36
  val alternative_compilation_of : Proof.context -> string -> mode ->
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
    37
    (compilation_funs -> typ -> term) option
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    38
  val functional_compilation : string -> mode -> compilation_funs -> typ -> term
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    39
  val force_modes_and_functions : string -> (mode * (string * bool)) list -> theory -> theory
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
    40
  val force_modes_and_compilations : string ->
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    41
    (mode * ((compilation_funs -> typ -> term) * bool)) 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
    42
  val preprocess_intro : theory -> thm -> thm
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
    43
  val print_stored_rules : Proof.context -> unit
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
    44
  val print_all_modes : compilation -> Proof.context -> 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
    45
  val mk_casesrule : Proof.context -> term -> thm list -> term
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    46
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    47
  val put_pred_result : (unit -> term Predicate.pred) -> Proof.context -> Proof.context
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    48
  val put_pred_random_result : (unit -> int * int -> term Predicate.pred * (int * int)) ->
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    49
    Proof.context -> Proof.context
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    50
  val put_dseq_result : (unit -> term DSequence.dseq) -> Proof.context -> Proof.context
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    51
  val put_dseq_random_result : (unit -> int -> int -> int * int -> term DSequence.dseq * (int * int)) ->
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    52
    Proof.context -> Proof.context
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    53
  val put_lseq_random_result :
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    54
    (unit -> int -> int -> int * int -> int -> term Lazy_Sequence.lazy_sequence) ->
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    55
    Proof.context -> Proof.context
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    56
  val put_lseq_random_stats_result :
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    57
    (unit -> int -> int -> int * int -> int -> (term * int) Lazy_Sequence.lazy_sequence) ->
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    58
    Proof.context -> Proof.context
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
    59
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
    60
  val code_pred_intro_attrib : attribute
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    61
  (* used by Quickcheck_Generator *) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    62
  (* temporary for testing of the compilation *)
36048
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    63
  val add_equations : options -> string list -> theory -> theory
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    64
  val add_depth_limited_random_equations : options -> string list -> theory -> theory
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    65
  val add_random_dseq_equations : options -> string list -> theory -> theory
1d2faa488166 clarifying the Predicate_Compile_Core signature
bulwahn
parents: 36046
diff changeset
    66
  val add_new_random_dseq_equations : options -> string list -> theory -> theory
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
    67
  val mk_tracing : string -> term -> term
39310
17ef4415b17c removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents: 39299
diff changeset
    68
  val prepare_intrs : options -> Proof.context -> string list -> thm list ->
38957
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    69
    ((string * typ) list * string list * string list * (string * mode list) list *
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    70
      (string *  (Term.term list * Predicate_Compile_Aux.indprem list) list) list)
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
    71
  type mode_analysis_options =
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
    72
   {use_generators : bool,
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
    73
    reorder_premises : bool,
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
    74
    infer_pos_and_neg_modes : bool}  
38957
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    75
  datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
39310
17ef4415b17c removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents: 39299
diff changeset
    76
    | Mode_Pair of mode_derivation * mode_derivation | Term of mode
39461
0ed0f015d140 adding mode inference to prolog compilation; separate between (ad-hoc) code modifications and system_configuration; adapting quickcheck
bulwahn
parents: 39403
diff changeset
    77
  val head_mode_of : mode_derivation -> mode
38957
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    78
  type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    79
  type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    80
2eb265efa1b0 exporting mode analysis for use in prolog generation
bulwahn
parents: 38864
diff changeset
    81
  val infer_modes : 
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
    82
    mode_analysis_options -> options ->
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
    83
     (string -> Predicate_Compile_Aux.mode list) * (string -> Predicate_Compile_Aux.mode list)
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
    84
       * (string -> Predicate_Compile_Aux.mode -> bool) -> Proof.context -> (string * typ) list ->
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
    85
      (string * mode list) list ->
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
    86
      string list -> (string * (Term.term list * Predicate_Compile_Aux.indprem list) list) list ->
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
    87
      ((moded_clause list pred_mode_table * (string * mode list) list) * string list)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    88
end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    89
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    90
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    91
struct
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    92
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
    93
open Predicate_Compile_Aux;
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
    94
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    95
(** auxiliary **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    96
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    97
(* debug stuff *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
    98
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
    99
fun print_tac options s = 
33127
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
   100
  if show_proof_trace options then Tactical.print_tac s else Seq.single;
eb91ec1ef6f0 added option show_proof_trace
bulwahn
parents: 33126
diff changeset
   101
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
   102
fun assert b = if not b then raise Fail "Assertion failed" else warning "Assertion holds"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   103
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
   104
datatype assertion = Max_number_of_subgoals of int
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
   105
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
   106
  if (nprems_of st <= i) then Seq.single st
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
   107
  else raise Fail ("assert_tac: Numbers of subgoals mismatch at goal state :"
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
   108
    ^ "\n" ^ Pretty.string_of (Pretty.chunks
39125
f45d332a90e3 pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
wenzelm
parents: 39021
diff changeset
   109
      (Goal_Display.pretty_goals_without_context st)));
33108
9d9afd478016 added test for higher-order function inductification; added debug messages
bulwahn
parents: 33106
diff changeset
   110
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   111
(** fundamentals **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   112
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   113
(* syntactic operations *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   114
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   115
fun mk_eq (x, xs) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   116
  let fun mk_eqs _ [] = []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   117
        | mk_eqs a (b::cs) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   118
            HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   119
  in mk_eqs x xs end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   120
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   121
fun mk_scomp (t, u) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   122
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   123
    val T = fastype_of t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   124
    val U = fastype_of u
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   125
    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
   126
    val D = body_type U                   
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   127
  in 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   128
    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   129
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   130
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   131
fun dest_funT (Type ("fun",[S, T])) = (S, T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   132
  | dest_funT T = raise TYPE ("dest_funT", [T], [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   133
 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   134
fun mk_fun_comp (t, u) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   135
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   136
    val (_, B) = dest_funT (fastype_of t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   137
    val (C, A) = dest_funT (fastype_of u)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   138
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   139
    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   140
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   141
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   142
fun dest_randomT (Type ("fun", [@{typ Random.seed},
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   143
  Type (@{type_name Product_Type.prod}, [Type (@{type_name Product_Type.prod}, [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   144
  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   145
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   146
fun mk_tracing s t =
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   147
  Const(@{const_name Code_Evaluation.tracing},
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   148
    @{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
   149
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   150
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
   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
(* 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
   153
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   154
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
   155
  | Mode_Pair of mode_derivation * mode_derivation | Term of mode
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   156
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   157
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
   158
  "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
   159
  | 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
   160
  "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
   161
  | 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
   162
  | string_of_derivation (Context m) = "Context (" ^ string_of_mode m ^ ")"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   163
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   164
fun strip_mode_derivation deriv =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   165
  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
   166
    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
   167
      | 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
   168
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   169
    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
   170
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   171
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   172
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
   173
  | 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
   174
  | 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
   175
    (case mode_of d1 of Fun (m, m') =>
39193
0e505a4e500c stating errors in error messages more verbose in predicate compiler
bulwahn
parents: 39191
diff changeset
   176
        (if eq_mode (m, mode_of d2) then m' else raise Fail "mode_of: derivation has mismatching modes")
0e505a4e500c stating errors in error messages more verbose in predicate compiler
bulwahn
parents: 39191
diff changeset
   177
      | _ => raise Fail "mode_of: derivation has a non-functional 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
   178
  | 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
   179
    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
   180
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   181
fun head_mode_of deriv = mode_of (fst (strip_mode_derivation deriv))
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
fun param_derivations_of deriv =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   184
  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
   185
    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
   186
    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
   187
        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
   188
      | 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
   189
      | 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
   190
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   191
    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
   192
  end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   193
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   194
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
   195
      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
   196
  | 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
   197
      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
   198
  | 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
   199
  | collect_context_modes (Term _) = []
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   200
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   201
(* representation of inferred clauses with modes *)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   202
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
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
   204
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   205
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   206
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   207
(* 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
   208
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   209
datatype predfun_data = PredfunData of {
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   210
  definition : thm,
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   211
  intro : thm,
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   212
  elim : thm,
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   213
  neg_intro : thm option
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   214
};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   215
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   216
fun rep_predfun_data (PredfunData data) = data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   217
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   218
fun mk_predfun_data (definition, ((intro, elim), neg_intro)) =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   219
  PredfunData {definition = definition, intro = intro, elim = elim, neg_intro = neg_intro}
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   220
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   221
datatype pred_data = PredData of {
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   222
  intros : (string option * thm) list,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   223
  elim : thm option,
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   224
  preprocessed : 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
   225
  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
   226
  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
   227
  needs_random : mode list
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   228
};
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   229
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   230
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
   231
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   232
fun mk_pred_data (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))) =
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   233
  PredData {intros = intros, elim = elim, preprocessed = preprocessed,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   234
    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
   235
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   236
fun map_pred_data f (PredData {intros, elim, preprocessed, function_names, predfun_data, needs_random}) =
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   237
  mk_pred_data (f (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))))
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   238
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   239
fun eq_option eq (NONE, NONE) = true
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   240
  | eq_option eq (SOME x, SOME y) = eq (x, y)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   241
  | eq_option eq _ = false
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   242
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   243
fun eq_pred_data (PredData d1, PredData d2) = 
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   244
  eq_list (eq_pair (op =) Thm.eq_thm) (#intros d1, #intros d2) andalso
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
   245
  eq_option Thm.eq_thm (#elim d1, #elim d2)
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   246
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   247
structure PredData = Theory_Data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   248
(
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   249
  type T = pred_data Graph.T;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   250
  val empty = Graph.empty;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   251
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   252
  val merge = Graph.merge eq_pred_data;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   253
);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   254
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   255
(* queries *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   256
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   257
fun lookup_pred_data ctxt name =
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   258
  Option.map rep_pred_data (try (Graph.get_node (PredData.get (ProofContext.theory_of ctxt))) name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   259
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   260
fun the_pred_data ctxt name = case lookup_pred_data ctxt name
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   261
 of NONE => error ("No such predicate " ^ quote name)  
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   262
  | SOME data => data;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   263
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   264
val is_registered = is_some oo lookup_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   265
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   266
val all_preds_of = Graph.keys o PredData.get o ProofContext.theory_of
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   267
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   268
val intros_of = map snd o #intros oo the_pred_data
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   269
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   270
val names_of = map fst o #intros oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   271
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   272
fun the_elim_of ctxt name = case #elim (the_pred_data ctxt name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   273
 of NONE => error ("No elimination rule for predicate " ^ quote name)
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   274
  | SOME thm => thm
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   275
  
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   276
val has_elim = is_some o #elim oo the_pred_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   277
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   278
fun function_names_of compilation ctxt name =
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   279
  case AList.lookup (op =) (#function_names (the_pred_data ctxt name)) compilation 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
   280
    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
   281
      ^ "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
   282
  | SOME fun_names => fun_names
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   283
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   284
fun function_name_of compilation ctxt name mode =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   285
  case AList.lookup eq_mode
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   286
    (function_names_of compilation ctxt name) mode 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
   287
    NONE => error ("No " ^ string_of_compilation compilation
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   288
      ^ " function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote 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
   289
  | 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
   290
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   291
fun modes_of compilation ctxt name = map fst (function_names_of compilation ctxt name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   292
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   293
fun all_modes_of compilation ctxt =
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   294
  map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name))
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   295
    (all_preds_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
   296
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   297
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
   298
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   299
fun defined_functions compilation ctxt name = case lookup_pred_data ctxt name of
37001
bcffdb899167 improved behaviour of defined_functions in the predicate compiler
bulwahn
parents: 36692
diff changeset
   300
    NONE => false
bcffdb899167 improved behaviour of defined_functions in the predicate compiler
bulwahn
parents: 36692
diff changeset
   301
  | SOME data => AList.defined (op =) (#function_names data) compilation
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   302
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   303
fun needs_random ctxt s m =
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   304
  member (op =) (#needs_random (the_pred_data ctxt s)) m
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   305
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   306
fun lookup_predfun_data ctxt 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
   307
  Option.map rep_predfun_data
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   308
    (AList.lookup (op =) (#predfun_data (the_pred_data ctxt name)) mode)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   309
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   310
fun the_predfun_data ctxt name mode =
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   311
  case lookup_predfun_data ctxt name mode 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
   312
    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
   313
      " 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
   314
  | SOME data => data;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   315
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   316
val predfun_definition_of = #definition ooo the_predfun_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   317
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   318
val predfun_intro_of = #intro ooo the_predfun_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   319
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   320
val predfun_elim_of = #elim ooo the_predfun_data
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   321
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   322
val predfun_neg_intro_of = #neg_intro ooo the_predfun_data
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   323
38072
7b8c295af291 exporting retrieval function for graph of introduction rules in the predicate compiler core
bulwahn
parents: 37678
diff changeset
   324
val intros_graph_of =
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   325
  Graph.map (K (map snd o #intros o rep_pred_data)) o PredData.get o ProofContext.theory_of
38072
7b8c295af291 exporting retrieval function for graph of introduction rules in the predicate compiler core
bulwahn
parents: 37678
diff changeset
   326
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   327
(* diagnostic display functions *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   328
37004
c62f743e37d4 removing unused argument in print_modes function
bulwahn
parents: 37003
diff changeset
   329
fun print_modes options modes =
33251
4b13ab778b78 added option show_modes to predicate compiler
bulwahn
parents: 33250
diff changeset
   330
  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
   331
    tracing ("Inferred modes:\n" ^
33251
4b13ab778b78 added option show_modes to predicate compiler
bulwahn
parents: 33250
diff changeset
   332
      cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   333
        (fn (p, m) => string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
33251
4b13ab778b78 added option show_modes to predicate compiler
bulwahn
parents: 33250
diff changeset
   334
  else ()
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   335
37004
c62f743e37d4 removing unused argument in print_modes function
bulwahn
parents: 37003
diff changeset
   336
fun print_pred_mode_table string_of_entry pred_mode_table =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   337
  let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   338
    fun print_mode pred ((pol, 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
   339
      ^ string_of_entry pred mode entry
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   340
    fun print_pred (pred, modes) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   341
      "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
   342
    val _ = tracing (cat_lines (map print_pred pred_mode_table))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   343
  in () end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   344
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   345
fun string_of_prem ctxt (Prem t) =
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   346
    (Syntax.string_of_term ctxt t) ^ "(premise)"
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   347
  | string_of_prem ctxt (Negprem t) =
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   348
    (Syntax.string_of_term ctxt (HOLogic.mk_not t)) ^ "(negative premise)"
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   349
  | string_of_prem ctxt (Sidecond t) =
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   350
    (Syntax.string_of_term ctxt t) ^ "(sidecondition)"
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   351
  | string_of_prem ctxt _ = raise Fail "string_of_prem: unexpected input"
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   352
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   353
fun string_of_clause ctxt pred (ts, prems) =
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   354
  (space_implode " --> "
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   355
  (map (string_of_prem ctxt) prems)) ^ " --> " ^ pred ^ " "
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   356
   ^ (space_implode " " (map (Syntax.string_of_term ctxt) ts))
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
   357
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
   358
fun print_compiled_terms options ctxt =
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   359
  if show_compilation options then
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
   360
    print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term ctxt)
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   361
  else K ()
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   362
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   363
fun print_stored_rules ctxt =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   364
  let
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   365
    val preds = Graph.keys (PredData.get (ProofContext.theory_of ctxt))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   366
    fun print pred () = let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   367
      val _ = writeln ("predicate: " ^ pred)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   368
      val _ = writeln ("introrules: ")
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   369
      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm ctxt thm))
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   370
        (rev (intros_of ctxt pred)) ()
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   371
    in
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   372
      if (has_elim ctxt pred) then
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   373
        writeln ("elimrule: " ^ Display.string_of_thm ctxt (the_elim_of ctxt pred))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   374
      else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   375
        writeln ("no elimrule defined")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   376
    end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   377
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   378
    fold print preds ()
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   379
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   380
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   381
fun print_all_modes compilation ctxt =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   382
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   383
    val _ = writeln ("Inferred modes:")
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   384
    fun print (pred, modes) u =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   385
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   386
        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
   387
        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
   388
      in u end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   389
  in
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   390
    fold print (all_modes_of compilation ctxt) ()
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   391
  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
   392
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
   393
(* validity checks *)
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
   394
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   395
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
   396
  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
   397
    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
   398
      SOME modes =>
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   399
        let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   400
          val modes' = map snd modes
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   401
        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
   402
          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
   403
            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
   404
            ^ "  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
   405
            ^ "  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
   406
          else ()
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   407
        end
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   408
      | NONE => ())
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   409
  | NONE => ()
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   410
39201
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
   411
fun check_proposed_modes preds options modes errors =
39382
c797f3ab2ae1 proposed modes for code_pred now supports modes for mutual predicates
bulwahn
parents: 39310
diff changeset
   412
  map (fn (s, _) => case proposed_modes options s of
c797f3ab2ae1 proposed modes for code_pred now supports modes for mutual predicates
bulwahn
parents: 39310
diff changeset
   413
    SOME ms => (case AList.lookup (op =) modes s of
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   414
      SOME inferred_ms =>
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   415
        let
39201
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
   416
          val preds_without_modes = map fst (filter (null o snd) modes)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
   417
          val modes' = map snd inferred_ms
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   418
        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
   419
          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
   420
            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
   421
            ^ "  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
   422
            ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms) ^ "\n"
39383
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   423
            ^ (if show_invalid_clauses options then
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   424
            ("For the following clauses, the following modes could not be inferred: " ^ "\n"
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   425
            ^ cat_lines errors) else "") ^
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   426
            (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
   427
              "\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
   428
            else ""))
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   429
          else ()
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   430
        end
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
   431
      | NONE => ())
39382
c797f3ab2ae1 proposed modes for code_pred now supports modes for mutual predicates
bulwahn
parents: 39310
diff changeset
   432
  | NONE => ()) preds
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
   433
39651
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   434
fun check_matches_type ctxt predname T ms =
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   435
  let
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   436
    fun check (m as Fun (m1, m2)) (Type("fun", [T1,T2])) = check m1 T1 andalso check m2 T2
39763
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   437
      | check m (T as Type("fun", _)) = (m = Input orelse m = Output)
39651
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   438
      | check (Pair (m1, m2)) (Type (@{type_name Product_Type.prod}, [T1, T2])) =
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   439
          check m1 T1 andalso check m2 T2 
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   440
      | check Input T = true
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   441
      | check Output T = true
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   442
      | check Bool @{typ bool} = true
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   443
      | check _ _ = false
39763
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   444
    fun check_consistent_modes ms =
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   445
      if forall (fn Fun (m1', m2') => true | _ => false) ms then
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   446
        pairself check_consistent_modes (split_list (map (fn Fun (m1, m2) => (m1, m2)) ms))
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   447
        |> (fn (res1, res2) => res1 andalso res2) 
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   448
      else if forall (fn Input => true | Output => true | Pair _ => true | _ => false) ms then
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   449
        true
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   450
      else if forall (fn Bool => true | _ => false) ms then
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   451
        true
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   452
      else
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   453
        false
39651
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   454
    val _ = map
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   455
      (fn mode =>
39763
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   456
        if length (strip_fun_mode mode) = length (binder_types T)
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   457
          andalso (forall (uncurry check) (strip_fun_mode mode ~~ binder_types T)) then ()
39651
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   458
        else error (string_of_mode mode ^ " is not a valid mode for " ^ Syntax.string_of_typ ctxt T
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   459
        ^ " at predicate " ^ predname)) ms
39763
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   460
    val _ =
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   461
     if check_consistent_modes ms then ()
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   462
     else error (commas (map string_of_mode ms) ^
03f95582ef63 weakening check for higher-order relations, but adding check for consistent modes
bulwahn
parents: 39762
diff changeset
   463
       " are inconsistent modes for predicate " ^ predname)
39651
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   464
  in
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   465
    ms
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   466
  end
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   467
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
   468
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
   469
(* 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
   470
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   471
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
   472
  (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
   473
     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
   474
     fun varify (t, (i, ts)) =
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35411
diff changeset
   475
       let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global [] t))
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   476
       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
   477
     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
   478
     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
   479
     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
   480
     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
   481
     fun unify (cname, cT) =
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33268
diff changeset
   482
       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
   483
       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
   484
     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
   485
     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
   486
   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
   487
   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
   488
     (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
   489
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   490
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
   491
  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
   492
    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
   493
    val T = fastype_of outp_pred
39299
e98a06145530 directly computing the values of interest instead of composing functions in an unintelligent way that causes exponential much garbage; using the latest theory
bulwahn
parents: 39273
diff changeset
   494
    val paramTs = ho_argsT_of_typ (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
   495
    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
   496
      (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
   497
    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
   498
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   499
    (((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
   500
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   501
  | 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
   502
    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
   503
      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
   504
      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
   505
      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
   506
      val T = fastype_of pred
39299
e98a06145530 directly computing the values of interest instead of composing functions in an unintelligent way that causes exponential much garbage; using the latest theory
bulwahn
parents: 39273
diff changeset
   507
      val ho_args = ho_args_of_typ T args
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   508
      fun subst_of (pred', pred) =
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   509
        let
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   510
          val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
39685
d8071cddb877 actually handle Type.TYPE_MATCH, not arbitrary exceptions;
wenzelm
parents: 39670
diff changeset
   511
            handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred)
39653
51e23b48a202 handling TYPE_MATCH error by raising user error message if user gives introduction rules with mismatching types
bulwahn
parents: 39652
diff changeset
   512
            ^ " (trying to match " ^ Syntax.string_of_typ ctxt (fastype_of pred')
51e23b48a202 handling TYPE_MATCH error by raising user error message if user gives introduction rules with mismatching types
bulwahn
parents: 39652
diff changeset
   513
            ^ " and " ^ Syntax.string_of_typ ctxt (fastype_of pred) ^ ")"
51e23b48a202 handling TYPE_MATCH error by raising user error message if user gives introduction rules with mismatching types
bulwahn
parents: 39652
diff changeset
   514
            ^ " in " ^ Display.string_of_thm ctxt th)
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   515
        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
   516
      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
   517
        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
   518
          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
   519
          val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
   520
            raise Fail "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
   521
        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
   522
      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
   523
        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
   524
          val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of) th
39299
e98a06145530 directly computing the values of interest instead of composing functions in an unintelligent way that causes exponential much garbage; using the latest theory
bulwahn
parents: 39273
diff changeset
   525
          val ho_args' = map dest_Var (ho_args_of_typ T 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
   526
        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
   527
      val outp_pred =
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   528
        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
   529
      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
   530
        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
   531
    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
   532
      (((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
   533
    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
   534
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
   535
(* 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
   536
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   537
fun mk_args2 (Type (@{type_name Product_Type.prod}, [T1, T2])) st =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   538
    let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   539
      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
   540
      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
   541
    in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   542
      (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
   543
    end
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   544
  (*| mk_args2 (T as Type ("fun", _)) (params, 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
   545
    let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   546
      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
   547
    in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   548
      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
   549
        (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
   550
      else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   551
        let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   552
          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
   553
        in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   554
          (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
   555
        end
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   556
    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
   557
  | 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
   558
    let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   559
      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
   560
    in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   561
      (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
   562
    end
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   563
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   564
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
   565
  let
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   566
    (* TODO: can be simplified if parameters are not treated specially ? *)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   567
    val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   568
    (* TODO: distinct required ? -- test case with more than one parameter! *)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   569
    val params = distinct (op aconv) params
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
   570
    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
   571
    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
   572
    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
   573
    val argsT = binder_types (fastype_of pred)
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   574
    (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   575
    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
   576
    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
   577
      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
   578
        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
   579
        val prems = Logic.strip_imp_prems intro
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   580
        val eqprems =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   581
          map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   582
        val frees = map Free (fold Term.add_frees (args @ prems) [])
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
   583
      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
   584
    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
   585
    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
   586
  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
   587
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   588
fun dest_conjunct_prem th =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   589
  case HOLogic.dest_Trueprop (prop_of th) of
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38558
diff changeset
   590
    (Const (@{const_name HOL.conj}, _) $ t $ t') =>
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   591
      dest_conjunct_prem (th RS @{thm conjunct1})
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   592
        @ dest_conjunct_prem (th RS @{thm conjunct2})
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   593
    | _ => [th]
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   594
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   595
fun prove_casesrule ctxt (pred, (pre_cases_rule, nparams)) cases_rule =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   596
  let
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   597
    val thy = ProofContext.theory_of ctxt
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   598
    val nargs = length (binder_types (fastype_of pred))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   599
    fun PEEK f dependent_tactic st = dependent_tactic (f st) st
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   600
    fun meta_eq_of th = th RS @{thm eq_reflection}
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   601
    val tuple_rew_rules = map meta_eq_of [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}]
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   602
    fun instantiate i n {context = ctxt, params = p, prems = prems,
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   603
      asms = a, concl = cl, schematics = s}  =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   604
      let
36510
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   605
        fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t)
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   606
        fun inst_of_matches tts = fold (Pattern.match thy) tts (Vartab.empty, Vartab.empty)
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   607
          |> snd |> Vartab.dest |> map (pairself (cterm_of thy) o term_pair_of)
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   608
        val (cases, (eqs, prems)) = apsnd (chop (nargs - nparams)) (chop n prems)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   609
        val case_th = MetaSimplifier.simplify true
36509
9cff57fc7113 improving proof procedure for transforming cases rule in the predicate compiler to handle free variables of function type
bulwahn
parents: 36323
diff changeset
   610
          (@{thm Predicate.eq_is_eq} :: map meta_eq_of eqs) (nth cases (i - 1))
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   611
        val prems' = maps (dest_conjunct_prem o MetaSimplifier.simplify true tuple_rew_rules) prems
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   612
        val pats = map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop) (take nargs (prems_of case_th))
36510
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   613
        val case_th' = Thm.instantiate ([], inst_of_matches pats) case_th
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   614
          OF (replicate nargs @{thm refl})
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   615
        val thesis =
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   616
          Thm.instantiate ([], inst_of_matches (prems_of case_th' ~~ map prop_of prems')) case_th'
d716eb002b9f removed local clone in the predicate compiler
bulwahn
parents: 36509
diff changeset
   617
            OF prems'
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   618
      in
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   619
        (rtac thesis 1)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   620
      end
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   621
    val tac =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   622
      etac pre_cases_rule 1
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   623
      THEN
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   624
      (PEEK nprems_of
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   625
        (fn n =>
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   626
          ALLGOALS (fn i =>
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   627
            MetaSimplifier.rewrite_goal_tac [@{thm split_paired_all}] i
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   628
            THEN (SUBPROOF (instantiate i n) ctxt i))))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   629
  in
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   630
    Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule (fn _ => tac)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   631
  end
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   632
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   633
(** preprocessing rules **)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   634
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   635
fun Trueprop_conv cv ct =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   636
  case Thm.term_of ct of
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   637
    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
   638
  | _ => raise Fail "Trueprop_conv"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   639
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   640
fun preprocess_equality thy rule =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   641
  Conv.fconv_rule
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   642
    (imp_prems_conv
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   643
      (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
   644
    (Thm.transfer thy rule)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   645
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   646
fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   647
39658
b3644e40f661 removing unneccessary expansion procedure for elimination rules; removing obsolete elim preprocessing; tuned
bulwahn
parents: 39657
diff changeset
   648
(* fetching introduction rules or registering introduction rules *)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   649
35887
f704ba9875f6 making flat triples to nested tuple to remove general triple functions
bulwahn
parents: 35886
diff changeset
   650
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
   651
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   652
fun fetch_pred_data ctxt name =
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   653
  case try (Inductive.the_inductive ctxt) name of
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   654
    SOME (info as (_, result)) => 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   655
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   656
        fun is_intro_of intro =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   657
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   658
            val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   659
          in (fst (dest_Const const) = name) end;
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
   660
        val thy = ProofContext.theory_of ctxt
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   661
        val intros = map (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
   662
        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
   663
        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
   664
        val pred = nth (#preds result) index
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   665
        val elim_t = mk_casesrule ctxt pred intros
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
   666
        val nparams = length (Inductive.params_of (#raw_induct result))
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   667
        val elim = prove_casesrule ctxt (pred, (pre_elim, nparams)) elim_t
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   668
      in
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   669
        mk_pred_data (((map (pair NONE) intros, SOME elim), true), 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
   670
      end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   671
  | NONE => error ("No such predicate: " ^ quote name)
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   672
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   673
fun add_predfun_data name mode data =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   674
  let
35887
f704ba9875f6 making flat triples to nested tuple to remove general triple functions
bulwahn
parents: 35886
diff changeset
   675
    val add = (apsnd o apsnd o apfst) (cons (mode, mk_predfun_data data))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   676
  in PredData.map (Graph.map_node name (map_pred_data add)) end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   677
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   678
fun is_inductive_predicate ctxt name =
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   679
  is_some (try (Inductive.the_inductive ctxt) name)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   680
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   681
fun depending_preds_of ctxt (key, value) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   682
  let
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   683
    val intros = map (Thm.prop_of o snd) ((#intros o rep_pred_data) value)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   684
  in
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   685
    fold Term.add_const_names intros []
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   686
      |> (fn cs =>
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   687
        if member (op =) cs @{const_name "HOL.eq"} then
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   688
          insert (op =) @{const_name Predicate.eq} cs
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   689
        else cs)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   690
      |> filter (fn c => (not (c = key)) andalso
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   691
        (is_inductive_predicate ctxt c orelse is_registered ctxt c))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   692
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   693
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
   694
fun add_intro (opt_case_name, thm) thy =
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
   695
  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
   696
    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
   697
    fun cons_intro gr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   698
     case try (Graph.get_node gr) name of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   699
       SOME pred_data => Graph.map_node name (map_pred_data
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   700
         (apfst (apfst (apfst (fn intros => intros @ [(opt_case_name, thm)]))))) gr
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   701
     | NONE => Graph.new_node (name, mk_pred_data ((([(opt_case_name, thm)], NONE), false), no_compilation)) gr
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   702
  in PredData.map cons_intro thy end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   703
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
   704
fun set_elim thm =
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
   705
  let
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   706
    val (name, _) = dest_Const (fst 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   707
      (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   708
  in PredData.map (Graph.map_node name (map_pred_data (apfst (apfst (apsnd (K (SOME thm))))))) end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   709
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   710
fun register_predicate (constname, intros, elim) thy =
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
   711
  let
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   712
    val named_intros = map (pair NONE) intros
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   713
  in
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   714
    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
   715
      PredData.map
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
   716
        (Graph.new_node (constname,
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   717
          mk_pred_data (((named_intros, SOME elim), false), no_compilation))) thy
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   718
    else thy
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   719
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   720
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   721
fun register_intros (constname, pre_intros) thy =
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
   722
  let
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   723
    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
   724
    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
   725
    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
   726
      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
   727
        "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
   728
          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
   729
      else ()
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   730
    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
   731
    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
   732
      (Drule.export_without_context o Skip_Proof.make_thm thy)
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36533
diff changeset
   733
      (mk_casesrule (ProofContext.init_global thy) pred 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
   734
  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
   735
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   736
fun defined_function_of compilation pred =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   737
  let
35887
f704ba9875f6 making flat triples to nested tuple to remove general triple functions
bulwahn
parents: 35886
diff changeset
   738
    val set = (apsnd o apfst) (cons (compilation, []))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   739
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   740
    PredData.map (Graph.map_node pred (map_pred_data set))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   741
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   742
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   743
fun set_function_name compilation pred mode name =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   744
  let
35887
f704ba9875f6 making flat triples to nested tuple to remove general triple functions
bulwahn
parents: 35886
diff changeset
   745
    val set = (apsnd o apfst)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   746
      (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
   747
  in
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   748
    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
   749
  end
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
   750
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   751
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
   752
  let
35887
f704ba9875f6 making flat triples to nested tuple to remove general triple functions
bulwahn
parents: 35886
diff changeset
   753
    val set = (apsnd o apsnd o apsnd) (K modes)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   754
  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
   755
    PredData.map (Graph.map_node name (map_pred_data set))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   756
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
   757
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   758
(* registration of alternative function names *)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   759
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   760
structure Alt_Compilations_Data = Theory_Data
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   761
(
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   762
  type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table;
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   763
  val empty = Symtab.empty;
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   764
  val extend = I;
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38757
diff changeset
   765
  fun merge data : T = Symtab.merge (K true) data;
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   766
);
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   767
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   768
fun alternative_compilation_of_global thy pred_name mode =
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   769
  AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get thy) pred_name) mode
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   770
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   771
fun alternative_compilation_of ctxt pred_name mode =
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   772
  AList.lookup eq_mode
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   773
    (Symtab.lookup_list (Alt_Compilations_Data.get (ProofContext.theory_of ctxt)) pred_name) mode
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
   774
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   775
fun force_modes_and_compilations pred_name compilations =
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   776
  let
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   777
    (* thm refl is a dummy thm *)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   778
    val modes = map fst compilations
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   779
    val (needs_random, non_random_modes) = pairself (map fst)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   780
      (List.partition (fn (m, (fun_name, random)) => random) compilations)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   781
    val non_random_dummys = map (rpair "dummy") non_random_modes
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   782
    val all_dummys = map (rpair "dummy") modes
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   783
    val dummy_function_names = map (rpair all_dummys) Predicate_Compile_Aux.random_compilations
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   784
      @ map (rpair non_random_dummys) Predicate_Compile_Aux.non_random_compilations
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   785
    val alt_compilations = map (apsnd fst) compilations
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   786
  in
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   787
    PredData.map (Graph.new_node
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
   788
      (pred_name, mk_pred_data ((([], SOME @{thm refl}), true), (dummy_function_names, ([], needs_random)))))
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   789
    #> Alt_Compilations_Data.map (Symtab.insert (K false) (pred_name, alt_compilations))
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   790
  end
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   791
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   792
fun functional_compilation fun_name mode compfuns T =
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   793
  let
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   794
    val (inpTs, outpTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE))
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   795
      mode (binder_types T)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   796
    val bs = map (pair "x") inpTs
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   797
    val bounds = map Bound (rev (0 upto (length bs) - 1))
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   798
    val f = Const (fun_name, inpTs ---> HOLogic.mk_tupleT outpTs)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   799
  in list_abs (bs, mk_single compfuns (list_comb (f, bounds))) end
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   800
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   801
fun register_alternative_function pred_name mode fun_name =
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   802
  Alt_Compilations_Data.map (Symtab.insert_list (eq_pair eq_mode (K false))
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   803
    (pred_name, (mode, functional_compilation fun_name mode)))
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   804
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   805
fun force_modes_and_functions pred_name fun_names =
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   806
  force_modes_and_compilations pred_name
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   807
    (map (fn (mode, (fun_name, random)) => (mode, (functional_compilation fun_name mode, random)))
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   808
    fun_names)
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
   809
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   810
(* compilation modifiers *)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   811
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   812
structure Comp_Mod =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   813
struct
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   814
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   815
datatype comp_modifiers = Comp_Modifiers of
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   816
{
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   817
  compilation : compilation,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   818
  function_name_prefix : string,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   819
  compfuns : compilation_funs,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   820
  mk_random : typ -> term list -> term,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   821
  modify_funT : typ -> typ,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   822
  additional_arguments : string list -> term list,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   823
  wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   824
  transform_additional_arguments : indprem -> term list -> term list
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   825
}
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   826
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   827
fun dest_comp_modifiers (Comp_Modifiers c) = c
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   828
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   829
val compilation = #compilation o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   830
val function_name_prefix = #function_name_prefix o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   831
val compfuns = #compfuns o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   832
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   833
val mk_random = #mk_random o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   834
val funT_of' = funT_of o compfuns
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   835
val modify_funT = #modify_funT o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   836
fun funT_of comp mode = modify_funT comp o funT_of' comp mode
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   837
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   838
val additional_arguments = #additional_arguments o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   839
val wrap_compilation = #wrap_compilation o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   840
val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   841
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   842
end;
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   843
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   844
val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   845
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   846
  compilation = Depth_Limited,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   847
  function_name_prefix = "depth_limited_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   848
  compfuns = PredicateCompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   849
  mk_random = (fn _ => error "no random generation"),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   850
  additional_arguments = fn names =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   851
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   852
      val depth_name = Name.variant names "depth"
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   853
    in [Free (depth_name, @{typ code_numeral})] end,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   854
  modify_funT = (fn T => let val (Ts, U) = strip_type T
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   855
  val Ts' = [@{typ code_numeral}] in (Ts @ Ts') ---> U end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   856
  wrap_compilation =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   857
    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   858
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   859
      val [depth] = additional_arguments
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   860
      val (_, Ts) = split_modeT' mode (binder_types T)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   861
      val T' = mk_predT compfuns (HOLogic.mk_tupleT Ts)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   862
      val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   863
    in
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   864
      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   865
        $ mk_bot compfuns (dest_predT compfuns T')
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   866
        $ compilation
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   867
    end,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   868
  transform_additional_arguments =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   869
    fn prem => fn additional_arguments =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   870
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   871
      val [depth] = additional_arguments
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   872
      val depth' =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   873
        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   874
          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   875
    in [depth'] end
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   876
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   877
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   878
val random_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   879
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   880
  compilation = Random,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   881
  function_name_prefix = "random_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   882
  compfuns = PredicateCompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   883
  mk_random = (fn T => fn additional_arguments =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   884
  list_comb (Const(@{const_name Quickcheck.iter},
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   885
  [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}] ---> 
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   886
    PredicateCompFuns.mk_predT T), additional_arguments)),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   887
  modify_funT = (fn T =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   888
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   889
      val (Ts, U) = strip_type T
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   890
      val Ts' = [@{typ code_numeral}, @{typ code_numeral}, @{typ "code_numeral * code_numeral"}]
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   891
    in (Ts @ Ts') ---> U end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   892
  additional_arguments = (fn names =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   893
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   894
      val [nrandom, size, seed] = Name.variant_list names ["nrandom", "size", "seed"]
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   895
    in
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   896
      [Free (nrandom, @{typ code_numeral}), Free (size, @{typ code_numeral}),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   897
        Free (seed, @{typ "code_numeral * code_numeral"})]
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   898
    end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   899
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   900
    : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   901
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   902
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   903
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   904
val depth_limited_random_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   905
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   906
  compilation = Depth_Limited_Random,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   907
  function_name_prefix = "depth_limited_random_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   908
  compfuns = PredicateCompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   909
  mk_random = (fn T => fn additional_arguments =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   910
  list_comb (Const(@{const_name Quickcheck.iter},
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   911
  [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}] ---> 
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   912
    PredicateCompFuns.mk_predT T), tl additional_arguments)),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   913
  modify_funT = (fn T =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   914
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   915
      val (Ts, U) = strip_type T
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   916
      val Ts' = [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   917
        @{typ "code_numeral * code_numeral"}]
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   918
    in (Ts @ Ts') ---> U end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   919
  additional_arguments = (fn names =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   920
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   921
      val [depth, nrandom, size, seed] = Name.variant_list names ["depth", "nrandom", "size", "seed"]
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   922
    in
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   923
      [Free (depth, @{typ code_numeral}), Free (nrandom, @{typ code_numeral}),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   924
        Free (size, @{typ code_numeral}), Free (seed, @{typ "code_numeral * code_numeral"})]
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   925
    end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   926
  wrap_compilation =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   927
  fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   928
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   929
      val depth = hd (additional_arguments)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   930
      val (_, Ts) = split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   931
        mode (binder_types T)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   932
      val T' = mk_predT compfuns (HOLogic.mk_tupleT Ts)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   933
      val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   934
    in
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   935
      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   936
        $ mk_bot compfuns (dest_predT compfuns T')
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   937
        $ compilation
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   938
    end,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   939
  transform_additional_arguments =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   940
    fn prem => fn additional_arguments =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   941
    let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   942
      val [depth, nrandom, size, seed] = additional_arguments
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   943
      val depth' =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   944
        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   945
          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   946
    in [depth', nrandom, size, seed] end
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   947
}
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   948
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   949
val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   950
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   951
  compilation = Pred,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   952
  function_name_prefix = "",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   953
  compfuns = PredicateCompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   954
  mk_random = (fn _ => error "no random generation"),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   955
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   956
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   957
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   958
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   959
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   960
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   961
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   962
val annotated_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   963
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   964
  compilation = Annotated,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   965
  function_name_prefix = "annotated_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   966
  compfuns = PredicateCompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   967
  mk_random = (fn _ => error "no random generation"),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   968
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   969
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   970
  wrap_compilation =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   971
    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   972
      mk_tracing ("calling predicate " ^ s ^
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   973
        " with mode " ^ string_of_mode mode) compilation,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   974
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   975
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   976
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   977
val dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   978
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   979
  compilation = DSeq,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   980
  function_name_prefix = "dseq_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   981
  compfuns = DSequence_CompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   982
  mk_random = (fn _ => error "no random generation"),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   983
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   984
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   985
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   986
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   987
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   988
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   989
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   990
val pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   991
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   992
  compilation = Pos_Random_DSeq,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   993
  function_name_prefix = "random_dseq_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   994
  compfuns = Random_Sequence_CompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   995
  mk_random = (fn T => fn additional_arguments =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   996
  let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   997
    val random = Const ("Quickcheck.random_class.random",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   998
      @{typ code_numeral} --> @{typ Random.seed} -->
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
   999
        HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed}))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1000
  in
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1001
    Const ("Random_Sequence.Random", (@{typ code_numeral} --> @{typ Random.seed} -->
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1002
      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1003
      Random_Sequence_CompFuns.mk_random_dseqT T) $ random
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1004
  end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1005
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1006
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1007
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1008
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1009
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1010
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1011
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1012
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1013
val neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1014
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1015
  compilation = Neg_Random_DSeq,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1016
  function_name_prefix = "random_dseq_neg_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1017
  compfuns = Random_Sequence_CompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1018
  mk_random = (fn _ => error "no random generation"),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1019
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1020
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1021
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1022
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1023
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1024
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1025
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1026
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1027
val new_pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1028
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1029
  compilation = New_Pos_Random_DSeq,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1030
  function_name_prefix = "new_random_dseq_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1031
  compfuns = New_Pos_Random_Sequence_CompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1032
  mk_random = (fn T => fn additional_arguments =>
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1033
  let
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1034
    val random = Const ("Quickcheck.random_class.random",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1035
      @{typ code_numeral} --> @{typ Random.seed} -->
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1036
        HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed}))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1037
  in
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1038
    Const ("New_Random_Sequence.Random", (@{typ code_numeral} --> @{typ Random.seed} -->
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1039
      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1040
      New_Pos_Random_Sequence_CompFuns.mk_pos_random_dseqT T) $ random
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1041
  end),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1042
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1043
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1044
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1045
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1046
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1047
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1048
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1049
val new_neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1050
  {
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1051
  compilation = New_Neg_Random_DSeq,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1052
  function_name_prefix = "new_random_dseq_neg_",
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1053
  compfuns = New_Neg_Random_Sequence_CompFuns.compfuns,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1054
  mk_random = (fn _ => error "no random generation"),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1055
  modify_funT = I,
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1056
  additional_arguments = K [],
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1057
  wrap_compilation = K (K (K (K (K I))))
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1058
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1059
  transform_additional_arguments = K I : (indprem -> term list -> term list)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1060
  }
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1061
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1062
fun negative_comp_modifiers_of comp_modifiers =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1063
    (case Comp_Mod.compilation comp_modifiers of
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1064
      Pos_Random_DSeq => neg_random_dseq_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1065
    | Neg_Random_DSeq => pos_random_dseq_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1066
    | New_Pos_Random_DSeq => new_neg_random_dseq_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1067
    | New_Neg_Random_DSeq => new_pos_random_dseq_comp_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1068
    | c => comp_modifiers)
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1069
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1070
(** mode analysis **)
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  1071
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  1072
type mode_analysis_options =
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  1073
  {use_generators : bool,
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  1074
  reorder_premises : bool,
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  1075
  infer_pos_and_neg_modes : bool}
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1076
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1077
fun is_constrt thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1078
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1079
    val cnstrs = flat (maps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1080
      (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
  1081
      (Symtab.dest (Datatype.get_all thy)));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1082
    fun check t = (case strip_comb t of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1083
        (Free _, []) => true
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1084
      | (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
  1085
            (SOME (i, Tname), Type (Tname', _)) =>
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  1086
              length ts = i andalso Tname = Tname' andalso forall check ts
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1087
          | _ => false)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1088
      | _ => false)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1089
  in check end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1090
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1091
(*** check if a type is an equality type (i.e. doesn't contain fun)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1092
  FIXME this is only an approximation ***)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1093
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1094
  | is_eqT _ = true;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1095
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1096
fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1097
val terms_vs = distinct (op =) o maps term_vs;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1098
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1099
(** collect all Frees in a term (with duplicates!) **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1100
fun term_vTs tm =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1101
  fold_aterms (fn Free xT => cons xT | _ => I) tm [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1102
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1103
fun subsets i j =
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1104
  if i <= j then
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1105
    let
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1106
      fun merge xs [] = xs
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1107
        | merge [] ys = ys
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1108
        | 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
  1109
            else y::merge (x::xs) ys;
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1110
      val is = subsets (i+1) j
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1111
    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
  1112
  else [[]];
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1113
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1114
fun print_failed_mode options thy modes p (pol, m) rs is =
33130
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1115
  if show_mode_inference options then
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1116
    let
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1117
      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
  1118
        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
  1119
    in () end
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1120
  else ()
7eac458c2b22 added option show_mode_inference; added splitting of conjunctions in expand_tuples
bulwahn
parents: 33129
diff changeset
  1121
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1122
fun error_of p (pol, m) is =
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
  1123
  "  Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
  1124
        p ^ " violates mode " ^ string_of_mode m
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1125
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 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
  1127
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1128
    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
  1129
      | 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
  1130
      | 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
  1131
      | 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
  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
    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
  1134
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1135
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1136
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
  1137
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1138
    val (Ts, U) = strip_type T
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
  1139
    fun input_of (Type (@{type_name Product_Type.prod}, [T1, T2])) = Pair (input_of T1, input_of T2)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1140
      | 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
  1141
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1142
    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
  1143
      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
  1144
    else
35885
7b39120a1494 some improvements thanks to Makarius source code review
bulwahn
parents: 35884
diff changeset
  1145
      raise Fail "all_input_of: not a predicate"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1146
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1147
39191
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1148
fun find_least ord xs =
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1149
  let
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1150
    fun find' x y = (case y of NONE => SOME x | SOME y' => if ord (x, y') = LESS then SOME x else y) 
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1151
  in
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1152
    fold find' xs NONE
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1153
  end
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1154
  
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1155
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
  1156
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
  1157
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1158
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
  1159
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1160
    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
  1161
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1162
    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
  1163
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1164
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1165
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
  1166
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1167
    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
  1168
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1169
    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
  1170
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1171
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1172
fun is_invertible_function ctxt (Const (f, _)) = is_constr ctxt f
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1173
  | is_invertible_function ctxt _ = false
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1174
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1175
fun non_invertible_subterms ctxt (t as Free _) = []
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1176
  | non_invertible_subterms ctxt t = 
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1177
  let
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1178
    val (f, args) = strip_comb t
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1179
  in
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1180
    if is_invertible_function ctxt f then
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1181
      maps (non_invertible_subterms ctxt) 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
  1182
    else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1183
      [t]
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1184
  end
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1185
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1186
fun collect_non_invertible_subterms ctxt (f as Free _) (names, eqs) = (f, (names, eqs))
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1187
  | collect_non_invertible_subterms ctxt t (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
  1188
    case (strip_comb t) of (f, args) =>
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1189
      if is_invertible_function ctxt 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
  1190
          let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1191
            val (args', (names', eqs')) =
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1192
              fold_map (collect_non_invertible_subterms ctxt) args (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
  1193
          in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1194
            (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
  1195
          end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1196
        else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1197
          let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1198
            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
  1199
            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
  1200
          in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1201
            (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
  1202
          end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1203
(*
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 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
  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 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
  1207
      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
  1208
    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
  1209
*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1210
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1211
fun is_possible_output ctxt vs 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
  1212
  forall
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1213
    (fn t => is_eqT (fastype_of t) andalso forall (member (op =) vs) (term_vs t))
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1214
      (non_invertible_subterms ctxt t)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1215
  andalso
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1216
    (forall (is_eqT o snd)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1217
      (inter (fn ((f', _), f) => f = f') vs (Term.add_frees t [])))
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1218
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1219
fun vars_of_destructable_term ctxt (Free (x, _)) = [x]
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1220
  | vars_of_destructable_term ctxt t =
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1221
  let
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1222
    val (f, args) = strip_comb t
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1223
  in
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1224
    if is_invertible_function ctxt f then
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1225
      maps (vars_of_destructable_term ctxt) 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
  1226
    else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1227
      []
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1228
  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
  1229
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1230
fun is_constructable vs t = forall (member (op =) vs) (term_vs 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
  1231
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1232
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
  1233
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37146
diff changeset
  1234
fun output_terms (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1235
    output_terms (t1, d1)  @ output_terms (t2, d2)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1236
  | output_terms (t1 $ t2, Mode_App (d1, d2)) =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1237
    output_terms (t1, d1)  @ output_terms (t2, d2)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1238
  | output_terms (t, Term Output) = [t]
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1239
  | output_terms _ = []
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1240
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1241
fun lookup_mode modes (Const (s, T)) =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1242
   (case (AList.lookup (op =) modes s) of
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1243
      SOME ms => SOME (map (fn m => (Context m, [])) ms)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1244
    | NONE => NONE)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1245
  | lookup_mode modes (Free (x, _)) =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1246
    (case (AList.lookup (op =) modes x) of
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1247
      SOME ms => SOME (map (fn m => (Context m , [])) ms)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1248
    | NONE => NONE)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1249
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37146
diff changeset
  1250
fun derivations_of (ctxt : Proof.context) modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) (Pair (m1, m2)) =
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
    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
  1252
      (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1253
        (derivations_of ctxt modes vs t1 m1) (derivations_of ctxt modes vs t2 m2)
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1254
  | derivations_of ctxt modes vs t (m as Fun _) =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1255
    (*let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1256
      val (p, args) = strip_comb t
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1257
    in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1258
      (case lookup_mode modes p of
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1259
        SOME ms => map_filter (fn (Context m, []) => let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1260
          val ms = strip_fun_mode m
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1261
          val (argms, restms) = chop (length args) ms
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1262
          val m' = fold_rev (curry Fun) restms Bool
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1263
        in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1264
          if forall (fn m => eq_mode (Input, m)) argms andalso eq_mode (m', mode) then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1265
            SOME (fold (curry Mode_App) (map Term argms) (Context m), missing_vars vs t)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1266
          else NONE
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1267
        end) ms
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1268
      | NONE => (if is_all_input mode then [(Context mode, [])] else []))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1269
    end*)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1270
    (case try (all_derivations_of ctxt modes vs) t  of
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1271
      SOME derivs =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1272
        filter (fn (d, mvars) => eq_mode (mode_of d, m) andalso null (output_terms (t, d))) derivs
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1273
    | NONE => (if is_all_input m then [(Context m, [])] else []))
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1274
  | derivations_of ctxt modes vs t m =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1275
    if eq_mode (m, Input) then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1276
      [(Term Input, missing_vars vs t)]
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1277
    else if eq_mode (m, Output) then
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1278
      (if is_possible_output ctxt vs t then [(Term Output, [])] else [])
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1279
    else []
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37146
diff changeset
  1280
and all_derivations_of ctxt modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1281
  let
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1282
    val derivs1 = all_derivations_of ctxt modes vs t1
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1283
    val derivs2 = all_derivations_of ctxt modes vs t2
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1284
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1285
    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
  1286
      (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
  1287
        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
  1288
  end
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1289
  | all_derivations_of ctxt modes vs (t1 $ t2) =
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  1290
  let
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1291
    val derivs1 = all_derivations_of ctxt modes vs t1
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1292
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1293
    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
  1294
      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
  1295
        Fun (m', _) => map (fn (d2, mvars2) =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1296
          (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of ctxt modes vs t2 m')
39193
0e505a4e500c stating errors in error messages more verbose in predicate compiler
bulwahn
parents: 39191
diff changeset
  1297
        | _ => raise Fail "all_derivations_of: derivation has an unexpected non-functional mode") derivs1
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1298
  end
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1299
  | all_derivations_of _ modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1300
  | all_derivations_of _ modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
39193
0e505a4e500c stating errors in error messages more verbose in predicate compiler
bulwahn
parents: 39191
diff changeset
  1301
  | all_derivations_of _ modes vs _ = raise Fail "all_derivations_of: unexpected term"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1302
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1303
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
  1304
  | 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
  1305
  | 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
  1306
  | 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
  1307
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1308
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
  1309
  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
  1310
    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
  1311
      (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
  1312
        SOME ms =>
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1313
          (case AList.lookup (op =) (map (fn ((p, m), r) => (m, r)) ms) (head_mode_of deriv) 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
  1314
            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
  1315
          | 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
  1316
      | 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
  1317
  | 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
  1318
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1319
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
  1320
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1321
    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
  1322
    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
  1323
      | 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
  1324
      | 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
  1325
      | 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
  1326
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1327
    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
  1328
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1329
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1330
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
  1331
  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
  1332
    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
  1333
  | 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
  1334
36247
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1335
fun lexl_ord [] (x, x') = EQUAL
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1336
  | lexl_ord (ord :: ords') (x, x') =
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1337
    case ord (x, x') of
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1338
      EQUAL => lexl_ord ords' (x, x')
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1339
    | ord => ord
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1340
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1341
fun deriv_ord' ctxt pol pred modes t1 t2 ((deriv1, mvars1), (deriv2, mvars2)) =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1342
  let
36247
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1343
    (* prefer functional modes if it is a function *)
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1344
    fun fun_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1345
      let
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1346
        fun is_functional t mode =
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1347
          case try (fst o dest_Const o fst o strip_comb) t of
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1348
            NONE => false
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1349
          | SOME c => is_some (alternative_compilation_of ctxt c mode)
36247
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1350
      in
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1351
        case (is_functional t1 (head_mode_of deriv1), is_functional t2 (head_mode_of deriv2)) of
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1352
          (true, true) => EQUAL
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1353
        | (true, false) => LESS
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1354
        | (false, true) => GREATER
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1355
        | (false, false) => EQUAL
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1356
      end
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1357
    (* prefer modes without requirement for generating random values *)
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
    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
  1359
      int_ord (length mvars1, length mvars2)
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1360
    (* prefer non-random 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
  1361
    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
  1362
      int_ord (if random_mode_in_deriv modes t1 deriv1 then 1 else 0,
39760
e6a370d35f45 fixed a typo that caused the preference of non-random modes to be ignored
bulwahn
parents: 39685
diff changeset
  1363
               if random_mode_in_deriv modes t2 deriv2 then 1 else 0)
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1364
    (* prefer modes with more input and less output *)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1365
    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
  1366
      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
  1367
        number_of_output_positions (head_mode_of deriv2))
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1368
    (* prefer recursive calls *)
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1369
    fun is_rec_premise t =
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1370
      case fst (strip_comb t) of Const (c, T) => c = pred | _ => false
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1371
    fun recursive_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1372
      int_ord (if is_rec_premise t1 then 0 else 1,
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1373
        if is_rec_premise t2 then 0 else 1)
36247
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1374
    val ord = lexl_ord [mvars_ord, fun_ord, random_mode_ord, output_mode_ord, recursive_ord]
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1375
  in
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1376
    ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1377
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1378
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1379
fun deriv_ord ctxt pol pred modes t = deriv_ord' ctxt pol pred modes 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
  1380
36247
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1381
fun premise_ord thy pol pred modes ((prem1, a1), (prem2, a2)) =
36258
f459a0cc3241 removing dead code; clarifying function names; removing clone
bulwahn
parents: 36254
diff changeset
  1382
  rev_option_ord (deriv_ord' thy pol pred modes (dest_indprem prem1) (dest_indprem prem2)) (a1, a2)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1383
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1384
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
  1385
  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
  1386
    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
  1387
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1388
fun select_mode_prem (mode_analysis_options : mode_analysis_options) (ctxt : Proof.context) pred
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1389
  pol (modes, (pos_modes, neg_modes)) vs 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
  1390
  let
39191
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1391
    fun choose_mode_of_prem (Prem t) =
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1392
          find_least (deriv_ord ctxt pol pred modes t) (all_derivations_of ctxt pos_modes vs t)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1393
      | choose_mode_of_prem (Sidecond t) = SOME (Context Bool, missing_vars vs t)
39191
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1394
      | choose_mode_of_prem (Negprem t) = find_least (deriv_ord ctxt (not pol) pred modes t)
36247
bcf23027bca2 prefer functional modes of functions in the mode analysis
bulwahn
parents: 36056
diff changeset
  1395
          (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
39191
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1396
             (all_derivations_of ctxt neg_modes vs t))
39193
0e505a4e500c stating errors in error messages more verbose in predicate compiler
bulwahn
parents: 39191
diff changeset
  1397
      | choose_mode_of_prem p = raise Fail ("choose_mode_of_prem: unexpected premise " ^ string_of_prem ctxt p)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1398
  in
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1399
    if #reorder_premises mode_analysis_options then
39191
edaf5a6ffa99 using linear find_least instead of sorting in the mode analysis of the predicate compiler
bulwahn
parents: 39125
diff changeset
  1400
      find_least (premise_ord ctxt pol pred modes) (ps ~~ map choose_mode_of_prem ps)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1401
    else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1402
      SOME (hd ps, choose_mode_of_prem (hd 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
  1403
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1404
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1405
fun check_mode_clause' (mode_analysis_options : mode_analysis_options) ctxt pred param_vs (modes :
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1406
  (string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, 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
  1407
  let
36258
f459a0cc3241 removing dead code; clarifying function names; removing clone
bulwahn
parents: 36254
diff changeset
  1408
    val vTs = distinct (op =) (fold Term.add_frees (map dest_indprem ps) (fold Term.add_frees ts []))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1409
    val modes' = modes @ (param_vs ~~ map (fn x => [((true, x), false), ((false, x), false)]) (ho_arg_modes_of mode))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1410
    fun retrieve_modes_of_pol pol = map (fn (s, ms) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1411
      (s, map_filter (fn ((p, m), r) => if p = pol then SOME m else NONE | _ => NONE) ms))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1412
    val (pos_modes', neg_modes') =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1413
      if #infer_pos_and_neg_modes mode_analysis_options then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1414
        (retrieve_modes_of_pol pol modes', retrieve_modes_of_pol (not pol) modes')
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1415
      else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1416
        let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1417
          val modes = map (fn (s, ms) => (s, map (fn ((p, m), r) => m) ms)) modes'
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1418
        in (modes, modes) end
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1419
    val (in_ts, out_ts) = split_mode mode ts
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1420
    val in_vs = union (op =) param_vs (maps (vars_of_destructable_term ctxt) in_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
  1421
    val out_vs = terms_vs out_ts
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1422
    fun known_vs_after p vs = (case p of
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1423
        Prem t => union (op =) vs (term_vs t)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1424
      | Sidecond t => union (op =) vs (term_vs t)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1425
      | Negprem t => union (op =) vs (term_vs t)
39193
0e505a4e500c stating errors in error messages more verbose in predicate compiler
bulwahn
parents: 39191
diff changeset
  1426
      | _ => raise Fail "unexpected premise")
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1427
    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
  1428
      | check_mode_prems acc_ps rnd vs ps =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1429
        (case
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1430
          (select_mode_prem mode_analysis_options ctxt pred pol (modes', (pos_modes', neg_modes')) vs ps) of
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1431
          SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1432
            (known_vs_after p vs) (filter_out (equal p) 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
  1433
        | SOME (p, SOME (deriv, missing_vars)) =>
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  1434
          if #use_generators mode_analysis_options andalso pol 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
  1435
            check_mode_prems ((p, deriv) :: (map
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1436
              (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1437
                (distinct (op =) missing_vars))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1438
                @ acc_ps) true (known_vs_after p vs) (filter_out (equal p) 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
  1439
          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
  1440
        | 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
  1441
        | 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
  1442
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1443
    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
  1444
      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
  1445
    | SOME (acc_ps, vs, rnd) =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1446
      if forall (is_constructable vs) (in_ts @ out_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
  1447
        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
  1448
      else
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  1449
        if #use_generators mode_analysis_options andalso pol 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
  1450
          let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1451
             val generators = 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
  1452
              (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1453
                (subtract (op =) vs (terms_vs (in_ts @ 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
  1454
          in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1455
            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
  1456
          end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1457
        else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1458
          NONE
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1459
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1460
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1461
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
  1462
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1463
fun check_modes_pred' mode_analysis_options options thy param_vs clauses modes (p, (ms : ((bool * mode) * bool) 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
  1464
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1465
    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
  1466
      let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1467
        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
  1468
          | split' ((m, Error z) :: xs) (ys, zs) = split' xs (ys, z :: zs)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1469
          | split' (((m : bool * mode), Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1470
       in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1471
         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
  1472
       end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1473
    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
  1474
    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
  1475
      let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1476
        val res = Output.cond_timeit false "work part of check_mode for one mode" (fn _ => 
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1477
          map (check_mode_clause' mode_analysis_options thy p param_vs modes m) rs)
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
      in
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1479
        Output.cond_timeit false "aux part of check_mode for one mode" (fn _ => 
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1480
        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
  1481
          [] => Success (exists (fn SOME (_, _, true) => true | _ => false) res)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1482
        | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is)))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1483
      end
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1484
    val _ = if show_mode_inference options then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1485
        tracing ("checking " ^ string_of_int (length ms) ^ " modes ...")
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1486
      else ()
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1487
    val res = Output.cond_timeit false "check_mode" (fn _ => map (fn (m, _) => (m, check_mode m)) ms)
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
    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
  1489
  in
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1490
    ((p, (ms' : ((bool * mode) * bool) list)), errors)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1491
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1492
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1493
fun get_modes_pred' mode_analysis_options thy param_vs clauses modes (p, ms) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1494
  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
  1495
    val rs = these (AList.lookup (op =) clauses p)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1496
  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
  1497
    (p, map (fn (m, rnd) =>
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1498
      (m, map
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1499
        ((fn (ts, ps, rnd) => (ts, ps)) o the o
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1500
          check_mode_clause' mode_analysis_options thy p param_vs modes m) rs)) ms)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1501
  end;
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33135
diff changeset
  1502
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1503
fun fixp f (x : (string * ((bool * mode) * bool) list) list) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1504
  let val y = f x
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1505
  in if x = y then x else fixp f y end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1506
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1507
fun fixp_with_state f (x : (string * ((bool * mode) * bool) list) list, state) =
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1508
  let
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1509
    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
  1510
  in
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1511
    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
  1512
  end
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1513
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1514
fun string_of_ext_mode ((pol, mode), rnd) =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1515
  string_of_mode mode ^ "(" ^ (if pol then "pos" else "neg") ^ ", "
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1516
  ^ (if rnd then "rnd" else "nornd") ^ ")"
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1517
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1518
fun print_extra_modes options modes =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1519
  if show_mode_inference options then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1520
    tracing ("Modes of inferred predicates: " ^
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1521
      cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_ext_mode ms)) modes))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1522
  else ()
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1523
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1524
fun infer_modes mode_analysis_options options (lookup_mode, lookup_neg_mode, needs_random) ctxt
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1525
  preds all_modes param_vs clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1526
  let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1527
    fun appair f (x1, x2) (y1, y2) = (f x1 y1, f x2 y2)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1528
    fun add_needs_random s (false, m) = ((false, m), false)
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1529
      | add_needs_random s (true, m) = ((true, m), needs_random s m)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1530
    fun add_polarity_and_random_bit s b ms = map (fn m => add_needs_random s (b, m)) ms
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1531
    val prednames = map fst preds
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1532
    (* extramodes contains all modes of all constants, should we only use the necessary ones
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1533
       - what is the impact on performance? *)
36251
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1534
    fun predname_of (Prem t) =
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1535
        (case try dest_Const (fst (strip_comb t)) of SOME (c, _) => insert (op =) c | NONE => I)
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1536
      | predname_of (Negprem t) =
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1537
        (case try dest_Const (fst (strip_comb t)) of SOME (c, _) => insert (op =) c | NONE => I)
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1538
      | predname_of _ = I
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1539
    val relevant_prednames = fold (fn (_, clauses') =>
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1540
      fold (fn (_, ps) => fold Term.add_const_names (map dest_indprem ps)) clauses') clauses []
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1541
      |> filter_out (fn name => member (op =) prednames name)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1542
    val extra_modes =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1543
      if #infer_pos_and_neg_modes mode_analysis_options then
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  1544
        let
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1545
          val pos_extra_modes =
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1546
            map_filter (fn name => Option.map (pair name) (try lookup_mode name))
36251
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  1547
            relevant_prednames
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1548
          val neg_extra_modes =
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1549
            map_filter (fn name => Option.map (pair name) (try lookup_neg_mode name))
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1550
              relevant_prednames
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1551
        in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1552
          map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1553
                @ add_polarity_and_random_bit s false (the (AList.lookup (op =) neg_extra_modes s))))
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1554
            pos_extra_modes
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1555
        end
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1556
      else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1557
        map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)))
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1558
          (map_filter (fn name => Option.map (pair name) (try lookup_mode name))
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1559
            relevant_prednames)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1560
    val _ = print_extra_modes options extra_modes
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1561
    val start_modes =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1562
      if #infer_pos_and_neg_modes mode_analysis_options then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1563
        map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms @
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1564
          (map (fn m => ((false, m), false)) ms))) all_modes
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1565
      else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1566
        map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms)) all_modes
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1567
    fun iteration modes = map
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1568
      (check_modes_pred' mode_analysis_options options ctxt param_vs clauses
36028
3837493fe4ab prefer recursive calls before others in the mode inference
bulwahn
parents: 36027
diff changeset
  1569
        (modes @ extra_modes)) modes
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1570
    val ((modes : (string * ((bool * mode) * bool) list) list), errors) =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1571
      Output.cond_timeit false "Fixpount computation of mode analysis" (fn () =>
39383
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
  1572
      if show_invalid_clauses options then
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1573
        fixp_with_state (fn (modes, errors) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1574
          let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1575
            val (modes', new_errors) = split_list (iteration modes)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1576
          in (modes', errors @ flat new_errors) end) (start_modes, [])
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1577
        else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1578
          (fixp (fn modes => map fst (iteration modes)) start_modes, []))
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1579
    val moded_clauses = map (get_modes_pred' mode_analysis_options ctxt param_vs clauses
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1580
      (modes @ extra_modes)) modes
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1581
    val need_random = fold (fn (s, ms) => if member (op =) (map fst preds) s then
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1582
      cons (s, (map_filter (fn ((true, m), true) => SOME m | _ => NONE) ms)) else I)
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1583
      modes []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1584
  in
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  1585
    ((moded_clauses, need_random), errors)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1586
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1587
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1588
(* term construction *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1589
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1590
fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1591
      NONE => (Free (s, T), (names, (s, [])::vs))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1592
    | SOME xs =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1593
        let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1594
          val s' = Name.variant names s;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1595
          val v = Free (s', T)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1596
        in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1597
          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1598
        end);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1599
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1600
fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1601
  | distinct_v (t $ u) nvs =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1602
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1603
        val (t', nvs') = distinct_v t nvs;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1604
        val (u', nvs'') = distinct_v u nvs';
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1605
      in (t' $ u', nvs'') end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1606
  | distinct_v x nvs = (x, nvs);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1607
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1608
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1609
(** specific rpred functions -- move them to the correct place in this file *)
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1610
fun mk_Eval_of (P as (Free (f, _)), T) mode =
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1611
let
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1612
  fun mk_bounds (Type (@{type_name Product_Type.prod}, [T1, T2])) i =
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1613
    let
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1614
      val (bs2, i') = mk_bounds T2 i 
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1615
      val (bs1, i'') = mk_bounds T1 i'
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1616
    in
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1617
      (HOLogic.pair_const T1 T2 $ bs1 $ bs2, i'' + 1)
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1618
    end
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1619
    | mk_bounds T i = (Bound i, i + 1)
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1620
  fun mk_prod ((t1, T1), (t2, T2)) = (HOLogic.pair_const T1 T2 $ t1 $ t2, HOLogic.mk_prodT (T1, T2))
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1621
  fun mk_tuple [] = (HOLogic.unit, HOLogic.unitT)
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1622
    | mk_tuple tTs = foldr1 mk_prod tTs;
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1623
  fun mk_split_abs (T as Type (@{type_name Product_Type.prod}, [T1, T2])) t = absdummy (T, HOLogic.split_const (T1, T2, @{typ bool}) $ (mk_split_abs T1 (mk_split_abs T2 t)))
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1624
    | mk_split_abs T t = absdummy (T, t)
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1625
  val args = rev (fst (fold_map mk_bounds (rev (binder_types T)) 0))
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1626
  val (inargs, outargs) = split_mode mode args
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1627
  val (inTs, outTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE)) mode (binder_types T)
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1628
  val inner_term = PredicateCompFuns.mk_Eval (list_comb (P, inargs), fst (mk_tuple (outargs ~~ outTs)))
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1629
in
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1630
  fold_rev mk_split_abs (binder_types T) inner_term  
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1631
end
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1632
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1633
fun compile_arg compilation_modifiers additional_arguments ctxt param_modes arg = 
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1634
  let
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1635
    fun map_params (t as Free (f, T)) =
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1636
      (case (AList.lookup (op =) param_modes f) of
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1637
          SOME mode =>
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1638
            let
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1639
              val T' = Comp_Mod.funT_of compilation_modifiers mode T
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1640
            in
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1641
              mk_Eval_of (Free (f, T'), T) mode
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1642
            end
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1643
        | NONE => t)
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1644
      | map_params t = t
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1645
  in
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1646
    map_aterms map_params arg
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1647
  end
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1648
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1649
fun compile_match compilation_modifiers additional_arguments ctxt param_modes
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1650
      eqs eqs' out_ts success_t =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1651
  let
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1652
    val compfuns = Comp_Mod.compfuns compilation_modifiers
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1653
    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
  1654
    val eqs'' =
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1655
      map (compile_arg compilation_modifiers additional_arguments ctxt param_modes) eqs''
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1656
    val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1657
    val name = Name.variant names "x";
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1658
    val name' = Name.variant (name :: names) "y";
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1659
    val T = HOLogic.mk_tupleT (map fastype_of out_ts);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1660
    val U = fastype_of success_t;
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1661
    val U' = dest_predT compfuns U;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1662
    val v = Free (name, T);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1663
    val v' = Free (name', T);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1664
  in
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1665
    lambda v (fst (Datatype.make_case ctxt Datatype_Case.Quiet [] v
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  1666
      [(HOLogic.mk_tuple out_ts,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1667
        if null eqs'' then success_t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1668
        else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1669
          foldr1 HOLogic.mk_conj eqs'' $ success_t $
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1670
            mk_bot compfuns U'),
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1671
       (v', mk_bot compfuns U')]))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1672
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1673
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1674
fun string_of_tderiv ctxt (t, deriv) = 
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1675
  (case (t, deriv) of
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1676
    (t1 $ t2, Mode_App (deriv1, deriv2)) =>
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1677
      string_of_tderiv ctxt (t1, deriv1) ^ " $ " ^ string_of_tderiv ctxt (t2, deriv2)
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37146
diff changeset
  1678
  | (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (deriv1, deriv2)) =>
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1679
    "(" ^ string_of_tderiv ctxt (t1, deriv1) ^ ", " ^ string_of_tderiv ctxt (t2, deriv2) ^ ")"
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1680
  | (t, Term Input) => Syntax.string_of_term ctxt t ^ "[Input]"
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1681
  | (t, Term Output) => Syntax.string_of_term ctxt t ^ "[Output]"
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1682
  | (t, Context m) => Syntax.string_of_term ctxt t ^ "[" ^ string_of_mode m ^ "]")
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  1683
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1684
fun compile_expr compilation_modifiers ctxt (t, deriv) param_modes additional_arguments =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1685
  let
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1686
    val compfuns = Comp_Mod.compfuns compilation_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
  1687
    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
  1688
      (case (t, deriv) of
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1689
        (t, Term Input) => SOME (compile_arg compilation_modifiers additional_arguments ctxt param_modes 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
  1690
      | (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
  1691
      | (Const (name, T), Context mode) =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1692
        (case alternative_compilation_of ctxt name mode of
36038
385f706eff24 generalized alternative functions to alternative compilation to handle arithmetic functions better
bulwahn
parents: 36037
diff changeset
  1693
          SOME alt_comp => SOME (alt_comp compfuns T)
36034
ee84eac07290 added book-keeping, registration and compilation with alternative functions for predicates in the predicate compiler
bulwahn
parents: 36031
diff changeset
  1694
        | NONE =>
ee84eac07290 added book-keeping, registration and compilation with alternative functions for predicates in the predicate compiler
bulwahn
parents: 36031
diff changeset
  1695
          SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  1696
            ctxt name mode,
36034
ee84eac07290 added book-keeping, registration and compilation with alternative functions for predicates in the predicate compiler
bulwahn
parents: 36031
diff changeset
  1697
            Comp_Mod.funT_of compilation_modifiers mode 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
  1698
      | (Free (s, T), Context m) =>
39785
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1699
        (case (AList.lookup (op =) param_modes s) of
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1700
          SOME mode => SOME (Free (s, Comp_Mod.funT_of compilation_modifiers m T))
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1701
        | NONE =>
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1702
        let
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1703
          val bs = map (pair "x") (binder_types (fastype_of t))
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1704
          val bounds = map Bound (rev (0 upto (length bs) - 1))
05c4e9ecf5f6 improving the compilation to handle predicate arguments in higher-order argument positions
bulwahn
parents: 39783
diff changeset
  1705
        in SOME (list_abs (bs, mk_if compfuns (list_comb (t, bounds)))) 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
  1706
      | (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
  1707
        let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1708
          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
  1709
          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
  1710
        in SOME (list_abs (bs, mk_if compfuns (list_comb (t, bounds)))) end
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37146
diff changeset
  1711
      | (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =>
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1712
        (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
  1713
          (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
  1714
        | (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
  1715
        | (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
  1716
        | (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
  1717
      | (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
  1718
        (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
  1719
          (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
  1720
         | (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
  1721
         | _ => error "something went wrong here!"))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1722
  in
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  1723
    list_comb (the (expr_of (t, deriv)), 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
  1724
  end
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33144
diff changeset
  1725
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1726
fun compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1727
  inp (in_ts, out_ts) moded_ps =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1728
  let
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1729
    val compfuns = Comp_Mod.compfuns compilation_modifiers
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1730
    val compile_match = compile_match compilation_modifiers
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1731
      additional_arguments ctxt param_modes
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1732
    val (in_ts', (all_vs', eqs)) =
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1733
      fold_map (collect_non_invertible_subterms ctxt) in_ts (all_vs, []);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1734
    fun compile_prems out_ts' vs names [] =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1735
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1736
            val (out_ts'', (names', eqs')) =
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1737
              fold_map (collect_non_invertible_subterms ctxt) out_ts' (names, []);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1738
            val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1739
              out_ts'' (names', map (rpair []) vs);
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  1740
            val processed_out_ts = map (compile_arg compilation_modifiers additional_arguments
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  1741
              ctxt param_modes) out_ts
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1742
          in
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1743
            compile_match constr_vs (eqs @ eqs') out_ts'''
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  1744
              (mk_single compfuns (HOLogic.mk_tuple processed_out_ts))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1745
          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
  1746
      | compile_prems out_ts vs names ((p, deriv) :: ps) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1747
          let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1748
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1749
            val (out_ts', (names', eqs)) =
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  1750
              fold_map (collect_non_invertible_subterms ctxt) out_ts (names, [])
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1751
            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1752
              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
  1753
            val mode = head_mode_of deriv
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1754
            val additional_arguments' =
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  1755
              Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1756
            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
  1757
               Prem t =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1758
                 let
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1759
                   val u =
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1760
                     compile_expr compilation_modifiers ctxt (t, deriv) param_modes 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
  1761
                   val (_, out_ts''') = split_mode mode (snd (strip_comb t))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1762
                   val rest = compile_prems out_ts''' vs' names'' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1763
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1764
                   (u, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1765
                 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
  1766
             | Negprem t =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1767
                 let
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1768
                   val neg_compilation_modifiers =
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1769
                     negative_comp_modifiers_of compilation_modifiers
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  1770
                   val u = mk_not compfuns
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1771
                     (compile_expr neg_compilation_modifiers ctxt (t, deriv) param_modes 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
  1772
                   val (_, out_ts''') = split_mode mode (snd (strip_comb t))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1773
                   val rest = compile_prems out_ts''' vs' names'' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1774
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1775
                   (u, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1776
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1777
             | Sidecond t =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1778
                 let
36019
64bbbd858c39 generalizing the compilation process of the predicate compiler
bulwahn
parents: 36018
diff changeset
  1779
                   val t = compile_arg compilation_modifiers additional_arguments
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1780
                     ctxt param_modes t
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1781
                   val rest = compile_prems [] vs' names'' ps;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1782
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1783
                   (mk_if compfuns t, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1784
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1785
             | Generator (v, T) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1786
                 let
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  1787
                   val u = Comp_Mod.mk_random compilation_modifiers T additional_arguments
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1788
                   val rest = compile_prems [Free (v, T)]  vs' names'' ps;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1789
                 in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1790
                   (u, rest)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1791
                 end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1792
          in
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
  1793
            compile_match constr_vs' eqs out_ts''
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1794
              (mk_bind compfuns (compiled_clause, rest))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1795
          end
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1796
    val prem_t = compile_prems in_ts' (map fst param_modes) all_vs' moded_ps;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1797
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1798
    mk_bind compfuns (mk_single compfuns inp, prem_t)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1799
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1800
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1801
(* switch detection *)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1802
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1803
(** argument position of an inductive predicates and the executable functions **)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1804
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1805
type position = int * int list
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1806
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1807
fun input_positions_pair Input = [[]]
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1808
  | input_positions_pair Output = []
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1809
  | input_positions_pair (Fun _) = []
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1810
  | input_positions_pair (Pair (m1, m2)) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1811
    map (cons 1) (input_positions_pair m1) @ map (cons 2) (input_positions_pair m2)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1812
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1813
fun input_positions_of_mode mode = flat (map_index
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1814
   (fn (i, Input) => [(i, [])]
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1815
   | (_, Output) => []
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1816
   | (_, Fun _) => []
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1817
   | (i, m as Pair (m1, m2)) => map (pair i) (input_positions_pair m))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1818
     (Predicate_Compile_Aux.strip_fun_mode mode))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1819
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1820
fun argument_position_pair mode [] = []
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1821
  | argument_position_pair (Pair (Fun _, m2)) (2 :: is) = argument_position_pair m2 is
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1822
  | argument_position_pair (Pair (m1, m2)) (i :: is) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1823
    (if eq_mode (m1, Output) andalso i = 2 then
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1824
      argument_position_pair m2 is
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1825
    else if eq_mode (m2, Output) andalso i = 1 then
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1826
      argument_position_pair m1 is
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1827
    else (i :: argument_position_pair (if i = 1 then m1 else m2) is))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1828
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1829
fun argument_position_of mode (i, is) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1830
  (i - (length (filter (fn Output => true | Fun _ => true | _ => false)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1831
    (List.take (strip_fun_mode mode, i)))),
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1832
  argument_position_pair (nth (strip_fun_mode mode) i) is)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1833
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1834
fun nth_pair [] t = t
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1835
  | nth_pair (1 :: is) (Const (@{const_name Pair}, _) $ t1 $ _) = nth_pair is t1
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1836
  | nth_pair (2 :: is) (Const (@{const_name Pair}, _) $ _ $ t2) = nth_pair is t2
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1837
  | nth_pair _ _ = raise Fail "unexpected input for nth_tuple"
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1838
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1839
(** switch detection analysis **)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1840
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1841
fun find_switch_test ctxt (i, is) (ts, prems) =
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1842
  let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1843
    val t = nth_pair is (nth ts i)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1844
    val T = fastype_of t
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1845
  in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1846
    case T of
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1847
      TFree _ => NONE
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1848
    | Type (Tcon, _) =>
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1849
      (case Datatype_Data.get_constrs (ProofContext.theory_of ctxt) Tcon of
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1850
        NONE => NONE
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1851
      | SOME cs =>
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1852
        (case strip_comb t of
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1853
          (Var _, []) => NONE
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1854
        | (Free _, []) => NONE
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1855
        | (Const (c, T), _) => if AList.defined (op =) cs c then SOME (c, T) else NONE))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1856
  end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1857
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1858
fun partition_clause ctxt pos moded_clauses =
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1859
  let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1860
    fun insert_list eq (key, value) = AList.map_default eq (key, []) (cons value)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1861
    fun find_switch_test' moded_clause (cases, left) =
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1862
      case find_switch_test ctxt pos moded_clause of
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1863
        SOME (c, T) => (insert_list (op =) ((c, T), moded_clause) cases, left)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1864
      | NONE => (cases, moded_clause :: left)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1865
  in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1866
    fold find_switch_test' moded_clauses ([], [])
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1867
  end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1868
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1869
datatype switch_tree =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1870
  Atom of moded_clause list | Node of (position * ((string * typ) * switch_tree) list) * switch_tree
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1871
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1872
fun mk_switch_tree ctxt mode moded_clauses =
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1873
  let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1874
    fun select_best_switch moded_clauses input_position best_switch =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1875
      let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1876
        val ord = option_ord (rev_order o int_ord o (pairself (length o snd o snd)))
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1877
        val partition = partition_clause ctxt input_position moded_clauses
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1878
        val switch = if (length (fst partition) > 1) then SOME (input_position, partition) else NONE
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1879
      in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1880
        case ord (switch, best_switch) of LESS => best_switch
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1881
          | EQUAL => best_switch | GREATER => switch
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1882
      end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1883
    fun detect_switches moded_clauses =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1884
      case fold (select_best_switch moded_clauses) (input_positions_of_mode mode) NONE of
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1885
        SOME (best_pos, (switched_on, left_clauses)) =>
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1886
          Node ((best_pos, map (apsnd detect_switches) switched_on),
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1887
            detect_switches left_clauses)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1888
      | NONE => Atom moded_clauses
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1889
  in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1890
    detect_switches moded_clauses
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1891
  end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1892
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1893
(** compilation of detected switches **)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1894
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1895
fun destruct_constructor_pattern (pat, obj) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1896
  (case strip_comb pat of
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1897
    (f as Free _, []) => cons (pat, obj)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1898
  | (Const (c, T), pat_args) =>
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1899
    (case strip_comb obj of
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1900
      (Const (c', T'), obj_args) =>
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1901
        (if c = c' andalso T = T' then
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1902
          fold destruct_constructor_pattern (pat_args ~~ obj_args)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1903
        else raise Fail "pattern and object mismatch")
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1904
    | _ => raise Fail "unexpected object")
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1905
  | _ => raise Fail "unexpected pattern")
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1906
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1907
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1908
fun compile_switch compilation_modifiers ctxt all_vs param_modes additional_arguments mode
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1909
  in_ts' outTs switch_tree =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1910
  let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1911
    val compfuns = Comp_Mod.compfuns compilation_modifiers
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1912
    val thy = ProofContext.theory_of ctxt
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1913
    fun compile_switch_tree _ _ (Atom []) = NONE
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1914
      | compile_switch_tree all_vs ctxt_eqs (Atom moded_clauses) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1915
        let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1916
          val in_ts' = map (Pattern.rewrite_term thy ctxt_eqs []) in_ts'
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1917
          fun compile_clause' (ts, moded_ps) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1918
            let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1919
              val (ts, out_ts) = split_mode mode ts
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1920
              val subst = fold destruct_constructor_pattern (in_ts' ~~ ts) []
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1921
              val (fsubst, pat') = List.partition (fn (_, Free _) => true | _ => false) subst
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1922
              val moded_ps' = (map o apfst o map_indprem)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1923
                (Pattern.rewrite_term thy (map swap fsubst) []) moded_ps
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1924
              val inp = HOLogic.mk_tuple (map fst pat')
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1925
              val in_ts' = map (Pattern.rewrite_term thy (map swap fsubst) []) (map snd pat')
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1926
              val out_ts' = map (Pattern.rewrite_term thy (map swap fsubst) []) out_ts
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1927
            in
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1928
              compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1929
                inp (in_ts', out_ts') moded_ps'
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1930
            end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1931
        in SOME (foldr1 (mk_sup compfuns) (map compile_clause' moded_clauses)) end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1932
    | compile_switch_tree all_vs ctxt_eqs (Node ((position, switched_clauses), left_clauses)) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1933
      let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1934
        val (i, is) = argument_position_of mode position
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1935
        val inp_var = nth_pair is (nth in_ts' i)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1936
        val x = Name.variant all_vs "x"
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1937
        val xt = Free (x, fastype_of inp_var)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1938
        fun compile_single_case ((c, T), switched) =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1939
          let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1940
            val Ts = binder_types T
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1941
            val argnames = Name.variant_list (x :: all_vs)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1942
              (map (fn i => "c" ^ string_of_int i) (1 upto length Ts))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1943
            val args = map2 (curry Free) argnames Ts
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1944
            val pattern = list_comb (Const (c, T), args)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1945
            val ctxt_eqs' = (inp_var, pattern) :: ctxt_eqs
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1946
            val compilation = the_default (mk_bot compfuns (HOLogic.mk_tupleT outTs))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1947
              (compile_switch_tree (argnames @ x :: all_vs) ctxt_eqs' switched)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1948
        in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1949
          (pattern, compilation)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1950
        end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1951
        val switch = fst (Datatype.make_case ctxt Datatype_Case.Quiet [] inp_var
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1952
          ((map compile_single_case switched_clauses) @
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1953
            [(xt, mk_bot compfuns (HOLogic.mk_tupleT outTs))]))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1954
      in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1955
        case compile_switch_tree all_vs ctxt_eqs left_clauses of
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1956
          NONE => SOME switch
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1957
        | SOME left_comp => SOME (mk_sup compfuns (switch, left_comp))
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1958
      end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1959
  in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1960
    compile_switch_tree all_vs [] switch_tree
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1961
  end
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1962
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1963
(* compilation of predicates *)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1964
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  1965
fun compile_pred options compilation_modifiers ctxt all_vs param_vs s T (pol, mode) moded_cls =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  1966
  let
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  1967
    val compilation_modifiers = if pol then compilation_modifiers else
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  1968
      negative_comp_modifiers_of compilation_modifiers
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  1969
    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
  1970
      (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
  1971
    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
  1972
    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
  1973
      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
  1974
      | 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
  1975
    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
  1976
      (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
  1977
    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
  1978
    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
  1979
    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
  1980
      (fn T => fn (param_vs, names) =>
36018
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  1981
        if is_param_type T 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
  1982
          (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
  1983
        else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1984
          let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  1985
            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
  1986
          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
  1987
        (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
  1988
    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
  1989
      (fn t as Free (x, _) => if member (op =) param_vs x then NONE else SOME t | t => SOME t)) in_ts
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1990
    val param_modes = param_vs ~~ ho_arg_modes_of mode
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1991
    val compilation =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1992
      if detect_switches options then
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1993
        the_default (mk_bot compfuns (HOLogic.mk_tupleT outTs))
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1994
          (compile_switch compilation_modifiers ctxt all_vs param_modes additional_arguments mode
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  1995
            in_ts' outTs (mk_switch_tree ctxt mode moded_cls))
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1996
      else
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1997
        let
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1998
          val cl_ts =
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  1999
            map (fn (ts, moded_prems) => 
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  2000
              compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  2001
                (HOLogic.mk_tuple in_ts') (split_mode mode ts) moded_prems) moded_cls;
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2002
        in
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2003
          Comp_Mod.wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2004
            (if null cl_ts then
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2005
              mk_bot compfuns (HOLogic.mk_tupleT outTs)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2006
            else
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2007
              foldr1 (mk_sup compfuns) cl_ts)
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2008
        end
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2009
    val fun_const =
35891
3122bdd95275 contextifying the compilation of the predicate compiler
bulwahn
parents: 35889
diff changeset
  2010
      Const (function_name_of (Comp_Mod.compilation compilation_modifiers)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2011
      ctxt s mode, funT)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2012
  in
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2013
    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
  2014
      (HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2015
  end;
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2016
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2017
(** special setup for simpset **)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 34963
diff changeset
  2018
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
  2019
  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
  2020
  setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2021
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2022
(* Definition of executable functions and their intro and elim rules *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2023
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2024
fun print_arities arities = tracing ("Arities:\n" ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2025
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2026
    space_implode " -> " (map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2027
      (fn NONE => "X" | SOME k' => string_of_int k')
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2028
        (ks @ [SOME k]))) arities));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2029
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37558
diff changeset
  2030
fun strip_split_abs (Const (@{const_name prod_case}, _) $ t) = strip_split_abs 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
  2031
  | 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
  2032
  | 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
  2033
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
  2034
fun mk_args is_eval (m as Pair (m1, m2), T as Type (@{type_name Product_Type.prod}, [T1, T2])) names =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2035
    if eq_mode (m, Input) orelse eq_mode (m, Output) then
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2036
      let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2037
        val x = Name.variant names "x"
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2038
      in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2039
        (Free (x, T), x :: names)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2040
      end
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2041
    else
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2042
      let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2043
        val (t1, names') = mk_args is_eval (m1, T1) names
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2044
        val (t2, names'') = mk_args is_eval (m2, T2) names'
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2045
      in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2046
        (HOLogic.mk_prod (t1, t2), names'')
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2047
      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
  2048
  | 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
  2049
    let
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 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
  2051
      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
  2052
      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
  2053
      val (inargs, outargs) = split_map_mode (fn _ => fn t => (SOME t, NONE)) m args
39756
6c8e83d94536 consolidated tupled_lambda; moved to structure HOLogic
krauss
parents: 39685
diff changeset
  2054
      val t = fold_rev HOLogic.tupled_lambda args (PredicateCompFuns.mk_Eval
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2055
        (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
  2056
    in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2057
      (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
  2058
    end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2059
  | 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
  2060
    let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2061
      val x = Name.variant names "x"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  2062
    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
  2063
      (Free (x, T), x :: names)
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33265
diff changeset
  2064
    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
  2065
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2066
fun create_intro_elim_rule ctxt mode defthm mode_id funT 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
  2067
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2068
    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
  2069
    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
  2070
    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
  2071
    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
  2072
      let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2073
        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
  2074
        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
  2075
      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
  2076
    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
  2077
    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
  2078
    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
  2079
    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
  2080
      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
  2081
    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
  2082
    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
  2083
    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
  2084
    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
  2085
    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
  2086
    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
  2087
                    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
  2088
    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
  2089
                     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
  2090
    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
  2091
    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
  2092
      @{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
  2093
    val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2094
    val introthm = Goal.prove 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
  2095
      (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
  2096
    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
  2097
    val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2098
    val elimthm = Goal.prove 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
  2099
      (argnames @ ["y", "P"]) [] elimtrm (fn _ => unfolddef_tac)
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2100
    val opt_neg_introthm =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2101
      if is_all_input mode then
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2102
        let
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2103
          val neg_predpropI = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (pred, args')))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2104
          val neg_funpropI =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2105
            HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2106
              (PredicateCompFuns.mk_not (list_comb (funtrm, inargs)), HOLogic.unit))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2107
          val neg_introtrm = Logic.list_implies (neg_predpropI :: param_eqs, neg_funpropI)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2108
          val tac =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2109
            Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2110
              (@{thm if_False} :: @{thm Predicate.not_pred_eq} :: simprules)) 1
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2111
            THEN rtac @{thm Predicate.singleI} 1
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2112
        in SOME (Goal.prove ctxt (argnames @ hoarg_names') []
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2113
            neg_introtrm (fn _ => tac))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2114
        end
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2115
      else 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
  2116
  in
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2117
    ((introthm, elimthm), opt_neg_introthm)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2118
  end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2119
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2120
fun create_constname_of_mode options thy prefix name T mode = 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2121
  let
33626
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33623
diff changeset
  2122
    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
  2123
      ^ "_" ^ 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
  2124
    val name = the_default system_proposal (proposed_names options name mode)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2125
  in
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2126
    Sign.full_bname thy name
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2127
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2128
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2129
fun create_definitions options preds (name, modes) thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2130
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2131
    val compfuns = PredicateCompFuns.compfuns
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2132
    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
  2133
    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
  2134
      let
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  2135
        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
  2136
        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
  2137
        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
  2138
        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
  2139
        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
  2140
          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
  2141
            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
  2142
            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
  2143
          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
  2144
        val (inargs, outargs) = split_map_mode strip_eval mode args
39756
6c8e83d94536 consolidated tupled_lambda; moved to structure HOLogic
krauss
parents: 39685
diff changeset
  2145
        val predterm = fold_rev HOLogic.tupled_lambda inargs
6c8e83d94536 consolidated tupled_lambda; moved to structure HOLogic
krauss
parents: 39685
diff changeset
  2146
          (PredicateCompFuns.mk_Enum (HOLogic.tupled_lambda (HOLogic.mk_tuple outargs)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2147
            (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
  2148
        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
  2149
        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
  2150
        val ([definition], thy') = thy |>
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  2151
          Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39545
diff changeset
  2152
          Global_Theory.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2153
        val ctxt' = ProofContext.init_global thy'
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2154
        val rules as ((intro, elim), _) =
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2155
          create_intro_elim_rule ctxt' mode definition mode_cname funT (Const (name, T))
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  2156
        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
  2157
          |> set_function_name Pred name mode mode_cname
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2158
          |> add_predfun_data name mode (definition, rules)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39545
diff changeset
  2159
          |> Global_Theory.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39545
diff changeset
  2160
          |> Global_Theory.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33671
diff changeset
  2161
          |> Theory.checkpoint
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2162
        end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2163
  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
  2164
    thy |> defined_function_of Pred name |> fold create_definition modes
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2165
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2166
33620
b6bf2dc5aed7 added interface of user proposals for names of generated constants
bulwahn
parents: 33619
diff changeset
  2167
fun define_functions comp_modifiers compfuns options preds (name, modes) thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2168
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2169
    val T = AList.lookup (op =) preds name |> the
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2170
    fun create_definition mode thy =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2171
      let
33485
0df4f6e46e5e made definition of functions generically for the different instances
bulwahn
parents: 33484
diff changeset
  2172
        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
  2173
        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
  2174
        val funT = Comp_Mod.funT_of comp_modifiers mode T
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2175
      in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2176
        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
  2177
        |> set_function_name (Comp_Mod.compilation comp_modifiers) name mode mode_cname
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2178
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2179
  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
  2180
    thy
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2181
    |> 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
  2182
    |> fold create_definition modes
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2183
  end;
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32671
diff changeset
  2184
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2185
(* Proving equivalence of term *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2186
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2187
fun is_Type (Type _) = true
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2188
  | is_Type _ = false
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2189
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2190
(* returns true if t is an application of an datatype constructor *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2191
(* which then consequently would be splitted *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2192
(* else false *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2193
fun is_constructor thy t =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2194
  if (is_Type (fastype_of t)) then
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2195
    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2196
      NONE => false
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2197
    | SOME info => (let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2198
      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2199
      val (c, _) = strip_comb t
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2200
      in (case c of
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
  2201
        Const (name, _) => member (op =) constr_consts name
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2202
        | _ => false) end))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2203
  else false
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2204
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2205
(* MAJOR FIXME:  prove_params should be simple
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2206
 - 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
  2207
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2208
fun prove_param options ctxt nargs t deriv =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2209
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2210
    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
  2211
    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
  2212
    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
  2213
    val ho_args = ho_args_of mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2214
    val f_tac = case f of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2215
      Const (name, T) => simp_tac (HOL_basic_ss addsimps 
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2216
         [@{thm eval_pred}, predfun_definition_of ctxt name mode,
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2217
         @{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2218
         @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2219
    | Free _ =>
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2220
      Subgoal.FOCUS_PREMS (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2221
        let
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2222
          val prems' = maps dest_conjunct_prem (take nargs prems)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2223
        in
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2224
          MetaSimplifier.rewrite_goal_tac
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2225
            (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2226
        end) ctxt 1
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2227
    | Abs _ => raise Fail "prove_param: No valid parameter term"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2228
  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
  2229
    REPEAT_DETERM (rtac @{thm ext} 1)
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2230
    THEN print_tac options "prove_param"
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2231
    THEN f_tac 
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2232
    THEN print_tac options "after prove_param"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2233
    THEN (REPEAT_DETERM (atac 1))
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2234
    THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2235
    THEN REPEAT_DETERM (rtac @{thm refl} 1)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2236
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2237
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2238
fun prove_expr options ctxt nargs (premposition : int) (t, deriv) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2239
  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
  2240
    (Const (name, T), args) =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2241
      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
  2242
        val mode = head_mode_of deriv
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2243
        val introrule = predfun_intro_of ctxt 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
  2244
        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
  2245
        val ho_args = ho_args_of mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2246
      in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2247
        print_tac options "before intro rule:"
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2248
        THEN rtac introrule 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2249
        THEN print_tac options "after intro rule"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2250
        (* for the right assumption in first position *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2251
        THEN rotate_tac premposition 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
  2252
        THEN atac 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2253
        THEN print_tac options "parameter goal"
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2254
        (* work with parameter arguments *)
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2255
        THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2256
        THEN (REPEAT_DETERM (atac 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2257
      end
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2258
  | (Free _, _) =>
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2259
    print_tac options "proving parameter call.."
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2260
    THEN Subgoal.FOCUS_PREMS (fn {context = ctxt, params, prems, asms, concl, schematics} =>
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2261
        let
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2262
          val param_prem = nth prems premposition
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2263
          val (param, _) = strip_comb (HOLogic.dest_Trueprop (prop_of param_prem))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2264
          val prems' = maps dest_conjunct_prem (take nargs prems)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2265
          fun param_rewrite prem =
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2266
            param = snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of prem)))
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2267
          val SOME rew_eq = find_first param_rewrite prems'
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2268
          val param_prem' = MetaSimplifier.rewrite_rule
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2269
            (map (fn th => th RS @{thm eq_reflection})
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2270
              [rew_eq RS @{thm sym}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}])
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2271
            param_prem
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2272
        in
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2273
          rtac param_prem' 1
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2274
        end) ctxt 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2275
    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
  2276
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2277
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
  2278
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2279
fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2280
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2281
fun check_format ctxt st =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2282
  let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2283
    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
  2284
    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
  2285
    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
  2286
    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
  2287
      | 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
  2288
      | 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
  2289
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2290
    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
  2291
      ((*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
  2292
    else
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2293
      ((*tracing "expression is not valid";*) Seq.empty) (* error "check_format: wrong format" *)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2294
  end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2295
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2296
fun prove_match options ctxt nargs 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
  2297
  let
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2298
    val thy = ProofContext.theory_of ctxt
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2299
    val eval_if_P =
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2300
      @{lemma "P ==> Predicate.eval x z ==> Predicate.eval (if P then x else y) z" by simp} 
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2301
    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
  2302
      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
  2303
        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
  2304
          ((fst o dest_Type o fastype_of) t)))
39194
c8406125193b handling collection of simprules as sets rather than as lists
bulwahn
parents: 39193
diff changeset
  2305
        in fold (union Thm.eq_thm) (case_rewrites :: map get_case_rewrite (snd (strip_comb t))) [] 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
  2306
      else []
39194
c8406125193b handling collection of simprules as sets rather than as lists
bulwahn
parents: 39193
diff changeset
  2307
    val simprules = insert Thm.eq_thm @{thm "unit.cases"} (insert Thm.eq_thm @{thm "prod.cases"}
c8406125193b handling collection of simprules as sets rather than as lists
bulwahn
parents: 39193
diff changeset
  2308
      (fold (union Thm.eq_thm) (map get_case_rewrite 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
  2309
  (* 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
  2310
  in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2311
     (* 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
  2312
    asm_full_simp_tac (HOL_basic_ss' addsimps simprules) 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2313
    THEN print_tac options "after prove_match:"
39767
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2314
    THEN (DETERM (TRY 
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2315
           (rtac eval_if_P 1
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2316
           THEN (SUBPROOF (fn {context = ctxt, params, prems, asms, concl, schematics} =>
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2317
             (REPEAT_DETERM (rtac @{thm conjI} 1
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2318
             THEN (SOLVED (asm_simp_tac HOL_basic_ss' 1))))
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2319
             THEN print_tac options "if condition to be solved:"
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2320
             THEN asm_simp_tac HOL_basic_ss' 1
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2321
             THEN TRY (
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2322
                let
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2323
                  val prems' = maps dest_conjunct_prem (take nargs prems)
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2324
                in
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2325
                  MetaSimplifier.rewrite_goal_tac
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2326
                    (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2327
                end
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2328
             THEN REPEAT_DETERM (rtac @{thm refl} 1))
327e463531e4 using SUBPROOF to ensure that rewriting is not done on the further subgoals of proof
bulwahn
parents: 39766
diff changeset
  2329
             THEN print_tac options "after if simp; in SUBPROOF") ctxt 1))))
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2330
    THEN print_tac options "after if simplification"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2331
  end;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2332
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2333
(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2334
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2335
fun prove_sidecond ctxt t =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2336
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2337
    fun preds_of t nameTs = case strip_comb t of 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2338
      (f as Const (name, T), args) =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2339
        if is_registered ctxt name then (name, T) :: nameTs
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2340
          else fold preds_of args nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2341
      | _ => nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2342
    val preds = preds_of t []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2343
    val defs = map
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2344
      (fn (pred, T) => predfun_definition_of ctxt 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
  2345
        (all_input_of T))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2346
        preds
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2347
  in 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2348
    simp_tac (HOL_basic_ss addsimps
39649
7186d338f2e1 removing duplicate rewrite rule from simpset in predicate compiler
bulwahn
parents: 39648
diff changeset
  2349
      (@{thms HOL.simp_thms} @ (@{thm eval_pred} :: defs))) 1 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2350
    (* need better control here! *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2351
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2352
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2353
fun prove_clause options ctxt nargs mode (_, clauses) (ts, moded_ps) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2354
  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
  2355
    val (in_ts, clause_out_ts) = split_mode mode ts;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2356
    fun prove_prems out_ts [] =
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2357
      (prove_match options ctxt nargs out_ts)
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2358
      THEN print_tac options "before simplifying assumptions"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2359
      THEN asm_full_simp_tac HOL_basic_ss' 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2360
      THEN print_tac options "before single intro rule"
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2361
      THEN Subgoal.FOCUS_PREMS
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2362
             (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2363
              let
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2364
                val prems' = maps dest_conjunct_prem (take nargs prems)
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2365
              in
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2366
                MetaSimplifier.rewrite_goal_tac
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2367
                  (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2368
              end) ctxt 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2369
      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
  2370
    | prove_prems out_ts ((p, deriv) :: ps) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2371
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2372
        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
  2373
        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
  2374
        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
  2375
          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
  2376
          THEN (case p of Prem t =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2377
            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
  2378
              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
  2379
              val (_, out_ts''') = split_mode mode us
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2380
              val rec_tac = prove_prems out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2381
            in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2382
              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
  2383
              (*THEN asm_simp_tac HOL_basic_ss 1*)
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2384
              THEN print_tac options "before prove_expr:"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2385
              THEN prove_expr options ctxt nargs premposition (t, deriv)
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2386
              THEN print_tac options "after prove_expr:"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2387
              THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2388
            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
  2389
          | Negprem t =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2390
            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
  2391
              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
  2392
              val (_, out_ts''') = split_mode mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2393
              val rec_tac = prove_prems out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2394
              val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2395
              val neg_intro_rule =
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2396
                Option.map (fn name =>
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2397
                  the (predfun_neg_intro_of ctxt name mode)) 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
  2398
              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
  2399
              val params = ho_args_of mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2400
            in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2401
              print_tac options "before prove_neg_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
  2402
              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
  2403
                [@{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
  2404
                 @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2405
              THEN (if (is_some name) then
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2406
                  print_tac options "before applying not introduction rule"
39766
66c1e526fd44 avoiding instable rotate_tac and using the nice Subgoal.FOCUS_PREMS instead
bulwahn
parents: 39764
diff changeset
  2407
                  THEN Subgoal.FOCUS_PREMS
66c1e526fd44 avoiding instable rotate_tac and using the nice Subgoal.FOCUS_PREMS instead
bulwahn
parents: 39764
diff changeset
  2408
                    (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
66c1e526fd44 avoiding instable rotate_tac and using the nice Subgoal.FOCUS_PREMS instead
bulwahn
parents: 39764
diff changeset
  2409
                      rtac (the neg_intro_rule) 1
66c1e526fd44 avoiding instable rotate_tac and using the nice Subgoal.FOCUS_PREMS instead
bulwahn
parents: 39764
diff changeset
  2410
                      THEN rtac (nth prems premposition) 1) ctxt 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2411
                  THEN print_tac options "after applying not introduction rule"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2412
                  THEN (EVERY (map2 (prove_param options ctxt nargs) params param_derivations))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2413
                  THEN (REPEAT_DETERM (atac 1))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2414
                else
35884
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2415
                  rtac @{thm not_predI'} 1
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2416
                  (* test: *)
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2417
                  THEN dtac @{thm sym} 1
362bfc2ca0ee adding proof procedure for cases rule with tuples; adding introduction rule for negated premises; improving proof procedure with negated premises
bulwahn
parents: 35881
diff changeset
  2418
                  THEN asm_full_simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2419
                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2420
              THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2421
            end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2422
          | 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
  2423
           rtac @{thm if_predI} 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2424
           THEN print_tac options "before sidecond:"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2425
           THEN prove_sidecond ctxt t
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2426
           THEN print_tac options "after sidecond:"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2427
           THEN prove_prems [] ps)
39762
02fb709ab3cb handling higher-order relations in output terms; improving proof procedure; added test case
bulwahn
parents: 39761
diff changeset
  2428
      in (prove_match options ctxt nargs out_ts)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2429
          THEN rest_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2430
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2431
    val prems_tac = prove_prems in_ts moded_ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2432
  in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2433
    print_tac options "Proving clause..."
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2434
    THEN rtac @{thm bindI} 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2435
    THEN rtac @{thm singleI} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2436
    THEN prems_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2437
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2438
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2439
fun select_sup 1 1 = []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2440
  | select_sup _ 1 = [rtac @{thm supI1}]
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2441
  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2442
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2443
fun prove_one_direction options ctxt clauses preds pred mode moded_clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2444
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2445
    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
  2446
    val nargs = length (binder_types T)
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2447
    val pred_case_rule = the_elim_of ctxt pred
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2448
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2449
    REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2450
    THEN print_tac options "before applying elim rule"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2451
    THEN etac (predfun_elim_of ctxt pred mode) 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2452
    THEN etac pred_case_rule 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2453
    THEN print_tac options "after applying elim rule"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2454
    THEN (EVERY (map
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2455
           (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2456
             (1 upto (length moded_clauses))))
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2457
    THEN (EVERY (map2 (prove_clause options ctxt nargs mode) clauses moded_clauses))
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2458
    THEN print_tac options "proved one direction"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2459
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2460
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2461
(** Proof in the other direction **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2462
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2463
fun prove_match2 options ctxt out_ts =
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2464
  let
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2465
    val thy = ProofContext.theory_of ctxt
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2466
    fun split_term_tac (Free _) = all_tac
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2467
      | split_term_tac t =
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2468
        if (is_constructor thy t) then
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2469
          let
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2470
            val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2471
            val num_of_constrs = length (#case_rewrites info)
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2472
            val (_, ts) = strip_comb t
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2473
          in
35889
c1f86c5d3827 avoiding fishing for split_asm rule in the predicate compiler
bulwahn
parents: 35888
diff changeset
  2474
            print_tac options ("Term " ^ (Syntax.string_of_term ctxt t) ^ 
c1f86c5d3827 avoiding fishing for split_asm rule in the predicate compiler
bulwahn
parents: 35888
diff changeset
  2475
              "splitting with rules \n" ^ Display.string_of_thm ctxt (#split_asm info))
c1f86c5d3827 avoiding fishing for split_asm rule in the predicate compiler
bulwahn
parents: 35888
diff changeset
  2476
            THEN TRY ((Splitter.split_asm_tac [#split_asm info] 1)
c1f86c5d3827 avoiding fishing for split_asm rule in the predicate compiler
bulwahn
parents: 35888
diff changeset
  2477
              THEN (print_tac options "after splitting with split_asm rules")
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2478
            (* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2479
              THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2480
              THEN (REPEAT_DETERM_N (num_of_constrs - 1)
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2481
                (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2482
            THEN (assert_tac (Max_number_of_subgoals 2))
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2483
            THEN (EVERY (map split_term_tac ts))
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2484
          end
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2485
      else all_tac
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2486
  in
33148
0808f7d0d0d7 removed tuple functions from the predicate compiler
bulwahn
parents: 33147
diff changeset
  2487
    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
  2488
    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
  2489
    THEN (etac @{thm botE} 2))))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2490
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2491
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2492
(* VERY LARGE SIMILIRATIY to function prove_param 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2493
-- join both functions
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2494
*)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2495
(* TODO: remove function *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2496
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2497
fun prove_param2 options ctxt t deriv =
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2498
  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
  2499
    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
  2500
    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
  2501
    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
  2502
    val ho_args = ho_args_of mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2503
    val f_tac = case f of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2504
        Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2505
           (@{thm eval_pred}::(predfun_definition_of ctxt name mode)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2506
           :: @{thm "Product_Type.split_conv"}::[])) 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2507
      | Free _ => all_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2508
      | _ => error "prove_param2: illegal parameter term"
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2509
  in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2510
    print_tac options "before simplification in prove_args:"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2511
    THEN f_tac
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2512
    THEN print_tac options "after simplification in prove_args"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2513
    THEN EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2514
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2515
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2516
fun prove_expr2 options ctxt (t, deriv) = 
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2517
  (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
  2518
      (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
  2519
        let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2520
          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
  2521
          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
  2522
          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
  2523
        in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2524
          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
  2525
          THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2526
          THEN print_tac options "prove_expr2-before"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2527
          THEN etac (predfun_elim_of ctxt name mode) 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2528
          THEN print_tac options "prove_expr2"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2529
          THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2530
          THEN print_tac options "finished prove_expr2"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2531
        end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2532
      | _ => 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
  2533
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2534
fun prove_sidecond2 options ctxt t = let
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2535
  fun preds_of t nameTs = case strip_comb t of 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2536
    (f as Const (name, T), args) =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2537
      if is_registered ctxt name then (name, T) :: nameTs
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2538
        else fold preds_of args nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2539
    | _ => nameTs
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2540
  val preds = preds_of t []
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2541
  val defs = map
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2542
    (fn (pred, T) => predfun_definition_of ctxt 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
  2543
      (all_input_of T))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2544
      preds
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2545
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2546
   (* only simplify the one assumption *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2547
   full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2548
   (* need better control here! *)
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2549
   THEN print_tac options "after sidecond2 simplification"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2550
   end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2551
  
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2552
fun prove_clause2 options ctxt pred mode (ts, ps) i =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2553
  let
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2554
    val pred_intro_rule = nth (intros_of ctxt 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
  2555
    val (in_ts, clause_out_ts) = split_mode mode ts;
37544
456dd03e8cce more precise tactic: do not escape to a different goal branch (REPEAT is still problematic, though)
haftmann
parents: 37391
diff changeset
  2556
    val split_ss = HOL_basic_ss' addsimps [@{thm split_eta}, @{thm split_beta},
456dd03e8cce more precise tactic: do not escape to a different goal branch (REPEAT is still problematic, though)
haftmann
parents: 37391
diff changeset
  2557
      @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}]
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2558
    fun prove_prems2 out_ts [] =
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2559
      print_tac options "before prove_match2 - last call:"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2560
      THEN prove_match2 options ctxt out_ts
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2561
      THEN print_tac options "after prove_match2 - last call:"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2562
      THEN (etac @{thm singleE} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2563
      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2564
      THEN (asm_full_simp_tac HOL_basic_ss' 1)
39783
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2565
      THEN TRY (
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2566
        (REPEAT_DETERM (etac @{thm Pair_inject} 1))
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2567
        THEN (asm_full_simp_tac HOL_basic_ss' 1)
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2568
        
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2569
        THEN SOLVED (print_tac options "state before applying intro rule:"
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2570
        THEN (rtac pred_intro_rule
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2571
        (* How to handle equality correctly? *)
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2572
        THEN_ALL_NEW (K (print_tac options "state before assumption matching")
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2573
        THEN' (atac ORELSE' ((CHANGED o asm_full_simp_tac split_ss) THEN' (TRY o atac)))
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2574
          THEN' (K (print_tac options "state after pre-simplification:"))
a8c52b982ff2 putting the last step of the proof in a TRY block to handle the case that the simplifier already solved the goal
bulwahn
parents: 39768
diff changeset
  2575
        THEN' (K (print_tac options "state after assumption matching:")))) 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
  2576
    | prove_prems2 out_ts ((p, deriv) :: ps) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2577
      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
  2578
        val mode = head_mode_of deriv
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2579
        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
  2580
          Prem t =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2581
          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
  2582
            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
  2583
            val (_, out_ts''') = split_mode mode us
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2584
            val rec_tac = prove_prems2 out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2585
          in
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2586
            (prove_expr2 options ctxt (t, deriv)) THEN rec_tac
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2587
          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
  2588
        | Negprem t =>
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2589
          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
  2590
            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
  2591
            val (_, out_ts''') = split_mode mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2592
            val rec_tac = prove_prems2 out_ts''' ps
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2593
            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
  2594
            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
  2595
            val ho_args = ho_args_of mode args
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2596
          in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2597
            print_tac options "before neg prem 2"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2598
            THEN etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2599
            THEN (if is_some name then
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2600
                full_simp_tac (HOL_basic_ss addsimps
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2601
                  [predfun_definition_of ctxt (the name) mode]) 1
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2602
                THEN etac @{thm not_predE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2603
                THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2604
                THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2605
              else
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2606
                etac @{thm not_predE'} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2607
            THEN rec_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2608
          end 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2609
        | Sidecond t =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2610
          etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2611
          THEN etac @{thm if_predE} 1
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2612
          THEN prove_sidecond2 options ctxt t
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2613
          THEN prove_prems2 [] ps)
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2614
      in print_tac options "before prove_match2:"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2615
         THEN prove_match2 options ctxt out_ts
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2616
         THEN print_tac options "after prove_match2:"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2617
         THEN rest_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2618
      end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2619
    val prems_tac = prove_prems2 in_ts ps 
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2620
  in
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2621
    print_tac options "starting prove_clause2"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2622
    THEN etac @{thm bindE} 1
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2623
    THEN (etac @{thm singleE'} 1)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2624
    THEN (TRY (etac @{thm Pair_inject} 1))
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2625
    THEN print_tac options "after singleE':"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2626
    THEN prems_tac
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2627
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2628
 
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2629
fun prove_other_direction options ctxt pred mode moded_clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2630
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2631
    fun prove_clause clause i =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2632
      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2633
      THEN (prove_clause2 options ctxt pred mode clause i)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2634
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2635
    (DETERM (TRY (rtac @{thm unit.induct} 1)))
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2636
     THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2637
     THEN (rtac (predfun_intro_of ctxt pred mode) 1)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2638
     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
  2639
     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
  2640
         etac @{thm botE} 1
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2641
       else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2642
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2643
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2644
(** proof procedure **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2645
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2646
fun prove_pred options thy clauses preds pred (pol, mode) (moded_clauses, compiled_term) =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2647
  let
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36533
diff changeset
  2648
    val ctxt = ProofContext.init_global thy
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2649
    val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2650
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2651
    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
  2652
      (if not (skip_proof options) then
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2653
        (fn _ =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2654
        rtac @{thm pred_iffI} 1
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2655
        THEN print_tac options "after pred_iffI"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2656
        THEN prove_one_direction options ctxt clauses preds pred mode moded_clauses
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2657
        THEN print_tac options "proved one direction"
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2658
        THEN prove_other_direction options ctxt pred mode moded_clauses
35886
f5422b736c44 reduced the debug output functions from 2 to 1
bulwahn
parents: 35885
diff changeset
  2659
        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
  2660
      else (fn _ => Skip_Proof.cheat_tac thy))
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2661
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2662
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2663
(* composition of mode inference, definition, compilation and proof *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2664
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2665
(** auxillary combinators for table of preds and modes **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2666
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2667
fun map_preds_modes f preds_modes_table =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2668
  map (fn (pred, modes) =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2669
    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2670
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2671
fun join_preds_modes table1 table2 =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2672
  map_preds_modes (fn pred => fn mode => fn value =>
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2673
    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2674
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2675
fun maps_modes preds_modes_table =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2676
  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
  2677
    (pred, map (fn (mode, value) => value) modes)) preds_modes_table
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2678
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  2679
fun compile_preds options comp_modifiers ctxt all_vs param_vs preds moded_clauses =
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  2680
  map_preds_modes (fn pred => compile_pred options comp_modifiers ctxt all_vs param_vs pred
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2681
      (the (AList.lookup (op =) preds pred))) moded_clauses
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2682
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2683
fun prove options thy clauses preds moded_clauses compiled_terms =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2684
  map_preds_modes (prove_pred options thy clauses preds)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2685
    (join_preds_modes moded_clauses compiled_terms)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2686
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2687
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
  2688
  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
  2689
    (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
  2690
    compiled_terms
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2691
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2692
(* preparation of introduction rules into special datastructures *)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2693
fun dest_prem ctxt params t =
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2694
  (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
  2695
    (v as Free _, ts) => if member (op =) params v then Prem t else Sidecond t
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2696
  | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem ctxt 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
  2697
      Prem t => Negprem t
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2698
    | Negprem _ => error ("Double negation not allowed in premise: " ^
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2699
        Syntax.string_of_term ctxt (c $ t)) 
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2700
    | Sidecond t => Sidecond (c $ t))
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2701
  | (c as Const (s, _), ts) =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2702
    if is_registered ctxt 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
  2703
  | _ => 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
  2704
39310
17ef4415b17c removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents: 39299
diff changeset
  2705
fun prepare_intrs options ctxt prednames intros =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2706
  let
39310
17ef4415b17c removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents: 39299
diff changeset
  2707
    val thy = ProofContext.theory_of ctxt
33126
bb8806eb5da7 importing of polymorphic introduction rules with different schematic variable names
bulwahn
parents: 33124
diff changeset
  2708
    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
  2709
    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
  2710
    val (preds, intrs) = unify_consts thy preds intrs
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2711
    val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] ctxt
33126
bb8806eb5da7 importing of polymorphic introduction rules with different schematic variable names
bulwahn
parents: 33124
diff changeset
  2712
    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
  2713
    val all_vs = terms_vs intrs
39201
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2714
    fun generate_modes s T =
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2715
      if member (op =) (no_higher_order_predicate options) s then
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2716
        all_smodes_of_typ T
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2717
      else
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2718
        all_modes_of_typ T
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2719
    val all_modes = 
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2720
      map (fn (s, T) =>
39382
c797f3ab2ae1 proposed modes for code_pred now supports modes for mutual predicates
bulwahn
parents: 39310
diff changeset
  2721
        (s, case proposed_modes options s of
39651
2e06dad03dd3 adding check if user-given modes match type of predicates; removed useless function expand_tuples_elim
bulwahn
parents: 39649
diff changeset
  2722
            SOME ms => check_matches_type ctxt s T ms
39201
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2723
          | NONE => generate_modes s T)) 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
  2724
    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
  2725
      case intrs of
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2726
        [] =>
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  2727
          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
  2728
            val T = snd (hd preds)
39764
1cf2088cf035 only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents: 39763
diff changeset
  2729
            val one_mode = hd (the (AList.lookup (op =) all_modes (fst (hd 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
  2730
            val paramTs =
39764
1cf2088cf035 only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents: 39763
diff changeset
  2731
              ho_argsT_of one_mode (binder_types T)
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2732
            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
  2733
              (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
  2734
          in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2735
            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
  2736
          end
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2737
      | (intr :: _) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2738
        let
39764
1cf2088cf035 only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents: 39763
diff changeset
  2739
          val (p, args) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))
1cf2088cf035 only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents: 39763
diff changeset
  2740
          val one_mode = hd (the (AList.lookup (op =) all_modes (fst (dest_Const p))))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2741
        in
39764
1cf2088cf035 only modes but not types are used to destruct terms and types; this allows to interpret arguments that are predicates simply as input
bulwahn
parents: 39763
diff changeset
  2742
          ho_args_of one_mode args
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2743
        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
  2744
    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
  2745
    fun add_clause intr clauses =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2746
      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
  2747
        val (Const (name, T), ts) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2748
        val prems = map (dest_prem ctxt params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2749
      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
  2750
        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
  2751
          [(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
  2752
      end;
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2753
    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
  2754
  in
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2755
    (preds, all_vs, param_vs, all_modes, clauses)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2756
  end;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2757
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2758
(* 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
  2759
(* 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
  2760
(*
33106
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2761
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
  2762
  let
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2763
    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
  2764
    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
  2765
    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
  2766
    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
  2767
      (Ts as _ :: _ :: _) =>
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2768
        if length (HOLogic.strip_tuple arg) = length Ts then
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2769
          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
  2770
        else
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2771
          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
  2772
          ^ (Syntax.string_of_term_global thy arg) ^ " in " ^
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2773
          (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
  2774
      | _ => true
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2775
    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
  2776
    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
  2777
      | 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
  2778
      | check_prem _ = true
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2779
  in
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2780
    forall check_arg args andalso
7a1636c3ffc9 extended core of predicate compiler to expand tuples in introduction rules
bulwahn
parents: 32740
diff changeset
  2781
    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
  2782
  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
  2783
*)
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2784
(*
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2785
fun check_intros_elim_match thy prednames =
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2786
  let
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2787
    fun check predname =
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2788
      let
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2789
        val intros = intros_of thy predname
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2790
        val elim = the_elim_of thy predname
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2791
        val nparams = nparams_of thy predname
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2792
        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
  2793
          (Drule.export_without_context o Skip_Proof.make_thm thy)
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36533
diff changeset
  2794
          (mk_casesrule (ProofContext.init_global thy) nparams intros)
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2795
      in
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2796
        if not (Thm.equiv_thm (elim, elim')) then
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2797
          error "Introduction and elimination rules do not match!"
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2798
        else true
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2799
      end
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2800
  in forall check prednames end
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2801
*)
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
  2802
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2803
(* create code equation *)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2804
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2805
fun add_code_equations ctxt preds result_thmss =
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2806
  let
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2807
    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
  2808
      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
  2809
        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
  2810
      in
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2811
        if member (op =) (modes_of Pred ctxt predname) full_mode then
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2812
          let
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2813
            val Ts = binder_types T
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2814
            val arg_names = Name.variant_list []
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2815
              (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
  2816
            val args = map2 (curry Free) arg_names Ts
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2817
            val predfun = Const (function_name_of Pred ctxt predname full_mode,
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2818
              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
  2819
            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
  2820
            val eq_term = HOLogic.mk_Trueprop
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2821
              (HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs))
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2822
            val def = predfun_definition_of ctxt predname full_mode
33441
99a5f22a967d eliminated funny record patterns and made SML/NJ happy;
wenzelm
parents: 33377
diff changeset
  2823
            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
  2824
              (HOL_basic_ss addsimps [def, @{thm holds_eq}, @{thm eval_pred}]) 1
35888
d902054e7ac6 contextifying the proof procedure in the predicate compiler
bulwahn
parents: 35887
diff changeset
  2825
            val eq = Goal.prove ctxt arg_names [] eq_term tac
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2826
          in
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2827
            (pred, result_thms @ [eq])
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2828
          end
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2829
        else
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2830
          (pred, result_thms)
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2831
      end
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2832
  in
33629
5f35cf91c6a4 improving code quality thanks to Florian's code review
bulwahn
parents: 33628
diff changeset
  2833
    map2 add_code_equation preds result_thmss
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2834
  end
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2835
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2836
(** main function of predicate compiler **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2837
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2838
datatype steps = Steps of
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2839
  {
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2840
  define_functions : options -> (string * typ) list -> string * (bool * mode) list -> theory -> theory,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2841
  prove : options -> theory -> (string * (term list * indprem list) list) list -> (string * typ) list
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2842
    -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2843
  add_code_equations : Proof.context -> (string * typ) list
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2844
    -> (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
  2845
  comp_modifiers : Comp_Mod.comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  2846
  use_generators : bool,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2847
  qname : bstring
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2848
  }
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2849
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2850
fun add_equations_of steps mode_analysis_options options prednames thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2851
  let
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2852
    fun dest_steps (Steps s) = s
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2853
    val compilation = Comp_Mod.compilation (#comp_modifiers (dest_steps steps))
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2854
    val ctxt = ProofContext.init_global thy
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2855
    val _ = print_step options
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2856
      ("Starting predicate compiler (compilation: " ^ string_of_compilation compilation
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2857
        ^ ") for predicates " ^ commas prednames ^ "...")
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  2858
      (*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
  2859
      (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2860
    val _ =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2861
      if show_intermediate_results options then
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2862
        tracing (commas (map (Display.string_of_thm ctxt) (maps (intros_of ctxt) prednames)))
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2863
      else ()
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2864
    val (preds, all_vs, param_vs, all_modes, clauses) =
39310
17ef4415b17c removing obsolete argument in prepare_intrs; passing context instead of theory in prepare_intrs
bulwahn
parents: 39299
diff changeset
  2865
      prepare_intrs options ctxt prednames (maps (intros_of ctxt) prednames)
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2866
    val _ = print_step options "Infering modes..."
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  2867
    val (lookup_mode, lookup_neg_mode, needs_random) = (modes_of compilation ctxt,
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  2868
      modes_of (negative_compilation_of compilation) ctxt, needs_random ctxt)
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  2869
    val ((moded_clauses, needs_random), errors) =
37009
4ba91ea2bf6d deactivated timing of infering modes
bulwahn
parents: 37007
diff changeset
  2870
      Output.cond_timeit (!Quickcheck.timing) "Infering modes"
36251
5fd5d732a4ea only add relevant predicates to the list of extra modes
bulwahn
parents: 36247
diff changeset
  2871
      (fn _ => infer_modes mode_analysis_options
39273
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  2872
        options (lookup_mode, lookup_neg_mode, needs_random) ctxt preds all_modes param_vs clauses)
92aa2a0f7399 refactoring mode inference so that the theory is not changed in the mode inference procedure
bulwahn
parents: 39201
diff changeset
  2873
    val thy' = fold (fn (s, ms) => set_needs_random s ms) needs_random 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
  2874
    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
  2875
    val _ = check_expected_modes preds options modes
39201
234e6ef4ff67 using the proposed modes for starting the fixpoint iteration in the mode analysis
bulwahn
parents: 39194
diff changeset
  2876
    val _ = check_proposed_modes preds options modes errors
37004
c62f743e37d4 removing unused argument in print_modes function
bulwahn
parents: 37003
diff changeset
  2877
    val _ = print_modes options modes
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2878
    val _ = print_step options "Defining executable functions..."
36252
beba03215d8f adding more profiling to the predicate compiler
bulwahn
parents: 36251
diff changeset
  2879
    val thy'' =
36261
01ccbaa3f49f make profiling depend on reference Quickcheck.timing
bulwahn
parents: 36258
diff changeset
  2880
      Output.cond_timeit (!Quickcheck.timing) "Defining executable functions..."
36252
beba03215d8f adding more profiling to the predicate compiler
bulwahn
parents: 36251
diff changeset
  2881
      (fn _ => fold (#define_functions (dest_steps steps) options preds) modes thy'
beba03215d8f adding more profiling to the predicate compiler
bulwahn
parents: 36251
diff changeset
  2882
      |> Theory.checkpoint)
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  2883
    val ctxt'' = ProofContext.init_global thy''
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2884
    val _ = print_step options "Compiling equations..."
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2885
    val compiled_terms =
36261
01ccbaa3f49f make profiling depend on reference Quickcheck.timing
bulwahn
parents: 36258
diff changeset
  2886
      Output.cond_timeit (!Quickcheck.timing) "Compiling equations...." (fn _ =>
36254
95ef0a3cf31c added switch detection to the predicate compiler
bulwahn
parents: 36252
diff changeset
  2887
        compile_preds options
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  2888
          (#comp_modifiers (dest_steps steps)) ctxt'' all_vs param_vs preds moded_clauses)
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 37004
diff changeset
  2889
    val _ = print_compiled_terms options ctxt'' compiled_terms
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
  2890
    val _ = print_step options "Proving equations..."
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2891
    val result_thms =
36261
01ccbaa3f49f make profiling depend on reference Quickcheck.timing
bulwahn
parents: 36258
diff changeset
  2892
      Output.cond_timeit (!Quickcheck.timing) "Proving equations...." (fn _ =>
36252
beba03215d8f adding more profiling to the predicate compiler
bulwahn
parents: 36251
diff changeset
  2893
      #prove (dest_steps steps) options thy'' clauses preds moded_clauses compiled_terms)
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2894
    val result_thms' = #add_code_equations (dest_steps steps) ctxt'' preds
33376
5cb663aa48ee predicate compiler creates code equations for predicates with full mode
bulwahn
parents: 33375
diff changeset
  2895
      (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
  2896
    val qname = #qname (dest_steps steps)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2897
    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2898
      (fn thm => Context.mapping (Code.add_eqn thm) I))))
36252
beba03215d8f adding more profiling to the predicate compiler
bulwahn
parents: 36251
diff changeset
  2899
    val thy''' =
36261
01ccbaa3f49f make profiling depend on reference Quickcheck.timing
bulwahn
parents: 36258
diff changeset
  2900
      Output.cond_timeit (!Quickcheck.timing) "Setting code equations...." (fn _ =>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39545
diff changeset
  2901
      fold (fn (name, result_thms) => fn thy => snd (Global_Theory.add_thmss
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2902
      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2903
        [attrib thy ])] thy))
36252
beba03215d8f adding more profiling to the predicate compiler
bulwahn
parents: 36251
diff changeset
  2904
      result_thms' thy'' |> Theory.checkpoint)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2905
  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
  2906
    thy'''
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2907
  end
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2908
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2909
fun extend' value_of edges_of key (G, visited) =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2910
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2911
    val (G', v) = case try (Graph.get_node G) key of
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2912
        SOME v => (G, v)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2913
      | 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
  2914
    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
  2915
      (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
  2916
      (G', key :: visited)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2917
  in
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2918
    (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2919
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2920
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2921
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
  2922
  
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  2923
fun gen_add_equations steps options names thy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2924
  let
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2925
    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
  2926
    val defined = defined_functions (Comp_Mod.compilation (#comp_modifiers (dest_steps steps)))
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2927
    val ctxt = ProofContext.init_global thy
33482
5029ec373036 strictly respecting the line margin in the predicate compiler core
bulwahn
parents: 33479
diff changeset
  2928
    val thy' = thy
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  2929
      |> PredData.map (fold (extend (fetch_pred_data ctxt) (depending_preds_of ctxt)) names)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2930
      |> Theory.checkpoint;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2931
    fun strong_conn_of gr keys =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2932
      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2933
    val scc = strong_conn_of (PredData.get thy') names
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2934
    fun preprocess name thy =
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2935
      PredData.map (Graph.map_node name (map_pred_data (apfst (fn (rules, preprocessed) =>
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2936
        if preprocessed then (rules, preprocessed)
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2937
        else
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2938
          let
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2939
            val (named_intros, SOME elim) = rules
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2940
            val named_intros' = map (apsnd (preprocess_intro thy)) named_intros
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2941
            val pred = Const (name, Sign.the_const_type thy name)
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2942
            val ctxt = ProofContext.init_global thy
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2943
            val elim_t = mk_casesrule ctxt pred (map snd named_intros')
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2944
            val nparams = (case try (Inductive.the_inductive ctxt) name of
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2945
                SOME (_, result) => length (Inductive.params_of (#raw_induct result))
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2946
              | NONE => 0)
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2947
            val elim' = prove_casesrule ctxt (pred, (elim, 0)) elim_t
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2948
          in
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2949
            ((named_intros', SOME elim'), true)
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2950
          end))))
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2951
        thy
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2952
    val thy'' = fold preprocess (flat scc) thy'
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2953
    val thy''' = fold_rev
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2954
      (fn preds => fn thy =>
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  2955
        if not (forall (defined (ProofContext.init_global thy)) preds) then
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2956
          let
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  2957
            val mode_analysis_options = {use_generators = #use_generators (dest_steps steps),
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2958
              reorder_premises =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2959
                not (no_topmost_reordering options andalso not (null (inter (op =) preds names))),
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  2960
              infer_pos_and_neg_modes = #use_generators (dest_steps steps)}
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2961
          in
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2962
            add_equations_of steps mode_analysis_options options preds thy
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2963
          end
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
  2964
        else thy)
39657
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2965
      scc thy'' |> Theory.checkpoint
5e57675b7e40 moving the preprocessing of introduction rules after the code_pred command; added tuple expansion preprocessing of elimination rule
bulwahn
parents: 39653
diff changeset
  2966
  in thy''' end
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2967
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2968
val add_equations = gen_add_equations
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2969
  (Steps {
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2970
  define_functions =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2971
    fn options => fn preds => fn (s, modes) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2972
      create_definitions
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2973
      options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) 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
  2974
  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
  2975
  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
  2976
  comp_modifiers = predicate_comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  2977
  use_generators = false,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  2978
  qname = "equation"})
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
  2979
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
  2980
val add_depth_limited_equations = gen_add_equations
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2981
  (Steps {
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2982
  define_functions =
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2983
    fn options => fn preds => fn (s, modes) =>
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2984
    define_functions depth_limited_comp_modifiers PredicateCompFuns.compfuns
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2985
    options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  2986
  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
  2987
  add_code_equations = K (K I),
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2988
  comp_modifiers = depth_limited_comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  2989
  use_generators = false,
33330
d6eb7f19bfc6 encapsulating records with datatype constructors and adding type annotations to make SML/NJ happy
bulwahn
parents: 33328
diff changeset
  2990
  qname = "depth_limited_equation"})
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  2991
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2992
val add_annotated_equations = gen_add_equations
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2993
  (Steps {
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2994
  define_functions =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2995
    fn options => fn preds => fn (s, modes) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2996
      define_functions annotated_comp_modifiers PredicateCompFuns.compfuns options preds
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  2997
      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  2998
  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
  2999
  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
  3000
  comp_modifiers = annotated_comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  3001
  use_generators = false,
33473
3b275a0bf18c adding tracing function for evaluated code; annotated compilation in the predicate compiler
bulwahn
parents: 33377
diff changeset
  3002
  qname = "annotated_equation"})
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3003
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3004
val add_random_equations = gen_add_equations
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3005
  (Steps {
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3006
  define_functions =
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3007
    fn options => fn preds => fn (s, modes) =>
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3008
      define_functions random_comp_modifiers PredicateCompFuns.compfuns options preds
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3009
      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3010
  comp_modifiers = random_comp_modifiers,
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3011
  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
  3012
  add_code_equations = K (K I),
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  3013
  use_generators = true,
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33333
diff changeset
  3014
  qname = "random_equation"})
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3015
35881
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3016
val add_depth_limited_random_equations = gen_add_equations
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3017
  (Steps {
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3018
  define_functions =
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3019
    fn options => fn preds => fn (s, modes) =>
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3020
      define_functions depth_limited_random_comp_modifiers PredicateCompFuns.compfuns options preds
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3021
      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3022
  comp_modifiers = depth_limited_random_comp_modifiers,
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3023
  prove = prove_by_skip,
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3024
  add_code_equations = K (K I),
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  3025
  use_generators = true,
35881
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3026
  qname = "depth_limited_random_equation"})
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3027
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3028
val add_dseq_equations = gen_add_equations
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3029
  (Steps {
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3030
  define_functions =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3031
  fn options => fn preds => fn (s, modes) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3032
    define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3033
    options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) 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
  3034
  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
  3035
  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
  3036
  comp_modifiers = dseq_comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  3037
  use_generators = false,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3038
  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
  3039
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3040
val add_random_dseq_equations = gen_add_equations
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3041
  (Steps {
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3042
  define_functions =
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3043
    fn options => fn preds => fn (s, modes) =>
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3044
    let
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3045
      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3046
      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3047
    in define_functions pos_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3048
      options preds (s, pos_modes)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3049
      #> define_functions neg_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3050
      options preds (s, neg_modes)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3051
    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
  3052
  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
  3053
  add_code_equations = K (K I),
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3054
  comp_modifiers = pos_random_dseq_comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  3055
  use_generators = true,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3056
  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
  3057
36018
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3058
val add_new_random_dseq_equations = gen_add_equations
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3059
  (Steps {
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3060
  define_functions =
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3061
    fn options => fn preds => fn (s, modes) =>
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3062
    let
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3063
      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3064
      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3065
    in define_functions new_pos_random_dseq_comp_modifiers New_Pos_Random_Sequence_CompFuns.compfuns
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3066
      options preds (s, pos_modes)
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3067
      #> define_functions new_neg_random_dseq_comp_modifiers New_Neg_Random_Sequence_CompFuns.compfuns
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3068
      options preds (s, neg_modes)
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3069
    end,
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3070
  prove = prove_by_skip,
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3071
  add_code_equations = K (K I),
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3072
  comp_modifiers = new_pos_random_dseq_comp_modifiers,
39761
c2a76ec6e2d9 renaming use_random to use_generators in the predicate compiler
bulwahn
parents: 39760
diff changeset
  3073
  use_generators = true,
36018
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3074
  qname = "new_random_dseq_equation"})
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3075
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3076
(** user interface **)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3077
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3078
(* code_pred_intro attribute *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3079
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3080
fun attrib' f opt_case_name =
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3081
  Thm.declaration_attribute (fn thm => Context.mapping (f (opt_case_name, thm)) I);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3082
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3083
val code_pred_intro_attrib = attrib' add_intro NONE;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3084
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3085
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3086
(*FIXME
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3087
- Naming of auxiliary rules necessary?
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3088
*)
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3089
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3090
val setup = PredData.put (Graph.empty) #>
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3091
  Attrib.setup @{binding code_pred_intro} (Scan.lift (Scan.option Args.name) >> attrib' add_intro)
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3092
    "adding alternative introduction rules for code generation of inductive predicates"
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3093
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
  3094
(* TODO: make Theory_Data to Generic_Data & remove duplication of local theory and theory *)
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38558
diff changeset
  3095
(* FIXME ... this is important to avoid changing the background theory below *)
33132
07efd452a698 moved argument expected_modes into options; improved mode check to only check mode of the named predicate
bulwahn
parents: 33131
diff changeset
  3096
fun generic_code_pred prep_const options raw_const lthy =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3097
  let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3098
    val thy = ProofContext.theory_of lthy
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3099
    val const = prep_const thy raw_const
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3100
    val ctxt = ProofContext.init_global thy
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38558
diff changeset
  3101
    val lthy' = Local_Theory.background_theory (PredData.map
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  3102
        (extend (fetch_pred_data ctxt) (depending_preds_of ctxt) const)) lthy
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3103
    val thy' = ProofContext.theory_of lthy'
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  3104
    val ctxt' = ProofContext.init_global thy'
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  3105
    val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim ctxt')
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3106
    fun mk_cases const =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3107
      let
39299
e98a06145530 directly computing the values of interest instead of composing functions in an unintelligent way that causes exponential much garbage; using the latest theory
bulwahn
parents: 39273
diff changeset
  3108
        val T = Sign.the_const_type thy' const
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
  3109
        val pred = Const (const, T)
37007
116670499530 changing operations for accessing data to work with contexts
bulwahn
parents: 37006
diff changeset
  3110
        val intros = intros_of ctxt' 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
  3111
      in mk_casesrule lthy' pred intros end  
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3112
    val cases_rules = map mk_cases preds
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3113
    val cases =
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3114
      map2 (fn pred_name => fn case_rule => Rule_Cases.Case {fixes = [],
39652
b1febbbda458 improving naming of assumptions in code_pred
bulwahn
parents: 39651
diff changeset
  3115
        assumes = ("that", tl (Logic.strip_imp_prems case_rule))
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3116
          :: (map_filter (fn (fact_name, fact) => Option.map (rpair [fact]) fact_name)
39652
b1febbbda458 improving naming of assumptions in code_pred
bulwahn
parents: 39651
diff changeset
  3117
            ((SOME (Long_Name.base_name pred_name ^ ".prems") :: names_of ctxt' pred_name) ~~ Logic.strip_imp_prems case_rule)),
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3118
        binds = [], cases = []}) preds cases_rules
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3119
    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
  3120
    val lthy'' = lthy'
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39477
diff changeset
  3121
      |> fold Variable.auto_fixes cases_rules
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3122
      |> ProofContext.add_cases true case_env
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3123
    fun after_qed thms goal_ctxt =
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3124
      let
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3125
        val global_thms = ProofContext.export goal_ctxt
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36533
diff changeset
  3126
          (ProofContext.init_global (ProofContext.theory_of goal_ctxt)) (map the_single thms)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3127
      in
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38558
diff changeset
  3128
        goal_ctxt |> Local_Theory.background_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
  3129
          ((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
  3130
             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
  3131
           | DSeq => add_dseq_equations
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3132
           | Pos_Random_DSeq => add_random_dseq_equations
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3133
           | Depth_Limited => add_depth_limited_equations
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3134
           | Random => add_random_equations
35881
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3135
           | Depth_Limited_Random => add_depth_limited_random_equations
36018
096ec83348f3 added new compilation to predicate_compiler
bulwahn
parents: 36004
diff changeset
  3136
           | New_Pos_Random_DSeq => add_new_random_dseq_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
  3137
           | 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
  3138
           ) options [const]))
33144
1831516747d4 removed obsolete GeneratorPrem; clean-up after modularization; tuned
bulwahn
parents: 33143
diff changeset
  3139
      end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3140
  in
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 36261
diff changeset
  3141
    Proof.theorem NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3142
  end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3143
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3144
val code_pred = generic_code_pred (K I);
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3145
val code_pred_cmd = generic_code_pred Code.read_const
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3146
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3147
(* transformation for code generation *)
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3148
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3149
structure Pred_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3150
  type T = unit -> term Predicate.pred
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3151
  fun init _ () = error "Pred_Result"
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3152
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3153
val put_pred_result = Pred_Result.put;
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3154
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3155
structure Pred_Random_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3156
  type T = unit -> int * int -> term Predicate.pred * (int * int)
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3157
  fun init _ () = error "Pred_Random_Result"
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3158
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3159
val put_pred_random_result = Pred_Random_Result.put;
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3160
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3161
structure Dseq_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3162
  type T = unit -> term DSequence.dseq
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3163
  fun init _ () = error "Dseq_Result"
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3164
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3165
val put_dseq_result = Dseq_Result.put;
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3166
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3167
structure Dseq_Random_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3168
  type T = unit -> int -> int -> int * int -> term DSequence.dseq * (int * int)
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3169
  fun init _ () = error "Dseq_Random_Result"
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3170
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3171
val put_dseq_random_result = Dseq_Random_Result.put;
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3172
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3173
structure Lseq_Random_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3174
  type T = unit -> int -> int -> int * int -> int -> term Lazy_Sequence.lazy_sequence
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3175
  fun init _ () = error "Lseq_Random_Result"
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3176
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3177
val put_lseq_random_result = Lseq_Random_Result.put;
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3178
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3179
structure Lseq_Random_Stats_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3180
  type T = unit -> int -> int -> int * int -> int -> (term * int) Lazy_Sequence.lazy_sequence
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3181
  fun init _ () = error "Lseq_Random_Stats_Result"
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3182
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3183
val put_lseq_random_stats_result = Lseq_Random_Stats_Result.put;
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3184
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3185
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3186
fun analyze_compr ctxt compfuns param_user_modes (compilation, arguments) t_compr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3187
  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
  3188
    val all_modes_of = all_modes_of compilation
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3189
    val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3190
      | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3191
    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
  3192
    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
  3193
      (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
34963
haftmann
parents: 34948 34962
diff changeset
  3194
    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
  3195
    val body = subst_bounds (output_frees, body)
34963
haftmann
parents: 34948 34962
diff changeset
  3196
    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
  3197
    val output_tuple = HOLogic.mk_ptuple fp T_compr (rev output_frees)
36031
199fe16cdaab returning an more understandable user error message in the values command
bulwahn
parents: 36030
diff changeset
  3198
    val (pred as Const (name, T), all_args) =
199fe16cdaab returning an more understandable user error message in the values command
bulwahn
parents: 36030
diff changeset
  3199
      case strip_comb body of
199fe16cdaab returning an more understandable user error message in the values command
bulwahn
parents: 36030
diff changeset
  3200
        (Const (name, T), all_args) => (Const (name, T), all_args)
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3201
      | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3202
  in
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3203
    if defined_functions compilation ctxt name then
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32667
diff changeset
  3204
      let
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37146
diff changeset
  3205
        fun extract_mode (Const (@{const_name Pair}, _) $ t1 $ t2) = Pair (extract_mode t1, extract_mode t2)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3206
          | 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
  3207
          | 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
  3208
        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
  3209
        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
  3210
          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
  3211
            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
  3212
          | 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
  3213
          | 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
  3214
            (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
  3215
        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
  3216
          let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3217
            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
  3218
              | 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
  3219
              | 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
  3220
              | 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
  3221
                  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
  3222
              | 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
  3223
                  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
  3224
              | 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
  3225
                  instance_of (m1, Output) andalso instance_of (m2, Output)
37002
34e73e8bab66 improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents: 37001
diff changeset
  3226
              | instance_of (Input, Pair (m1, m2)) =
34e73e8bab66 improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents: 37001
diff changeset
  3227
                  instance_of (Input, m1) andalso instance_of (Input, m2)
34e73e8bab66 improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents: 37001
diff changeset
  3228
              | instance_of (Output, Pair (m1, m2)) =
34e73e8bab66 improved values command to handle a special case with tuples and polymorphic predicates more correctly
bulwahn
parents: 37001
diff changeset
  3229
                  instance_of (Output, m1) andalso instance_of (Output, m2)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3230
              | 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
  3231
          in forall instance_of (strip_fun_mode m1 ~~ strip_fun_mode m2) end
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3232
        val derivs = all_derivations_of ctxt (all_modes_of ctxt) [] body
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3233
          |> 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
  3234
            let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3235
              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
  3236
            in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3237
              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
  3238
              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
  3239
              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
  3240
            end)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3241
          |> 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
  3242
        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
  3243
            [] => error ("No mode possible for comprehension "
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3244
                    ^ Syntax.string_of_term ctxt t_compr)
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3245
          | [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
  3246
          | d :: _ :: _ => (warning ("Multiple modes possible for comprehension "
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3247
                    ^ Syntax.string_of_term ctxt t_compr); d);
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3248
        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
  3249
        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
  3250
          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
  3251
            Pred => []
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3252
          | Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3253
            [@{term "(1, 1) :: code_numeral * code_numeral"}]
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3254
          | Annotated => []
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3255
          | Depth_Limited => [HOLogic.mk_number @{typ "code_numeral"} (hd arguments)]
35881
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3256
          | Depth_Limited_Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3257
            [@{term "(1, 1) :: code_numeral * code_numeral"}]
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3258
          | DSeq => []
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3259
          | Pos_Random_DSeq => []
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3260
          | New_Pos_Random_DSeq => []
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3261
        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
  3262
          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
  3263
            Pred => predicate_comp_modifiers
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3264
          | Random => random_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
  3265
          | Depth_Limited => depth_limited_comp_modifiers
35881
aa412e08bfee adding depth_limited_random compilation to predicate compiler
bulwahn
parents: 35880
diff changeset
  3266
          | Depth_Limited_Random => depth_limited_random_comp_modifiers
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3267
          (*| Annotated => annotated_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
  3268
          | DSeq => dseq_comp_modifiers
35879
99818df5b8f5 reviving the classical depth-limited computation in the predicate compiler
bulwahn
parents: 35845
diff changeset
  3269
          | Pos_Random_DSeq => pos_random_dseq_comp_modifiers
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3270
          | New_Pos_Random_DSeq => new_pos_random_dseq_comp_modifiers
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3271
        val t_pred = compile_expr comp_modifiers ctxt
39648
655307cb8489 rewriting function mk_Eval_of in predicate compiler
bulwahn
parents: 39557
diff changeset
  3272
          (body, deriv) [] 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
  3273
        val T_pred = dest_predT compfuns (fastype_of t_pred)
39756
6c8e83d94536 consolidated tupled_lambda; moved to structure HOLogic
krauss
parents: 39685
diff changeset
  3274
        val arrange = HOLogic.tupled_lambda (HOLogic.mk_tuple outargs) output_tuple
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3275
      in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3276
        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
  3277
      end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3278
    else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3279
      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
  3280
  end
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3281
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3282
fun eval ctxt stats param_user_modes (options as (compilation, arguments)) k t_compr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3283
  let
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3284
    fun count xs x =
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3285
      let
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3286
        fun count' i [] = i
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3287
          | count' i (x' :: xs) = if x = x' then count' (i + 1) xs else count' i xs
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3288
      in count' 0 xs end
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3289
    fun accumulate xs = map (fn x => (x, count xs x)) (sort int_ord (distinct (op =) xs))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3290
    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
  3291
      case compilation of
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3292
        Random => 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
  3293
      | DSeq => DSequence_CompFuns.compfuns
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 35267
diff changeset
  3294
      | Pos_Random_DSeq => Random_Sequence_CompFuns.compfuns
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3295
      | New_Pos_Random_DSeq => New_Pos_Random_Sequence_CompFuns.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
  3296
      | _ => PredicateCompFuns.compfuns
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3297
    val t = analyze_compr ctxt 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
  3298
    val T = dest_predT compfuns (fastype_of t);
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3299
    val t' =
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3300
      if stats andalso compilation = New_Pos_Random_DSeq then
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3301
        mk_map compfuns T (HOLogic.mk_prodT (HOLogic.termT, @{typ code_numeral}))
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3302
          (absdummy (T, HOLogic.mk_prod (HOLogic.term_of_const T $ Bound 0,
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3303
            @{term Code_Numeral.of_nat} $ (HOLogic.size_const T $ Bound 0)))) t
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3304
      else
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3305
        mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3306
    val thy = ProofContext.theory_of ctxt
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3307
    val (ts, statistics) =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3308
      case compilation of
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3309
       (* Random =>
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3310
          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
  3311
          (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
  3312
            (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
35880
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3313
            |> Random_Engine.run))*)
2623b23e41fc a new simpler random compilation for the predicate compiler
bulwahn
parents: 35879
diff changeset
  3314
        Pos_Random_DSeq =>
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3315
          let
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3316
            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
  3317
          in
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3318
            rpair NONE (fst (DSequence.yieldn k
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3319
              (Code_Runtime.dynamic_value_strict (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result")
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3320
                thy NONE (fn proc => fn g => fn nrandom => fn size => fn s => g nrandom size s |>> DSequence.map proc)
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3321
                  t' [] nrandom size
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3322
                |> Random_Engine.run)
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3323
              depth true))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3324
          end
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3325
      | DSeq =>
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3326
          rpair NONE (fst (DSequence.yieldn k
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3327
            (Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result")
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3328
              thy NONE DSequence.map t' []) (the_single arguments) true))
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3329
      | New_Pos_Random_DSeq =>
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3330
          let
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3331
            val [nrandom, size, depth] = arguments
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3332
            val seed = Random_Engine.next_seed ()
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3333
          in
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3334
            if stats then
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3335
              apsnd (SOME o accumulate) (split_list
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3336
              (fst (Lazy_Sequence.yieldn k
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3337
                (Code_Runtime.dynamic_value_strict
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3338
                  (Lseq_Random_Stats_Result.get, put_lseq_random_stats_result, "Predicate_Compile_Core.put_lseq_random_stats_result")
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3339
                  thy NONE
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3340
                  (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
36533
f8df589ca2a5 dropped unnecessary ML code
haftmann
parents: 36510
diff changeset
  3341
                    |> Lazy_Sequence.mapa (apfst proc))
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3342
                    t' [] nrandom size seed depth))))
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3343
            else rpair NONE
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3344
              (fst (Lazy_Sequence.yieldn k
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3345
                (Code_Runtime.dynamic_value_strict
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39310
diff changeset
  3346
                  (Lseq_Random_Result.get, put_lseq_random_result, "Predicate_Compile_Core.put_lseq_random_result")
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3347
                  thy NONE 
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3348
                  (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
36533
f8df589ca2a5 dropped unnecessary ML code
haftmann
parents: 36510
diff changeset
  3349
                    |> Lazy_Sequence.mapa proc)
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3350
                    t' [] nrandom size seed depth)))
36020
3ee4c29ead7f adding values command for new monad; added new random monad compilation to predicate_compile_quickcheck
bulwahn
parents: 36019
diff changeset
  3351
          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
  3352
      | _ =>
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3353
          rpair NONE (fst (Predicate.yieldn k
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3354
            (Code_Runtime.dynamic_value_strict (Pred_Result.get, put_pred_result, "Predicate_Compile_Core.put_pred_result")
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
  3355
              thy NONE Predicate.map t' [])))
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3356
  in ((T, ts), statistics) end;
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3357
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3358
fun values ctxt param_user_modes ((raw_expected, stats), comp_options) k t_compr =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3359
  let
37003
a393a588b82e moving towards working with proof contexts in the predicate compiler
bulwahn
parents: 37002
diff changeset
  3360
    val ((T, ts), statistics) = eval ctxt stats param_user_modes comp_options k t_compr
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3361
    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
  3362
    val elems = HOLogic.mk_set T ts
33476
27cca5416a88 Adopted output of values command
bulwahn
parents: 33474
diff changeset
  3363
    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
  3364
    (* 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
  3365
    val () =
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3366
      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
  3367
        NONE => ()
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3368
      | 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
  3369
        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
  3370
        else
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
  3371
          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
  3372
            "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
  3373
            "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
  3374
  in
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3375
    (if k = ~1 orelse length ts < k then elems
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3376
    else Const (@{const_abbrev Set.union}, setT --> setT --> setT) $ elems $ cont, statistics)
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3377
  end;
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33620
diff changeset
  3378
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33478
diff changeset
  3379
fun values_cmd print_modes param_user_modes options k raw_t state =
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3380
  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
  3381
    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
  3382
    val t = Syntax.read_term ctxt raw_t
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3383
    val (t', stats) = values ctxt param_user_modes options k 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
  3384
    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
  3385
    val ctxt' = Variable.auto_fixes t' ctxt
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3386
    val pretty_stat =
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3387
      case stats of
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3388
          NONE => []
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3389
        | SOME xs =>
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3390
          let
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3391
            val total = fold (curry (op +)) (map snd xs) 0
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3392
            fun pretty_entry (s, n) =
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3393
              [Pretty.str "size", Pretty.brk 1,
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3394
               Pretty.str (string_of_int s), Pretty.str ":", Pretty.brk 1,
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3395
               Pretty.str (string_of_int n), Pretty.fbrk]
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3396
          in
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3397
            [Pretty.fbrk, Pretty.str "Statistics:", Pretty.fbrk,
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3398
             Pretty.str "total:", Pretty.brk 1, Pretty.str (string_of_int total), Pretty.fbrk]
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3399
             @ maps pretty_entry xs
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3400
          end
37146
f652333bbf8e renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
wenzelm
parents: 37135
diff changeset
  3401
    val p = Print_Mode.with_modes print_modes (fn () =>
36027
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3402
      Pretty.block ([Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3403
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]
29a15da9c63d added statistics to values command for random generation
bulwahn
parents: 36025
diff changeset
  3404
        @ pretty_stat)) ();
32667
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3405
  in Pretty.writeln p end;
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3406
09546e654222 moved predicate compiler to Tools
bulwahn
parents:
diff changeset
  3407
end;