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