src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML
author krauss
Tue, 19 Jul 2011 00:07:21 +0200
changeset 44767 8955dcac6c71
parent 44500 3803869014aa
child 44819 8f5add916a99
permissions -rw-r--r--
values_timeout defaults to 600.0 on SML/NJ -- saves us from cluttering all theories equivalent declarations
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@40234
    19
bulwahn@43012
    20
  val values_timeout : real Config.T
bulwahn@43012
    21
  
bulwahn@36998
    22
  val print_stored_rules : Proof.context -> unit
bulwahn@36994
    23
  val print_all_modes : compilation -> Proof.context -> unit
haftmann@39634
    24
haftmann@39634
    25
  val put_pred_result : (unit -> term Predicate.pred) -> Proof.context -> Proof.context
haftmann@39634
    26
  val put_pred_random_result : (unit -> int * int -> term Predicate.pred * (int * int)) ->
haftmann@39634
    27
    Proof.context -> Proof.context
haftmann@39634
    28
  val put_dseq_result : (unit -> term DSequence.dseq) -> Proof.context -> Proof.context
haftmann@39634
    29
  val put_dseq_random_result : (unit -> int -> int -> int * int -> term DSequence.dseq * (int * int)) ->
haftmann@39634
    30
    Proof.context -> Proof.context
bulwahn@40348
    31
  val put_new_dseq_result : (unit -> int -> term Lazy_Sequence.lazy_sequence) ->
bulwahn@40348
    32
    Proof.context -> Proof.context
haftmann@39634
    33
  val put_lseq_random_result :
haftmann@39634
    34
    (unit -> int -> int -> int * int -> int -> term Lazy_Sequence.lazy_sequence) ->
haftmann@39634
    35
    Proof.context -> Proof.context
haftmann@39634
    36
  val put_lseq_random_stats_result :
haftmann@39634
    37
    (unit -> int -> int -> int * int -> int -> (term * int) Lazy_Sequence.lazy_sequence) ->
haftmann@39634
    38
    Proof.context -> Proof.context
haftmann@39634
    39
bulwahn@33628
    40
  val code_pred_intro_attrib : attribute
bulwahn@32667
    41
  (* used by Quickcheck_Generator *) 
bulwahn@32667
    42
  (* temporary for testing of the compilation *)
bulwahn@36048
    43
  val add_equations : options -> string list -> theory -> theory
bulwahn@36048
    44
  val add_depth_limited_random_equations : options -> string list -> theory -> theory
bulwahn@36048
    45
  val add_random_dseq_equations : options -> string list -> theory -> theory
bulwahn@36048
    46
  val add_new_random_dseq_equations : options -> string list -> theory -> theory
bulwahn@40349
    47
  val add_generator_dseq_equations : options -> string list -> theory -> theory
bulwahn@33473
    48
  val mk_tracing : string -> term -> term
bulwahn@39556
    49
  val prepare_intrs : options -> Proof.context -> string list -> thm list ->
bulwahn@39183
    50
    ((string * typ) list * string list * string list * (string * mode list) list *
bulwahn@39183
    51
      (string *  (Term.term list * Predicate_Compile_Aux.indprem list) list) list)
bulwahn@39994
    52
  type mode_analysis_options =
bulwahn@39994
    53
   {use_generators : bool,
bulwahn@39994
    54
    reorder_premises : bool,
bulwahn@39994
    55
    infer_pos_and_neg_modes : bool}  
bulwahn@39183
    56
  datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
bulwahn@39556
    57
    | Mode_Pair of mode_derivation * mode_derivation | Term of mode
bulwahn@39701
    58
  val head_mode_of : mode_derivation -> mode
bulwahn@39183
    59
  type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
bulwahn@39183
    60
  type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
bulwahn@39183
    61
bulwahn@32667
    62
end;
bulwahn@32667
    63
bulwahn@32667
    64
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
bulwahn@32667
    65
struct
bulwahn@32667
    66
bulwahn@32668
    67
open Predicate_Compile_Aux;
bulwahn@40233
    68
open Core_Data;
bulwahn@40233
    69
open Mode_Inference;
bulwahn@40233
    70
open Predicate_Compile_Proof;
bulwahn@33108
    71
bulwahn@32667
    72
(** fundamentals **)
bulwahn@32667
    73
bulwahn@32667
    74
(* syntactic operations *)
bulwahn@32667
    75
bulwahn@32667
    76
fun mk_eq (x, xs) =
bulwahn@32667
    77
  let fun mk_eqs _ [] = []
bulwahn@32667
    78
        | mk_eqs a (b::cs) =
bulwahn@32667
    79
            HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
bulwahn@32667
    80
  in mk_eqs x xs end;
bulwahn@32667
    81
bulwahn@32667
    82
fun mk_scomp (t, u) =
bulwahn@32667
    83
  let
bulwahn@32667
    84
    val T = fastype_of t
bulwahn@32667
    85
    val U = fastype_of u
bulwahn@32667
    86
    val [A] = binder_types T
bulwahn@34935
    87
    val D = body_type U                   
bulwahn@32667
    88
  in 
bulwahn@32667
    89
    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
bulwahn@32667
    90
  end;
bulwahn@32667
    91
bulwahn@32667
    92
fun dest_randomT (Type ("fun", [@{typ Random.seed},
haftmann@37678
    93
  Type (@{type_name Product_Type.prod}, [Type (@{type_name Product_Type.prod}, [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
bulwahn@32667
    94
  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
bulwahn@32667
    95
bulwahn@33473
    96
fun mk_tracing s t =
bulwahn@33473
    97
  Const(@{const_name Code_Evaluation.tracing},
bulwahn@33473
    98
    @{typ String.literal} --> (fastype_of t) --> (fastype_of t)) $ (HOLogic.mk_literal s) $ t
bulwahn@33473
    99
bulwahn@34935
   100
(* representation of inferred clauses with modes *)
bulwahn@32667
   101
bulwahn@34935
   102
type moded_clause = term list * (indprem * mode_derivation) list
bulwahn@32668
   103
bulwahn@35324
   104
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
bulwahn@32667
   105
bulwahn@32667
   106
(* diagnostic display functions *)
bulwahn@32667
   107
bulwahn@36995
   108
fun print_modes options modes =
bulwahn@33243
   109
  if show_modes options then
bulwahn@33326
   110
    tracing ("Inferred modes:\n" ^
bulwahn@33243
   111
      cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
bulwahn@35324
   112
        (fn (p, m) => string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
bulwahn@33243
   113
  else ()
bulwahn@32667
   114
bulwahn@36995
   115
fun print_pred_mode_table string_of_entry pred_mode_table =
bulwahn@32667
   116
  let
bulwahn@35324
   117
    fun print_mode pred ((pol, mode), entry) =  "mode : " ^ string_of_mode mode
bulwahn@33619
   118
      ^ string_of_entry pred mode entry
bulwahn@32667
   119
    fun print_pred (pred, modes) =
bulwahn@32667
   120
      "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
bulwahn@33326
   121
    val _ = tracing (cat_lines (map print_pred pred_mode_table))
bulwahn@32667
   122
  in () end;
bulwahn@32667
   123
bulwahn@36996
   124
fun print_compiled_terms options ctxt =
bulwahn@33139
   125
  if show_compilation options then
bulwahn@36996
   126
    print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term ctxt)
bulwahn@33139
   127
  else K ()
bulwahn@33139
   128
bulwahn@36998
   129
fun print_stored_rules ctxt =
bulwahn@32667
   130
  let
wenzelm@43232
   131
    val preds = Graph.keys (PredData.get (Proof_Context.theory_of ctxt))
bulwahn@32667
   132
    fun print pred () = let
bulwahn@32667
   133
      val _ = writeln ("predicate: " ^ pred)
bulwahn@32667
   134
      val _ = writeln ("introrules: ")
bulwahn@36998
   135
      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm ctxt thm))
bulwahn@36998
   136
        (rev (intros_of ctxt pred)) ()
bulwahn@32667
   137
    in
bulwahn@36998
   138
      if (has_elim ctxt pred) then
bulwahn@36998
   139
        writeln ("elimrule: " ^ Display.string_of_thm ctxt (the_elim_of ctxt pred))
bulwahn@32667
   140
      else
bulwahn@32667
   141
        writeln ("no elimrule defined")
bulwahn@32667
   142
    end
bulwahn@32667
   143
  in
bulwahn@32667
   144
    fold print preds ()
bulwahn@32667
   145
  end;
bulwahn@32667
   146
bulwahn@36994
   147
fun print_all_modes compilation ctxt =
bulwahn@32667
   148
  let
bulwahn@32667
   149
    val _ = writeln ("Inferred modes:")
bulwahn@32667
   150
    fun print (pred, modes) u =
bulwahn@32667
   151
      let
bulwahn@32667
   152
        val _ = writeln ("predicate: " ^ pred)
bulwahn@34935
   153
        val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
bulwahn@33619
   154
      in u end
bulwahn@32667
   155
  in
bulwahn@36994
   156
    fold print (all_modes_of compilation ctxt) ()
bulwahn@32667
   157
  end
bulwahn@33129
   158
bulwahn@33132
   159
(* validity checks *)
bulwahn@33132
   160
bulwahn@40380
   161
fun check_expected_modes options preds modes =
bulwahn@33752
   162
  case expected_modes options of
bulwahn@33752
   163
    SOME (s, ms) => (case AList.lookup (op =) modes s of
bulwahn@33752
   164
      SOME modes =>
bulwahn@33752
   165
        let
bulwahn@35324
   166
          val modes' = map snd modes
bulwahn@33752
   167
        in
bulwahn@34935
   168
          if not (eq_set eq_mode (ms, modes')) then
bulwahn@33752
   169
            error ("expected modes were not inferred:\n"
bulwahn@34935
   170
            ^ "  inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes')  ^ "\n"
bulwahn@34935
   171
            ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms))
bulwahn@33752
   172
          else ()
bulwahn@33752
   173
        end
bulwahn@33752
   174
      | NONE => ())
bulwahn@33752
   175
  | NONE => ()
bulwahn@33752
   176
bulwahn@40380
   177
fun check_proposed_modes options preds modes errors =
bulwahn@39628
   178
  map (fn (s, _) => case proposed_modes options s of
bulwahn@39628
   179
    SOME ms => (case AList.lookup (op =) modes s of
bulwahn@33752
   180
      SOME inferred_ms =>
bulwahn@33752
   181
        let
bulwahn@39431
   182
          val preds_without_modes = map fst (filter (null o snd) modes)
bulwahn@35324
   183
          val modes' = map snd inferred_ms
bulwahn@33752
   184
        in
bulwahn@34935
   185
          if not (eq_set eq_mode (ms, modes')) then
bulwahn@33752
   186
            error ("expected modes were not inferred:\n"
bulwahn@34935
   187
            ^ "  inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes')  ^ "\n"
bulwahn@34935
   188
            ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms) ^ "\n"
bulwahn@39629
   189
            ^ (if show_invalid_clauses options then
bulwahn@39629
   190
            ("For the following clauses, the following modes could not be inferred: " ^ "\n"
bulwahn@39629
   191
            ^ cat_lines errors) else "") ^
bulwahn@33752
   192
            (if not (null preds_without_modes) then
bulwahn@33752
   193
              "\n" ^ "No mode inferred for the predicates " ^ commas preds_without_modes
bulwahn@33752
   194
            else ""))
bulwahn@33752
   195
          else ()
bulwahn@33752
   196
        end
bulwahn@33752
   197
      | NONE => ())
bulwahn@39628
   198
  | NONE => ()) preds
bulwahn@33132
   199
bulwahn@39875
   200
fun check_matches_type ctxt predname T ms =
bulwahn@39875
   201
  let
bulwahn@39875
   202
    fun check (m as Fun (m1, m2)) (Type("fun", [T1,T2])) = check m1 T1 andalso check m2 T2
bulwahn@39996
   203
      | check m (T as Type("fun", _)) = (m = Input orelse m = Output)
bulwahn@39875
   204
      | check (Pair (m1, m2)) (Type (@{type_name Product_Type.prod}, [T1, T2])) =
bulwahn@39875
   205
          check m1 T1 andalso check m2 T2 
bulwahn@39875
   206
      | check Input T = true
bulwahn@39875
   207
      | check Output T = true
bulwahn@39875
   208
      | check Bool @{typ bool} = true
bulwahn@39875
   209
      | check _ _ = false
bulwahn@39996
   210
    fun check_consistent_modes ms =
bulwahn@39996
   211
      if forall (fn Fun (m1', m2') => true | _ => false) ms then
bulwahn@39996
   212
        pairself check_consistent_modes (split_list (map (fn Fun (m1, m2) => (m1, m2)) ms))
bulwahn@39996
   213
        |> (fn (res1, res2) => res1 andalso res2) 
bulwahn@39996
   214
      else if forall (fn Input => true | Output => true | Pair _ => true | _ => false) ms then
bulwahn@39996
   215
        true
bulwahn@39996
   216
      else if forall (fn Bool => true | _ => false) ms then
bulwahn@39996
   217
        true
bulwahn@39996
   218
      else
bulwahn@39996
   219
        false
bulwahn@39875
   220
    val _ = map
bulwahn@39875
   221
      (fn mode =>
bulwahn@39996
   222
        if length (strip_fun_mode mode) = length (binder_types T)
bulwahn@39996
   223
          andalso (forall (uncurry check) (strip_fun_mode mode ~~ binder_types T)) then ()
bulwahn@39875
   224
        else error (string_of_mode mode ^ " is not a valid mode for " ^ Syntax.string_of_typ ctxt T
bulwahn@39875
   225
        ^ " at predicate " ^ predname)) ms
bulwahn@39996
   226
    val _ =
bulwahn@39996
   227
     if check_consistent_modes ms then ()
bulwahn@39996
   228
     else error (commas (map string_of_mode ms) ^
bulwahn@39996
   229
       " are inconsistent modes for predicate " ^ predname)
bulwahn@39875
   230
  in
bulwahn@39875
   231
    ms
bulwahn@39875
   232
  end
bulwahn@39875
   233
bulwahn@36013
   234
(* compilation modifiers *)
bulwahn@36013
   235
bulwahn@36013
   236
structure Comp_Mod =
bulwahn@36013
   237
struct
bulwahn@36013
   238
bulwahn@36013
   239
datatype comp_modifiers = Comp_Modifiers of
bulwahn@36013
   240
{
bulwahn@36013
   241
  compilation : compilation,
bulwahn@36013
   242
  function_name_prefix : string,
bulwahn@36013
   243
  compfuns : compilation_funs,
bulwahn@36013
   244
  mk_random : typ -> term list -> term,
bulwahn@36013
   245
  modify_funT : typ -> typ,
bulwahn@36013
   246
  additional_arguments : string list -> term list,
bulwahn@36013
   247
  wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term,
bulwahn@36013
   248
  transform_additional_arguments : indprem -> term list -> term list
bulwahn@36013
   249
}
bulwahn@36013
   250
bulwahn@36013
   251
fun dest_comp_modifiers (Comp_Modifiers c) = c
bulwahn@36013
   252
bulwahn@36013
   253
val compilation = #compilation o dest_comp_modifiers
bulwahn@36013
   254
val function_name_prefix = #function_name_prefix o dest_comp_modifiers
bulwahn@36013
   255
val compfuns = #compfuns o dest_comp_modifiers
bulwahn@36013
   256
bulwahn@36013
   257
val mk_random = #mk_random o dest_comp_modifiers
bulwahn@36013
   258
val funT_of' = funT_of o compfuns
bulwahn@36013
   259
val modify_funT = #modify_funT o dest_comp_modifiers
bulwahn@36013
   260
fun funT_of comp mode = modify_funT comp o funT_of' comp mode
bulwahn@36013
   261
bulwahn@36013
   262
val additional_arguments = #additional_arguments o dest_comp_modifiers
bulwahn@36013
   263
val wrap_compilation = #wrap_compilation o dest_comp_modifiers
bulwahn@36013
   264
val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers
bulwahn@36013
   265
bulwahn@40230
   266
fun set_compfuns compfuns' (Comp_Modifiers {compilation, function_name_prefix, compfuns, mk_random,
bulwahn@40230
   267
    modify_funT, additional_arguments, wrap_compilation, transform_additional_arguments}) =
bulwahn@40230
   268
    (Comp_Modifiers {compilation = compilation, function_name_prefix = function_name_prefix,
bulwahn@40230
   269
    compfuns = compfuns', mk_random = mk_random, modify_funT = modify_funT,
bulwahn@40230
   270
    additional_arguments = additional_arguments, wrap_compilation = wrap_compilation,
bulwahn@40230
   271
    transform_additional_arguments = transform_additional_arguments})
bulwahn@40230
   272
bulwahn@36013
   273
end;
bulwahn@36013
   274
bulwahn@40232
   275
fun unlimited_compfuns_of true New_Pos_Random_DSeq =
bulwahn@40230
   276
    New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
bulwahn@40232
   277
  | unlimited_compfuns_of false New_Pos_Random_DSeq =
bulwahn@40230
   278
    New_Neg_Random_Sequence_CompFuns.depth_unlimited_compfuns
bulwahn@40232
   279
  | unlimited_compfuns_of true Pos_Generator_DSeq =
bulwahn@40232
   280
    New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
bulwahn@40232
   281
  | unlimited_compfuns_of false Pos_Generator_DSeq =
bulwahn@40232
   282
    New_Neg_DSequence_CompFuns.depth_unlimited_compfuns
bulwahn@40232
   283
  | unlimited_compfuns_of _ c =
bulwahn@40232
   284
    raise Fail ("No unlimited compfuns for compilation " ^ string_of_compilation c)
bulwahn@40232
   285
bulwahn@40230
   286
fun limited_compfuns_of true Predicate_Compile_Aux.New_Pos_Random_DSeq =
bulwahn@40230
   287
    New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
bulwahn@40230
   288
  | limited_compfuns_of false Predicate_Compile_Aux.New_Pos_Random_DSeq =
bulwahn@40230
   289
    New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns
bulwahn@40232
   290
  | limited_compfuns_of true Pos_Generator_DSeq =
bulwahn@40232
   291
    New_Pos_DSequence_CompFuns.depth_limited_compfuns
bulwahn@40232
   292
  | limited_compfuns_of false Pos_Generator_DSeq =
bulwahn@40232
   293
    New_Neg_DSequence_CompFuns.depth_limited_compfuns
bulwahn@40232
   294
  | limited_compfuns_of _ c =
bulwahn@40232
   295
    raise Fail ("No limited compfuns for compilation " ^ string_of_compilation c)
bulwahn@40230
   296
bulwahn@36013
   297
val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   298
  {
bulwahn@36013
   299
  compilation = Depth_Limited,
bulwahn@36013
   300
  function_name_prefix = "depth_limited_",
bulwahn@36013
   301
  compfuns = PredicateCompFuns.compfuns,
bulwahn@36013
   302
  mk_random = (fn _ => error "no random generation"),
bulwahn@36013
   303
  additional_arguments = fn names =>
bulwahn@36013
   304
    let
wenzelm@44206
   305
      val depth_name = singleton (Name.variant_list names) "depth"
bulwahn@36013
   306
    in [Free (depth_name, @{typ code_numeral})] end,
bulwahn@36013
   307
  modify_funT = (fn T => let val (Ts, U) = strip_type T
bulwahn@36013
   308
  val Ts' = [@{typ code_numeral}] in (Ts @ Ts') ---> U end),
bulwahn@36013
   309
  wrap_compilation =
bulwahn@36013
   310
    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
bulwahn@36013
   311
    let
bulwahn@36013
   312
      val [depth] = additional_arguments
bulwahn@40381
   313
      val (_, Ts) = split_modeT mode (binder_types T)
bulwahn@36013
   314
      val T' = mk_predT compfuns (HOLogic.mk_tupleT Ts)
bulwahn@36013
   315
      val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
bulwahn@36013
   316
    in
bulwahn@36013
   317
      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
bulwahn@36013
   318
        $ mk_bot compfuns (dest_predT compfuns T')
bulwahn@36013
   319
        $ compilation
bulwahn@36013
   320
    end,
bulwahn@36013
   321
  transform_additional_arguments =
bulwahn@36013
   322
    fn prem => fn additional_arguments =>
bulwahn@36013
   323
    let
bulwahn@36013
   324
      val [depth] = additional_arguments
bulwahn@36013
   325
      val depth' =
bulwahn@36013
   326
        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
bulwahn@36013
   327
          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
bulwahn@36013
   328
    in [depth'] end
bulwahn@36013
   329
  }
bulwahn@36013
   330
bulwahn@36013
   331
val random_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   332
  {
bulwahn@36013
   333
  compilation = Random,
bulwahn@36013
   334
  function_name_prefix = "random_",
bulwahn@36013
   335
  compfuns = PredicateCompFuns.compfuns,
bulwahn@36013
   336
  mk_random = (fn T => fn additional_arguments =>
bulwahn@36013
   337
  list_comb (Const(@{const_name Quickcheck.iter},
bulwahn@36013
   338
  [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}] ---> 
bulwahn@36013
   339
    PredicateCompFuns.mk_predT T), additional_arguments)),
bulwahn@36013
   340
  modify_funT = (fn T =>
bulwahn@36013
   341
    let
bulwahn@36013
   342
      val (Ts, U) = strip_type T
bulwahn@36013
   343
      val Ts' = [@{typ code_numeral}, @{typ code_numeral}, @{typ "code_numeral * code_numeral"}]
bulwahn@36013
   344
    in (Ts @ Ts') ---> U end),
bulwahn@36013
   345
  additional_arguments = (fn names =>
bulwahn@36013
   346
    let
bulwahn@36013
   347
      val [nrandom, size, seed] = Name.variant_list names ["nrandom", "size", "seed"]
bulwahn@36013
   348
    in
bulwahn@36013
   349
      [Free (nrandom, @{typ code_numeral}), Free (size, @{typ code_numeral}),
bulwahn@36013
   350
        Free (seed, @{typ "code_numeral * code_numeral"})]
bulwahn@36013
   351
    end),
bulwahn@36013
   352
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   353
    : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   354
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   355
  }
bulwahn@36013
   356
bulwahn@36013
   357
val depth_limited_random_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   358
  {
bulwahn@36013
   359
  compilation = Depth_Limited_Random,
bulwahn@36013
   360
  function_name_prefix = "depth_limited_random_",
bulwahn@36013
   361
  compfuns = PredicateCompFuns.compfuns,
bulwahn@36013
   362
  mk_random = (fn T => fn additional_arguments =>
bulwahn@36013
   363
  list_comb (Const(@{const_name Quickcheck.iter},
bulwahn@36013
   364
  [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}] ---> 
bulwahn@36013
   365
    PredicateCompFuns.mk_predT T), tl additional_arguments)),
bulwahn@36013
   366
  modify_funT = (fn T =>
bulwahn@36013
   367
    let
bulwahn@36013
   368
      val (Ts, U) = strip_type T
bulwahn@36013
   369
      val Ts' = [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
bulwahn@36013
   370
        @{typ "code_numeral * code_numeral"}]
bulwahn@36013
   371
    in (Ts @ Ts') ---> U end),
bulwahn@36013
   372
  additional_arguments = (fn names =>
bulwahn@36013
   373
    let
bulwahn@36013
   374
      val [depth, nrandom, size, seed] = Name.variant_list names ["depth", "nrandom", "size", "seed"]
bulwahn@36013
   375
    in
bulwahn@36013
   376
      [Free (depth, @{typ code_numeral}), Free (nrandom, @{typ code_numeral}),
bulwahn@36013
   377
        Free (size, @{typ code_numeral}), Free (seed, @{typ "code_numeral * code_numeral"})]
bulwahn@36013
   378
    end),
bulwahn@36013
   379
  wrap_compilation =
bulwahn@36013
   380
  fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
bulwahn@36013
   381
    let
bulwahn@36013
   382
      val depth = hd (additional_arguments)
bulwahn@36013
   383
      val (_, Ts) = split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE))
bulwahn@36013
   384
        mode (binder_types T)
bulwahn@36013
   385
      val T' = mk_predT compfuns (HOLogic.mk_tupleT Ts)
bulwahn@36013
   386
      val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
bulwahn@36013
   387
    in
bulwahn@36013
   388
      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
bulwahn@36013
   389
        $ mk_bot compfuns (dest_predT compfuns T')
bulwahn@36013
   390
        $ compilation
bulwahn@36013
   391
    end,
bulwahn@36013
   392
  transform_additional_arguments =
bulwahn@36013
   393
    fn prem => fn additional_arguments =>
bulwahn@36013
   394
    let
bulwahn@36013
   395
      val [depth, nrandom, size, seed] = additional_arguments
bulwahn@36013
   396
      val depth' =
bulwahn@36013
   397
        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
bulwahn@36013
   398
          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
bulwahn@36013
   399
    in [depth', nrandom, size, seed] end
bulwahn@36013
   400
}
bulwahn@36013
   401
bulwahn@36013
   402
val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   403
  {
bulwahn@36013
   404
  compilation = Pred,
bulwahn@36013
   405
  function_name_prefix = "",
bulwahn@36013
   406
  compfuns = PredicateCompFuns.compfuns,
bulwahn@36013
   407
  mk_random = (fn _ => error "no random generation"),
bulwahn@36013
   408
  modify_funT = I,
bulwahn@36013
   409
  additional_arguments = K [],
bulwahn@36013
   410
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   411
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   412
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   413
  }
bulwahn@36013
   414
bulwahn@36013
   415
val annotated_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   416
  {
bulwahn@36013
   417
  compilation = Annotated,
bulwahn@36013
   418
  function_name_prefix = "annotated_",
bulwahn@36013
   419
  compfuns = PredicateCompFuns.compfuns,
bulwahn@36013
   420
  mk_random = (fn _ => error "no random generation"),
bulwahn@36013
   421
  modify_funT = I,
bulwahn@36013
   422
  additional_arguments = K [],
bulwahn@36013
   423
  wrap_compilation =
bulwahn@36013
   424
    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
bulwahn@36013
   425
      mk_tracing ("calling predicate " ^ s ^
bulwahn@36013
   426
        " with mode " ^ string_of_mode mode) compilation,
bulwahn@36013
   427
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   428
  }
bulwahn@36013
   429
bulwahn@36013
   430
val dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   431
  {
bulwahn@36013
   432
  compilation = DSeq,
bulwahn@36013
   433
  function_name_prefix = "dseq_",
bulwahn@36013
   434
  compfuns = DSequence_CompFuns.compfuns,
bulwahn@40232
   435
  mk_random = (fn _ => error "no random generation in dseq"),
bulwahn@36013
   436
  modify_funT = I,
bulwahn@36013
   437
  additional_arguments = K [],
bulwahn@36013
   438
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   439
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   440
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   441
  }
bulwahn@36013
   442
bulwahn@36013
   443
val pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   444
  {
bulwahn@36013
   445
  compilation = Pos_Random_DSeq,
bulwahn@36013
   446
  function_name_prefix = "random_dseq_",
bulwahn@36013
   447
  compfuns = Random_Sequence_CompFuns.compfuns,
bulwahn@36013
   448
  mk_random = (fn T => fn additional_arguments =>
bulwahn@36013
   449
  let
bulwahn@36013
   450
    val random = Const ("Quickcheck.random_class.random",
bulwahn@36013
   451
      @{typ code_numeral} --> @{typ Random.seed} -->
bulwahn@36013
   452
        HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed}))
bulwahn@36013
   453
  in
bulwahn@36013
   454
    Const ("Random_Sequence.Random", (@{typ code_numeral} --> @{typ Random.seed} -->
bulwahn@36013
   455
      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
bulwahn@36013
   456
      Random_Sequence_CompFuns.mk_random_dseqT T) $ random
bulwahn@36013
   457
  end),
bulwahn@36013
   458
bulwahn@36013
   459
  modify_funT = I,
bulwahn@36013
   460
  additional_arguments = K [],
bulwahn@36013
   461
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   462
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   463
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   464
  }
bulwahn@36013
   465
bulwahn@36013
   466
val neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   467
  {
bulwahn@36013
   468
  compilation = Neg_Random_DSeq,
bulwahn@36013
   469
  function_name_prefix = "random_dseq_neg_",
bulwahn@36013
   470
  compfuns = Random_Sequence_CompFuns.compfuns,
bulwahn@36013
   471
  mk_random = (fn _ => error "no random generation"),
bulwahn@36013
   472
  modify_funT = I,
bulwahn@36013
   473
  additional_arguments = K [],
bulwahn@36013
   474
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   475
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   476
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   477
  }
bulwahn@36013
   478
bulwahn@36013
   479
bulwahn@36013
   480
val new_pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   481
  {
bulwahn@36013
   482
  compilation = New_Pos_Random_DSeq,
bulwahn@36013
   483
  function_name_prefix = "new_random_dseq_",
bulwahn@40230
   484
  compfuns = New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns,
bulwahn@36013
   485
  mk_random = (fn T => fn additional_arguments =>
bulwahn@36013
   486
  let
bulwahn@36013
   487
    val random = Const ("Quickcheck.random_class.random",
bulwahn@36013
   488
      @{typ code_numeral} --> @{typ Random.seed} -->
bulwahn@36013
   489
        HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed}))
bulwahn@36013
   490
  in
bulwahn@36013
   491
    Const ("New_Random_Sequence.Random", (@{typ code_numeral} --> @{typ Random.seed} -->
bulwahn@36013
   492
      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
bulwahn@36013
   493
      New_Pos_Random_Sequence_CompFuns.mk_pos_random_dseqT T) $ random
bulwahn@36013
   494
  end),
bulwahn@36013
   495
  modify_funT = I,
bulwahn@36013
   496
  additional_arguments = K [],
bulwahn@36013
   497
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   498
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   499
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   500
  }
bulwahn@36013
   501
bulwahn@36013
   502
val new_neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@36013
   503
  {
bulwahn@36013
   504
  compilation = New_Neg_Random_DSeq,
bulwahn@36013
   505
  function_name_prefix = "new_random_dseq_neg_",
bulwahn@40230
   506
  compfuns = New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns,
bulwahn@36013
   507
  mk_random = (fn _ => error "no random generation"),
bulwahn@36013
   508
  modify_funT = I,
bulwahn@36013
   509
  additional_arguments = K [],
bulwahn@36013
   510
  wrap_compilation = K (K (K (K (K I))))
bulwahn@36013
   511
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@36013
   512
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@36013
   513
  }
bulwahn@36013
   514
bulwahn@40232
   515
val pos_generator_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@40232
   516
  {
bulwahn@40232
   517
  compilation = Pos_Generator_DSeq,
bulwahn@40232
   518
  function_name_prefix = "generator_dseq_",
bulwahn@40232
   519
  compfuns = New_Pos_DSequence_CompFuns.depth_limited_compfuns,
bulwahn@40232
   520
  mk_random =
bulwahn@40232
   521
    (fn T => fn additional_arguments =>
bulwahn@40232
   522
      let
bulwahn@40232
   523
        val small_lazy =
bulwahn@40232
   524
         Const (@{const_name "Lazy_Sequence.small_lazy_class.small_lazy"},
bulwahn@40232
   525
         @{typ "Int.int"} --> Type (@{type_name "Lazy_Sequence.lazy_sequence"}, [T])) 
bulwahn@40232
   526
      in
bulwahn@40232
   527
        absdummy (@{typ code_numeral}, small_lazy $ HOLogic.mk_number @{typ int} 3)
bulwahn@40232
   528
      end
bulwahn@40232
   529
    ),
bulwahn@40232
   530
  modify_funT = I,
bulwahn@40232
   531
  additional_arguments = K [],
bulwahn@40232
   532
  wrap_compilation = K (K (K (K (K I))))
bulwahn@40232
   533
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@40232
   534
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@40232
   535
  }
bulwahn@40232
   536
bulwahn@40232
   537
val neg_generator_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
bulwahn@40232
   538
  {
bulwahn@40232
   539
  compilation = Neg_Generator_DSeq,
bulwahn@40232
   540
  function_name_prefix = "generator_dseq_neg_",
bulwahn@40232
   541
  compfuns = New_Neg_DSequence_CompFuns.depth_limited_compfuns,
bulwahn@40232
   542
  mk_random = (fn _ => error "no random generation"),
bulwahn@40232
   543
  modify_funT = I,
bulwahn@40232
   544
  additional_arguments = K [],
bulwahn@40232
   545
  wrap_compilation = K (K (K (K (K I))))
bulwahn@40232
   546
   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
bulwahn@40232
   547
  transform_additional_arguments = K I : (indprem -> term list -> term list)
bulwahn@40232
   548
  }
bulwahn@40232
   549
bulwahn@36013
   550
fun negative_comp_modifiers_of comp_modifiers =
bulwahn@36013
   551
    (case Comp_Mod.compilation comp_modifiers of
bulwahn@36013
   552
      Pos_Random_DSeq => neg_random_dseq_comp_modifiers
bulwahn@36013
   553
    | Neg_Random_DSeq => pos_random_dseq_comp_modifiers
bulwahn@36013
   554
    | New_Pos_Random_DSeq => new_neg_random_dseq_comp_modifiers
bulwahn@40232
   555
    | New_Neg_Random_DSeq => new_pos_random_dseq_comp_modifiers 
bulwahn@40232
   556
    | Pos_Generator_DSeq => neg_generator_dseq_comp_modifiers
bulwahn@40232
   557
    | Neg_Generator_DSeq => pos_generator_dseq_comp_modifiers
bulwahn@36013
   558
    | c => comp_modifiers)
bulwahn@36013
   559
bulwahn@32667
   560
(* term construction *)
bulwahn@32667
   561
bulwahn@32667
   562
fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
bulwahn@32667
   563
      NONE => (Free (s, T), (names, (s, [])::vs))
bulwahn@32667
   564
    | SOME xs =>
bulwahn@32667
   565
        let
wenzelm@44206
   566
          val s' = singleton (Name.variant_list names) s;
bulwahn@32667
   567
          val v = Free (s', T)
bulwahn@32667
   568
        in
bulwahn@32667
   569
          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
bulwahn@32667
   570
        end);
bulwahn@32667
   571
bulwahn@32667
   572
fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
bulwahn@32667
   573
  | distinct_v (t $ u) nvs =
bulwahn@32667
   574
      let
bulwahn@32667
   575
        val (t', nvs') = distinct_v t nvs;
bulwahn@32667
   576
        val (u', nvs'') = distinct_v u nvs';
bulwahn@32667
   577
      in (t' $ u', nvs'') end
bulwahn@32667
   578
  | distinct_v x nvs = (x, nvs);
bulwahn@32667
   579
bulwahn@39872
   580
bulwahn@33147
   581
(** specific rpred functions -- move them to the correct place in this file *)
bulwahn@39872
   582
fun mk_Eval_of (P as (Free (f, _)), T) mode =
bulwahn@39872
   583
let
bulwahn@39872
   584
  fun mk_bounds (Type (@{type_name Product_Type.prod}, [T1, T2])) i =
bulwahn@39872
   585
    let
bulwahn@39872
   586
      val (bs2, i') = mk_bounds T2 i 
bulwahn@39872
   587
      val (bs1, i'') = mk_bounds T1 i'
bulwahn@39872
   588
    in
bulwahn@39872
   589
      (HOLogic.pair_const T1 T2 $ bs1 $ bs2, i'' + 1)
bulwahn@39872
   590
    end
bulwahn@39872
   591
    | mk_bounds T i = (Bound i, i + 1)
bulwahn@39872
   592
  fun mk_prod ((t1, T1), (t2, T2)) = (HOLogic.pair_const T1 T2 $ t1 $ t2, HOLogic.mk_prodT (T1, T2))
bulwahn@39872
   593
  fun mk_tuple [] = (HOLogic.unit, HOLogic.unitT)
bulwahn@39872
   594
    | mk_tuple tTs = foldr1 mk_prod tTs;
bulwahn@39872
   595
  fun mk_split_abs (T as Type (@{type_name Product_Type.prod}, [T1, T2])) t = absdummy (T, HOLogic.split_const (T1, T2, @{typ bool}) $ (mk_split_abs T1 (mk_split_abs T2 t)))
bulwahn@39872
   596
    | mk_split_abs T t = absdummy (T, t)
bulwahn@39872
   597
  val args = rev (fst (fold_map mk_bounds (rev (binder_types T)) 0))
bulwahn@39872
   598
  val (inargs, outargs) = split_mode mode args
bulwahn@39872
   599
  val (inTs, outTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE)) mode (binder_types T)
bulwahn@39872
   600
  val inner_term = PredicateCompFuns.mk_Eval (list_comb (P, inargs), fst (mk_tuple (outargs ~~ outTs)))
bulwahn@39872
   601
in
bulwahn@39872
   602
  fold_rev mk_split_abs (binder_types T) inner_term  
bulwahn@39872
   603
end
bulwahn@33147
   604
bulwahn@39872
   605
fun compile_arg compilation_modifiers additional_arguments ctxt param_modes arg = 
bulwahn@33147
   606
  let
bulwahn@33147
   607
    fun map_params (t as Free (f, T)) =
bulwahn@39872
   608
      (case (AList.lookup (op =) param_modes f) of
bulwahn@39872
   609
          SOME mode =>
bulwahn@33147
   610
            let
bulwahn@39872
   611
              val T' = Comp_Mod.funT_of compilation_modifiers mode T
bulwahn@39872
   612
            in
bulwahn@39872
   613
              mk_Eval_of (Free (f, T'), T) mode
bulwahn@39872
   614
            end
bulwahn@39872
   615
        | NONE => t)
bulwahn@33147
   616
      | map_params t = t
bulwahn@39872
   617
  in
bulwahn@39872
   618
    map_aterms map_params arg
bulwahn@39872
   619
  end
bulwahn@33147
   620
bulwahn@39872
   621
fun compile_match compilation_modifiers additional_arguments ctxt param_modes
bulwahn@39872
   622
      eqs eqs' out_ts success_t =
bulwahn@32667
   623
  let
bulwahn@36013
   624
    val compfuns = Comp_Mod.compfuns compilation_modifiers
bulwahn@32667
   625
    val eqs'' = maps mk_eq eqs @ eqs'
bulwahn@33147
   626
    val eqs'' =
bulwahn@39872
   627
      map (compile_arg compilation_modifiers additional_arguments ctxt param_modes) eqs''
bulwahn@32667
   628
    val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
wenzelm@44206
   629
    val name = singleton (Name.variant_list names) "x";
wenzelm@44206
   630
    val name' = singleton (Name.variant_list (name :: names)) "y";
bulwahn@33148
   631
    val T = HOLogic.mk_tupleT (map fastype_of out_ts);
bulwahn@32667
   632
    val U = fastype_of success_t;
bulwahn@36254
   633
    val U' = dest_predT compfuns U;
bulwahn@32667
   634
    val v = Free (name, T);
bulwahn@32667
   635
    val v' = Free (name', T);
bulwahn@32667
   636
  in
krauss@44094
   637
    lambda v (Datatype.make_case ctxt Datatype_Case.Quiet [] v
bulwahn@33148
   638
      [(HOLogic.mk_tuple out_ts,
bulwahn@32667
   639
        if null eqs'' then success_t
bulwahn@32667
   640
        else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
bulwahn@32667
   641
          foldr1 HOLogic.mk_conj eqs'' $ success_t $
bulwahn@32667
   642
            mk_bot compfuns U'),
krauss@44094
   643
       (v', mk_bot compfuns U')])
bulwahn@32667
   644
  end;
bulwahn@32667
   645
bulwahn@35891
   646
fun string_of_tderiv ctxt (t, deriv) = 
bulwahn@35324
   647
  (case (t, deriv) of
bulwahn@35324
   648
    (t1 $ t2, Mode_App (deriv1, deriv2)) =>
bulwahn@35891
   649
      string_of_tderiv ctxt (t1, deriv1) ^ " $ " ^ string_of_tderiv ctxt (t2, deriv2)
haftmann@37366
   650
  | (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (deriv1, deriv2)) =>
bulwahn@35891
   651
    "(" ^ string_of_tderiv ctxt (t1, deriv1) ^ ", " ^ string_of_tderiv ctxt (t2, deriv2) ^ ")"
bulwahn@35891
   652
  | (t, Term Input) => Syntax.string_of_term ctxt t ^ "[Input]"
bulwahn@35891
   653
  | (t, Term Output) => Syntax.string_of_term ctxt t ^ "[Output]"
bulwahn@35891
   654
  | (t, Context m) => Syntax.string_of_term ctxt t ^ "[" ^ string_of_mode m ^ "]")
bulwahn@35324
   655
bulwahn@39872
   656
fun compile_expr compilation_modifiers ctxt (t, deriv) param_modes additional_arguments =
bulwahn@32667
   657
  let
bulwahn@36013
   658
    val compfuns = Comp_Mod.compfuns compilation_modifiers
bulwahn@34935
   659
    fun expr_of (t, deriv) =
bulwahn@34935
   660
      (case (t, deriv) of
bulwahn@39872
   661
        (t, Term Input) => SOME (compile_arg compilation_modifiers additional_arguments ctxt param_modes t)
bulwahn@34935
   662
      | (t, Term Output) => NONE
bulwahn@34935
   663
      | (Const (name, T), Context mode) =>
bulwahn@36994
   664
        (case alternative_compilation_of ctxt name mode of
bulwahn@36032
   665
          SOME alt_comp => SOME (alt_comp compfuns T)
bulwahn@36028
   666
        | NONE =>
bulwahn@36028
   667
          SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers)
bulwahn@36994
   668
            ctxt name mode,
bulwahn@36028
   669
            Comp_Mod.funT_of compilation_modifiers mode T)))
bulwahn@34935
   670
      | (Free (s, T), Context m) =>
bulwahn@40018
   671
        (case (AList.lookup (op =) param_modes s) of
bulwahn@40018
   672
          SOME mode => SOME (Free (s, Comp_Mod.funT_of compilation_modifiers m T))
bulwahn@40018
   673
        | NONE =>
bulwahn@40018
   674
        let
bulwahn@40018
   675
          val bs = map (pair "x") (binder_types (fastype_of t))
bulwahn@40018
   676
          val bounds = map Bound (rev (0 upto (length bs) - 1))
bulwahn@40018
   677
        in SOME (list_abs (bs, mk_if compfuns (list_comb (t, bounds)))) end)
bulwahn@34935
   678
      | (t, Context m) =>
bulwahn@34935
   679
        let
bulwahn@34935
   680
          val bs = map (pair "x") (binder_types (fastype_of t))
bulwahn@34935
   681
          val bounds = map Bound (rev (0 upto (length bs) - 1))
bulwahn@34935
   682
        in SOME (list_abs (bs, mk_if compfuns (list_comb (t, bounds)))) end
haftmann@37366
   683
      | (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =>
bulwahn@34935
   684
        (case (expr_of (t1, d1), expr_of (t2, d2)) of
bulwahn@34935
   685
          (NONE, NONE) => NONE
bulwahn@34935
   686
        | (NONE, SOME t) => SOME t
bulwahn@34935
   687
        | (SOME t, NONE) => SOME t
bulwahn@34935
   688
        | (SOME t1, SOME t2) => SOME (HOLogic.mk_prod (t1, t2)))
bulwahn@34935
   689
      | (t1 $ t2, Mode_App (deriv1, deriv2)) =>
bulwahn@34935
   690
        (case (expr_of (t1, deriv1), expr_of (t2, deriv2)) of
bulwahn@34935
   691
          (SOME t, NONE) => SOME t
bulwahn@34935
   692
         | (SOME t, SOME u) => SOME (t $ u)
bulwahn@34935
   693
         | _ => error "something went wrong here!"))
bulwahn@32667
   694
  in
bulwahn@35879
   695
    list_comb (the (expr_of (t, deriv)), additional_arguments)
bulwahn@34935
   696
  end
bulwahn@33145
   697
bulwahn@39872
   698
fun compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments
bulwahn@39872
   699
  inp (in_ts, out_ts) moded_ps =
bulwahn@32667
   700
  let
bulwahn@36013
   701
    val compfuns = Comp_Mod.compfuns compilation_modifiers
bulwahn@36013
   702
    val compile_match = compile_match compilation_modifiers
bulwahn@39872
   703
      additional_arguments ctxt param_modes
bulwahn@32667
   704
    val (in_ts', (all_vs', eqs)) =
bulwahn@35891
   705
      fold_map (collect_non_invertible_subterms ctxt) in_ts (all_vs, []);
bulwahn@32667
   706
    fun compile_prems out_ts' vs names [] =
bulwahn@32667
   707
          let
bulwahn@32667
   708
            val (out_ts'', (names', eqs')) =
bulwahn@35891
   709
              fold_map (collect_non_invertible_subterms ctxt) out_ts' (names, []);
bulwahn@32667
   710
            val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
bulwahn@32667
   711
              out_ts'' (names', map (rpair []) vs);
bulwahn@39995
   712
            val processed_out_ts = map (compile_arg compilation_modifiers additional_arguments
bulwahn@39995
   713
              ctxt param_modes) out_ts
bulwahn@32667
   714
          in
bulwahn@33147
   715
            compile_match constr_vs (eqs @ eqs') out_ts'''
bulwahn@39995
   716
              (mk_single compfuns (HOLogic.mk_tuple processed_out_ts))
bulwahn@32667
   717
          end
bulwahn@34935
   718
      | compile_prems out_ts vs names ((p, deriv) :: ps) =
bulwahn@32667
   719
          let
bulwahn@32667
   720
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
bulwahn@32667
   721
            val (out_ts', (names', eqs)) =
bulwahn@35891
   722
              fold_map (collect_non_invertible_subterms ctxt) out_ts (names, [])
bulwahn@32667
   723
            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
bulwahn@32667
   724
              out_ts' ((names', map (rpair []) vs))
bulwahn@34935
   725
            val mode = head_mode_of deriv
bulwahn@33143
   726
            val additional_arguments' =
bulwahn@33330
   727
              Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments
bulwahn@32667
   728
            val (compiled_clause, rest) = case p of
bulwahn@34935
   729
               Prem t =>
bulwahn@32667
   730
                 let
bulwahn@33138
   731
                   val u =
bulwahn@39872
   732
                     compile_expr compilation_modifiers ctxt (t, deriv) param_modes additional_arguments'
bulwahn@34935
   733
                   val (_, out_ts''') = split_mode mode (snd (strip_comb t))
bulwahn@32667
   734
                   val rest = compile_prems out_ts''' vs' names'' ps
bulwahn@32667
   735
                 in
bulwahn@32667
   736
                   (u, rest)
bulwahn@32667
   737
                 end
bulwahn@34935
   738
             | Negprem t =>
bulwahn@32667
   739
                 let
bulwahn@36013
   740
                   val neg_compilation_modifiers =
bulwahn@36013
   741
                     negative_comp_modifiers_of compilation_modifiers
bulwahn@33143
   742
                   val u = mk_not compfuns
bulwahn@39872
   743
                     (compile_expr neg_compilation_modifiers ctxt (t, deriv) param_modes additional_arguments')
bulwahn@34935
   744
                   val (_, out_ts''') = split_mode mode (snd (strip_comb t))
bulwahn@32667
   745
                   val rest = compile_prems out_ts''' vs' names'' ps
bulwahn@32667
   746
                 in
bulwahn@32667
   747
                   (u, rest)
bulwahn@32667
   748
                 end
bulwahn@32667
   749
             | Sidecond t =>
bulwahn@32667
   750
                 let
bulwahn@36013
   751
                   val t = compile_arg compilation_modifiers additional_arguments
bulwahn@39872
   752
                     ctxt param_modes t
bulwahn@32667
   753
                   val rest = compile_prems [] vs' names'' ps;
bulwahn@32667
   754
                 in
bulwahn@32667
   755
                   (mk_if compfuns t, rest)
bulwahn@32667
   756
                 end
bulwahn@32667
   757
             | Generator (v, T) =>
bulwahn@32667
   758
                 let
bulwahn@35880
   759
                   val u = Comp_Mod.mk_random compilation_modifiers T additional_arguments
bulwahn@32667
   760
                   val rest = compile_prems [Free (v, T)]  vs' names'' ps;
bulwahn@32667
   761
                 in
bulwahn@32667
   762
                   (u, rest)
bulwahn@32667
   763
                 end
bulwahn@32667
   764
          in
bulwahn@33147
   765
            compile_match constr_vs' eqs out_ts''
bulwahn@32667
   766
              (mk_bind compfuns (compiled_clause, rest))
bulwahn@32667
   767
          end
bulwahn@39872
   768
    val prem_t = compile_prems in_ts' (map fst param_modes) all_vs' moded_ps;
bulwahn@32667
   769
  in
bulwahn@32667
   770
    mk_bind compfuns (mk_single compfuns inp, prem_t)
bulwahn@32667
   771
  end
bulwahn@32667
   772
bulwahn@36254
   773
(* switch detection *)
bulwahn@36254
   774
bulwahn@36254
   775
(** argument position of an inductive predicates and the executable functions **)
bulwahn@36254
   776
bulwahn@36254
   777
type position = int * int list
bulwahn@36254
   778
bulwahn@36254
   779
fun input_positions_pair Input = [[]]
bulwahn@36254
   780
  | input_positions_pair Output = []
bulwahn@36254
   781
  | input_positions_pair (Fun _) = []
bulwahn@36254
   782
  | input_positions_pair (Pair (m1, m2)) =
bulwahn@36254
   783
    map (cons 1) (input_positions_pair m1) @ map (cons 2) (input_positions_pair m2)
bulwahn@36254
   784
bulwahn@36254
   785
fun input_positions_of_mode mode = flat (map_index
bulwahn@36254
   786
   (fn (i, Input) => [(i, [])]
bulwahn@36254
   787
   | (_, Output) => []
bulwahn@36254
   788
   | (_, Fun _) => []
bulwahn@36254
   789
   | (i, m as Pair (m1, m2)) => map (pair i) (input_positions_pair m))
bulwahn@36254
   790
     (Predicate_Compile_Aux.strip_fun_mode mode))
bulwahn@36254
   791
bulwahn@36254
   792
fun argument_position_pair mode [] = []
bulwahn@36254
   793
  | argument_position_pair (Pair (Fun _, m2)) (2 :: is) = argument_position_pair m2 is
bulwahn@36254
   794
  | argument_position_pair (Pair (m1, m2)) (i :: is) =
bulwahn@36254
   795
    (if eq_mode (m1, Output) andalso i = 2 then
bulwahn@36254
   796
      argument_position_pair m2 is
bulwahn@36254
   797
    else if eq_mode (m2, Output) andalso i = 1 then
bulwahn@36254
   798
      argument_position_pair m1 is
bulwahn@36254
   799
    else (i :: argument_position_pair (if i = 1 then m1 else m2) is))
bulwahn@36254
   800
bulwahn@36254
   801
fun argument_position_of mode (i, is) =
bulwahn@36254
   802
  (i - (length (filter (fn Output => true | Fun _ => true | _ => false)
bulwahn@36254
   803
    (List.take (strip_fun_mode mode, i)))),
bulwahn@36254
   804
  argument_position_pair (nth (strip_fun_mode mode) i) is)
bulwahn@36254
   805
bulwahn@36254
   806
fun nth_pair [] t = t
bulwahn@36254
   807
  | nth_pair (1 :: is) (Const (@{const_name Pair}, _) $ t1 $ _) = nth_pair is t1
bulwahn@36254
   808
  | nth_pair (2 :: is) (Const (@{const_name Pair}, _) $ _ $ t2) = nth_pair is t2
bulwahn@36254
   809
  | nth_pair _ _ = raise Fail "unexpected input for nth_tuple"
bulwahn@36254
   810
bulwahn@36254
   811
(** switch detection analysis **)
bulwahn@36254
   812
bulwahn@36996
   813
fun find_switch_test ctxt (i, is) (ts, prems) =
bulwahn@36254
   814
  let
bulwahn@36254
   815
    val t = nth_pair is (nth ts i)
bulwahn@36254
   816
    val T = fastype_of t
bulwahn@36254
   817
  in
bulwahn@36254
   818
    case T of
bulwahn@36254
   819
      TFree _ => NONE
bulwahn@36254
   820
    | Type (Tcon, _) =>
wenzelm@43232
   821
      (case Datatype_Data.get_constrs (Proof_Context.theory_of ctxt) Tcon of
bulwahn@36254
   822
        NONE => NONE
bulwahn@36254
   823
      | SOME cs =>
bulwahn@36254
   824
        (case strip_comb t of
bulwahn@36254
   825
          (Var _, []) => NONE
bulwahn@36254
   826
        | (Free _, []) => NONE
bulwahn@36254
   827
        | (Const (c, T), _) => if AList.defined (op =) cs c then SOME (c, T) else NONE))
bulwahn@36254
   828
  end
bulwahn@36254
   829
bulwahn@36996
   830
fun partition_clause ctxt pos moded_clauses =
bulwahn@36254
   831
  let
bulwahn@36254
   832
    fun insert_list eq (key, value) = AList.map_default eq (key, []) (cons value)
bulwahn@36254
   833
    fun find_switch_test' moded_clause (cases, left) =
bulwahn@36996
   834
      case find_switch_test ctxt pos moded_clause of
bulwahn@36254
   835
        SOME (c, T) => (insert_list (op =) ((c, T), moded_clause) cases, left)
bulwahn@36254
   836
      | NONE => (cases, moded_clause :: left)
bulwahn@36254
   837
  in
bulwahn@36254
   838
    fold find_switch_test' moded_clauses ([], [])
bulwahn@36254
   839
  end
bulwahn@36254
   840
bulwahn@36254
   841
datatype switch_tree =
bulwahn@36254
   842
  Atom of moded_clause list | Node of (position * ((string * typ) * switch_tree) list) * switch_tree
bulwahn@36254
   843
bulwahn@36996
   844
fun mk_switch_tree ctxt mode moded_clauses =
bulwahn@36254
   845
  let
bulwahn@36254
   846
    fun select_best_switch moded_clauses input_position best_switch =
bulwahn@36254
   847
      let
bulwahn@36254
   848
        val ord = option_ord (rev_order o int_ord o (pairself (length o snd o snd)))
bulwahn@36996
   849
        val partition = partition_clause ctxt input_position moded_clauses
bulwahn@36254
   850
        val switch = if (length (fst partition) > 1) then SOME (input_position, partition) else NONE
bulwahn@36254
   851
      in
bulwahn@36254
   852
        case ord (switch, best_switch) of LESS => best_switch
bulwahn@36254
   853
          | EQUAL => best_switch | GREATER => switch
bulwahn@36254
   854
      end
bulwahn@36254
   855
    fun detect_switches moded_clauses =
bulwahn@36254
   856
      case fold (select_best_switch moded_clauses) (input_positions_of_mode mode) NONE of
bulwahn@36254
   857
        SOME (best_pos, (switched_on, left_clauses)) =>
bulwahn@36254
   858
          Node ((best_pos, map (apsnd detect_switches) switched_on),
bulwahn@36254
   859
            detect_switches left_clauses)
bulwahn@36254
   860
      | NONE => Atom moded_clauses
bulwahn@36254
   861
  in
bulwahn@36254
   862
    detect_switches moded_clauses
bulwahn@36254
   863
  end
bulwahn@36254
   864
bulwahn@36254
   865
(** compilation of detected switches **)
bulwahn@36254
   866
bulwahn@36254
   867
fun destruct_constructor_pattern (pat, obj) =
bulwahn@36254
   868
  (case strip_comb pat of
bulwahn@36254
   869
    (f as Free _, []) => cons (pat, obj)
bulwahn@36254
   870
  | (Const (c, T), pat_args) =>
bulwahn@36254
   871
    (case strip_comb obj of
bulwahn@36254
   872
      (Const (c', T'), obj_args) =>
bulwahn@36254
   873
        (if c = c' andalso T = T' then
bulwahn@36254
   874
          fold destruct_constructor_pattern (pat_args ~~ obj_args)
bulwahn@36254
   875
        else raise Fail "pattern and object mismatch")
bulwahn@36254
   876
    | _ => raise Fail "unexpected object")
bulwahn@36254
   877
  | _ => raise Fail "unexpected pattern")
bulwahn@36254
   878
bulwahn@36254
   879
bulwahn@39872
   880
fun compile_switch compilation_modifiers ctxt all_vs param_modes additional_arguments mode
bulwahn@36254
   881
  in_ts' outTs switch_tree =
bulwahn@36254
   882
  let
bulwahn@36254
   883
    val compfuns = Comp_Mod.compfuns compilation_modifiers
wenzelm@43232
   884
    val thy = Proof_Context.theory_of ctxt
bulwahn@36254
   885
    fun compile_switch_tree _ _ (Atom []) = NONE
bulwahn@36254
   886
      | compile_switch_tree all_vs ctxt_eqs (Atom moded_clauses) =
bulwahn@36254
   887
        let
bulwahn@36254
   888
          val in_ts' = map (Pattern.rewrite_term thy ctxt_eqs []) in_ts'
bulwahn@36254
   889
          fun compile_clause' (ts, moded_ps) =
bulwahn@36254
   890
            let
bulwahn@36254
   891
              val (ts, out_ts) = split_mode mode ts
bulwahn@36254
   892
              val subst = fold destruct_constructor_pattern (in_ts' ~~ ts) []
bulwahn@36254
   893
              val (fsubst, pat') = List.partition (fn (_, Free _) => true | _ => false) subst
bulwahn@36254
   894
              val moded_ps' = (map o apfst o map_indprem)
bulwahn@36254
   895
                (Pattern.rewrite_term thy (map swap fsubst) []) moded_ps
bulwahn@36254
   896
              val inp = HOLogic.mk_tuple (map fst pat')
bulwahn@36254
   897
              val in_ts' = map (Pattern.rewrite_term thy (map swap fsubst) []) (map snd pat')
bulwahn@36254
   898
              val out_ts' = map (Pattern.rewrite_term thy (map swap fsubst) []) out_ts
bulwahn@36254
   899
            in
bulwahn@39872
   900
              compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments
bulwahn@39872
   901
                inp (in_ts', out_ts') moded_ps'
bulwahn@36254
   902
            end
bulwahn@36254
   903
        in SOME (foldr1 (mk_sup compfuns) (map compile_clause' moded_clauses)) end
bulwahn@36254
   904
    | compile_switch_tree all_vs ctxt_eqs (Node ((position, switched_clauses), left_clauses)) =
bulwahn@36254
   905
      let
bulwahn@36254
   906
        val (i, is) = argument_position_of mode position
bulwahn@36254
   907
        val inp_var = nth_pair is (nth in_ts' i)
wenzelm@44206
   908
        val x = singleton (Name.variant_list all_vs) "x"
bulwahn@36254
   909
        val xt = Free (x, fastype_of inp_var)
bulwahn@36254
   910
        fun compile_single_case ((c, T), switched) =
bulwahn@36254
   911
          let
bulwahn@36254
   912
            val Ts = binder_types T
bulwahn@36254
   913
            val argnames = Name.variant_list (x :: all_vs)
bulwahn@36254
   914
              (map (fn i => "c" ^ string_of_int i) (1 upto length Ts))
bulwahn@36254
   915
            val args = map2 (curry Free) argnames Ts
bulwahn@36254
   916
            val pattern = list_comb (Const (c, T), args)
bulwahn@36254
   917
            val ctxt_eqs' = (inp_var, pattern) :: ctxt_eqs
bulwahn@36254
   918
            val compilation = the_default (mk_bot compfuns (HOLogic.mk_tupleT outTs))
bulwahn@36254
   919
              (compile_switch_tree (argnames @ x :: all_vs) ctxt_eqs' switched)
bulwahn@36254
   920
        in
bulwahn@36254
   921
          (pattern, compilation)
bulwahn@36254
   922
        end
krauss@44094
   923
        val switch = Datatype.make_case ctxt Datatype_Case.Quiet [] inp_var
bulwahn@36254
   924
          ((map compile_single_case switched_clauses) @
krauss@44094
   925
            [(xt, mk_bot compfuns (HOLogic.mk_tupleT outTs))])
bulwahn@36254
   926
      in
bulwahn@36254
   927
        case compile_switch_tree all_vs ctxt_eqs left_clauses of
bulwahn@36254
   928
          NONE => SOME switch
bulwahn@36254
   929
        | SOME left_comp => SOME (mk_sup compfuns (switch, left_comp))
bulwahn@36254
   930
      end
bulwahn@36254
   931
  in
bulwahn@36254
   932
    compile_switch_tree all_vs [] switch_tree
bulwahn@36254
   933
  end
bulwahn@36254
   934
bulwahn@36254
   935
(* compilation of predicates *)
bulwahn@36254
   936
bulwahn@36996
   937
fun compile_pred options compilation_modifiers ctxt all_vs param_vs s T (pol, mode) moded_cls =
bulwahn@32667
   938
  let
bulwahn@40231
   939
    val is_terminating = false (* FIXME: requires an termination analysis *)  
bulwahn@40230
   940
    val compilation_modifiers =
bulwahn@40230
   941
      (if pol then compilation_modifiers else
bulwahn@40230
   942
        negative_comp_modifiers_of compilation_modifiers)
bulwahn@40230
   943
      |> (if is_depth_limited_compilation (Comp_Mod.compilation compilation_modifiers) then
bulwahn@40230
   944
           (if is_terminating then
bulwahn@40230
   945
             (Comp_Mod.set_compfuns (unlimited_compfuns_of pol (Comp_Mod.compilation compilation_modifiers)))
bulwahn@40230
   946
           else
bulwahn@40230
   947
             (Comp_Mod.set_compfuns (limited_compfuns_of pol (Comp_Mod.compilation compilation_modifiers))))
bulwahn@40230
   948
         else I)
bulwahn@35879
   949
    val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers
bulwahn@33482
   950
      (all_vs @ param_vs)
bulwahn@34935
   951
    val compfuns = Comp_Mod.compfuns compilation_modifiers
bulwahn@34935
   952
    fun is_param_type (T as Type ("fun",[_ , T'])) =
bulwahn@34935
   953
      is_some (try (dest_predT compfuns) T) orelse is_param_type T'
bulwahn@34935
   954
      | is_param_type T = is_some (try (dest_predT compfuns) T)
bulwahn@34935
   955
    val (inpTs, outTs) = split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE)) mode
bulwahn@34935
   956
      (binder_types T)
bulwahn@34935
   957
    val predT = mk_predT compfuns (HOLogic.mk_tupleT outTs)
bulwahn@34935
   958
    val funT = Comp_Mod.funT_of compilation_modifiers mode T
bulwahn@34935
   959
    val (in_ts, _) = fold_map (fold_map_aterms_prodT (curry HOLogic.mk_prod)
bulwahn@34935
   960
      (fn T => fn (param_vs, names) =>
bulwahn@36012
   961
        if is_param_type T then
bulwahn@34935
   962
          (Free (hd param_vs, T), (tl param_vs, names))
bulwahn@34935
   963
        else
bulwahn@34935
   964
          let
wenzelm@44206
   965
            val new = singleton (Name.variant_list names) "x"
bulwahn@34935
   966
          in (Free (new, T), (param_vs, new :: names)) end)) inpTs
bulwahn@34935
   967
        (param_vs, (all_vs @ param_vs))
bulwahn@34935
   968
    val in_ts' = map_filter (map_filter_prod
bulwahn@34935
   969
      (fn t as Free (x, _) => if member (op =) param_vs x then NONE else SOME t | t => SOME t)) in_ts
bulwahn@39872
   970
    val param_modes = param_vs ~~ ho_arg_modes_of mode
bulwahn@36254
   971
    val compilation =
bulwahn@36254
   972
      if detect_switches options then
bulwahn@36254
   973
        the_default (mk_bot compfuns (HOLogic.mk_tupleT outTs))
bulwahn@39872
   974
          (compile_switch compilation_modifiers ctxt all_vs param_modes additional_arguments mode
bulwahn@39872
   975
            in_ts' outTs (mk_switch_tree ctxt mode moded_cls))
bulwahn@36254
   976
      else
bulwahn@36254
   977
        let
bulwahn@36254
   978
          val cl_ts =
bulwahn@36254
   979
            map (fn (ts, moded_prems) => 
bulwahn@39872
   980
              compile_clause compilation_modifiers ctxt all_vs param_modes additional_arguments
bulwahn@39872
   981
                (HOLogic.mk_tuple in_ts') (split_mode mode ts) moded_prems) moded_cls;
bulwahn@36254
   982
        in
bulwahn@36254
   983
          Comp_Mod.wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
bulwahn@36254
   984
            (if null cl_ts then
bulwahn@36254
   985
              mk_bot compfuns (HOLogic.mk_tupleT outTs)
bulwahn@36254
   986
            else
bulwahn@36254
   987
              foldr1 (mk_sup compfuns) cl_ts)
bulwahn@36254
   988
        end
bulwahn@33143
   989
    val fun_const =
bulwahn@35891
   990
      Const (function_name_of (Comp_Mod.compilation compilation_modifiers)
bulwahn@36994
   991
      ctxt s mode, funT)
bulwahn@32667
   992
  in
bulwahn@33143
   993
    HOLogic.mk_Trueprop
bulwahn@34935
   994
      (HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation))
bulwahn@32667
   995
  end;
bulwahn@33143
   996
bulwahn@32667
   997
(* Definition of executable functions and their intro and elim rules *)
bulwahn@32667
   998
bulwahn@32667
   999
fun print_arities arities = tracing ("Arities:\n" ^
bulwahn@32667
  1000
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
bulwahn@32667
  1001
    space_implode " -> " (map
bulwahn@32667
  1002
      (fn NONE => "X" | SOME k' => string_of_int k')
bulwahn@32667
  1003
        (ks @ [SOME k]))) arities));
bulwahn@32667
  1004
haftmann@37591
  1005
fun strip_split_abs (Const (@{const_name prod_case}, _) $ t) = strip_split_abs t
bulwahn@34935
  1006
  | strip_split_abs (Abs (_, _, t)) = strip_split_abs t
bulwahn@34935
  1007
  | strip_split_abs t = t
bulwahn@34935
  1008
haftmann@37678
  1009
fun mk_args is_eval (m as Pair (m1, m2), T as Type (@{type_name Product_Type.prod}, [T1, T2])) names =
bulwahn@35324
  1010
    if eq_mode (m, Input) orelse eq_mode (m, Output) then
bulwahn@35324
  1011
      let
wenzelm@44206
  1012
        val x = singleton (Name.variant_list names) "x"
bulwahn@35324
  1013
      in
bulwahn@35324
  1014
        (Free (x, T), x :: names)
bulwahn@35324
  1015
      end
bulwahn@35324
  1016
    else
bulwahn@35324
  1017
      let
bulwahn@35324
  1018
        val (t1, names') = mk_args is_eval (m1, T1) names
bulwahn@35324
  1019
        val (t2, names'') = mk_args is_eval (m2, T2) names'
bulwahn@35324
  1020
      in
bulwahn@35324
  1021
        (HOLogic.mk_prod (t1, t2), names'')
bulwahn@35324
  1022
      end
bulwahn@34935
  1023
  | mk_args is_eval ((m as Fun _), T) names =
bulwahn@34935
  1024
    let
bulwahn@34935
  1025
      val funT = funT_of PredicateCompFuns.compfuns m T
wenzelm@44206
  1026
      val x = singleton (Name.variant_list names) "x"
bulwahn@34935
  1027
      val (args, _) = fold_map (mk_args is_eval) (strip_fun_mode m ~~ binder_types T) (x :: names)
bulwahn@34935
  1028
      val (inargs, outargs) = split_map_mode (fn _ => fn t => (SOME t, NONE)) m args
krauss@39989
  1029
      val t = fold_rev HOLogic.tupled_lambda args (PredicateCompFuns.mk_Eval
bulwahn@34935
  1030
        (list_comb (Free (x, funT), inargs), HOLogic.mk_tuple outargs))
bulwahn@34935
  1031
    in
bulwahn@34935
  1032
      (if is_eval then t else Free (x, funT), x :: names)
bulwahn@34935
  1033
    end
bulwahn@34935
  1034
  | mk_args is_eval (_, T) names =
bulwahn@34935
  1035
    let
wenzelm@44206
  1036
      val x = singleton (Name.variant_list names) "x"
bulwahn@34935
  1037
    in
bulwahn@34935
  1038
      (Free (x, T), x :: names)
bulwahn@34935
  1039
    end
bulwahn@34935
  1040
bulwahn@36994
  1041
fun create_intro_elim_rule ctxt mode defthm mode_id funT pred =
bulwahn@34935
  1042
  let
bulwahn@34935
  1043
    val funtrm = Const (mode_id, funT)
bulwahn@34935
  1044
    val Ts = binder_types (fastype_of pred)
bulwahn@34935
  1045
    val (args, argnames) = fold_map (mk_args true) (strip_fun_mode mode ~~ Ts) []
bulwahn@34935
  1046
    fun strip_eval _ t =
bulwahn@34935
  1047
      let
bulwahn@34935
  1048
        val t' = strip_split_abs t
bulwahn@34935
  1049
        val (r, _) = PredicateCompFuns.dest_Eval t'
bulwahn@34935
  1050
      in (SOME (fst (strip_comb r)), NONE) end
bulwahn@34935
  1051
    val (inargs, outargs) = split_map_mode strip_eval mode args
bulwahn@34935
  1052
    val eval_hoargs = ho_args_of mode args
bulwahn@34935
  1053
    val hoargTs = ho_argsT_of mode Ts
bulwahn@34935
  1054
    val hoarg_names' =
bulwahn@34935
  1055
      Name.variant_list argnames ((map (fn i => "x" ^ string_of_int i)) (1 upto (length hoargTs)))
bulwahn@34935
  1056
    val hoargs' = map2 (curry Free) hoarg_names' hoargTs
bulwahn@34935
  1057
    val args' = replace_ho_args mode hoargs' args
bulwahn@34935
  1058
    val predpropI = HOLogic.mk_Trueprop (list_comb (pred, args'))
bulwahn@34935
  1059
    val predpropE = HOLogic.mk_Trueprop (list_comb (pred, args))
bulwahn@34935
  1060
    val param_eqs = map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) eval_hoargs hoargs'
bulwahn@34935
  1061
    val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, inargs),
bulwahn@34935
  1062
                    if null outargs then Free("y", HOLogic.unitT) else HOLogic.mk_tuple outargs))
bulwahn@34935
  1063
    val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, inargs),
bulwahn@34935
  1064
                     HOLogic.mk_tuple outargs))
bulwahn@34935
  1065
    val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
bulwahn@34935
  1066
    val simprules = [defthm, @{thm eval_pred},
bulwahn@34935
  1067
      @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
bulwahn@34935
  1068
    val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
bulwahn@36994
  1069
    val introthm = Goal.prove ctxt
bulwahn@34935
  1070
      (argnames @ hoarg_names' @ ["y"]) [] introtrm (fn _ => unfolddef_tac)
bulwahn@34935
  1071
    val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
bulwahn@34935
  1072
    val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
bulwahn@36994
  1073
    val elimthm = Goal.prove ctxt
bulwahn@34935
  1074
      (argnames @ ["y", "P"]) [] elimtrm (fn _ => unfolddef_tac)
bulwahn@35884
  1075
    val opt_neg_introthm =
bulwahn@35884
  1076
      if is_all_input mode then
bulwahn@35884
  1077
        let
bulwahn@35884
  1078
          val neg_predpropI = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (pred, args')))
bulwahn@35884
  1079
          val neg_funpropI =
bulwahn@35884
  1080
            HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval
bulwahn@35884
  1081
              (PredicateCompFuns.mk_not (list_comb (funtrm, inargs)), HOLogic.unit))
bulwahn@35884
  1082
          val neg_introtrm = Logic.list_implies (neg_predpropI :: param_eqs, neg_funpropI)
bulwahn@35884
  1083
          val tac =
bulwahn@35884
  1084
            Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps
bulwahn@35884
  1085
              (@{thm if_False} :: @{thm Predicate.not_pred_eq} :: simprules)) 1
bulwahn@35884
  1086
            THEN rtac @{thm Predicate.singleI} 1
bulwahn@36994
  1087
        in SOME (Goal.prove ctxt (argnames @ hoarg_names') []
bulwahn@35884
  1088
            neg_introtrm (fn _ => tac))
bulwahn@35884
  1089
        end
bulwahn@35884
  1090
      else NONE
bulwahn@34935
  1091
  in
bulwahn@35884
  1092
    ((introthm, elimthm), opt_neg_introthm)
bulwahn@34935
  1093
  end
bulwahn@32667
  1094
bulwahn@33620
  1095
fun create_constname_of_mode options thy prefix name T mode = 
bulwahn@32667
  1096
  let
bulwahn@33626
  1097
    val system_proposal = prefix ^ (Long_Name.base_name name)
bulwahn@34935
  1098
      ^ "_" ^ ascii_string_of_mode mode
bulwahn@34935
  1099
    val name = the_default system_proposal (proposed_names options name mode)
bulwahn@32667
  1100
  in
bulwahn@33620
  1101
    Sign.full_bname thy name
bulwahn@32667
  1102
  end;
bulwahn@32667
  1103
bulwahn@33620
  1104
fun create_definitions options preds (name, modes) thy =
bulwahn@32667
  1105
  let
bulwahn@32667
  1106
    val compfuns = PredicateCompFuns.compfuns
bulwahn@32667
  1107
    val T = AList.lookup (op =) preds name |> the
bulwahn@34935
  1108
    fun create_definition mode thy =
bulwahn@33752
  1109
      let
bulwahn@33752
  1110
        val mode_cname = create_constname_of_mode options thy "" name T mode
bulwahn@33752
  1111
        val mode_cbasename = Long_Name.base_name mode_cname
bulwahn@34935
  1112
        val funT = funT_of compfuns mode T
bulwahn@34935
  1113
        val (args, _) = fold_map (mk_args true) ((strip_fun_mode mode) ~~ (binder_types T)) []
bulwahn@34935
  1114
        fun strip_eval m t =
bulwahn@33752
  1115
          let
bulwahn@34935
  1116
            val t' = strip_split_abs t
bulwahn@34935
  1117
            val (r, _) = PredicateCompFuns.dest_Eval t'
bulwahn@34935
  1118
          in (SOME (fst (strip_comb r)), NONE) end
bulwahn@34935
  1119
        val (inargs, outargs) = split_map_mode strip_eval mode args
krauss@39989
  1120
        val predterm = fold_rev HOLogic.tupled_lambda inargs
krauss@39989
  1121
          (PredicateCompFuns.mk_Enum (HOLogic.tupled_lambda (HOLogic.mk_tuple outargs)
bulwahn@34935
  1122
            (list_comb (Const (name, T), args))))
bulwahn@34935
  1123
        val lhs = Const (mode_cname, funT)
bulwahn@33752
  1124
        val def = Logic.mk_equals (lhs, predterm)
bulwahn@33752
  1125
        val ([definition], thy') = thy |>
bulwahn@33752
  1126
          Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
wenzelm@39814
  1127
          Global_Theory.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
wenzelm@43232
  1128
        val ctxt' = Proof_Context.init_global thy'
bulwahn@35884
  1129
        val rules as ((intro, elim), _) =
bulwahn@36994
  1130
          create_intro_elim_rule ctxt' mode definition mode_cname funT (Const (name, T))
bulwahn@33752
  1131
        in thy'
bulwahn@34935
  1132
          |> set_function_name Pred name mode mode_cname
bulwahn@35884
  1133
          |> add_predfun_data name mode (definition, rules)
wenzelm@39814
  1134
          |> Global_Theory.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
wenzelm@39814
  1135
          |> Global_Theory.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
bulwahn@33752
  1136
          |> Theory.checkpoint
bulwahn@32667
  1137
        end;
bulwahn@32667
  1138
  in
bulwahn@34935
  1139
    thy |> defined_function_of Pred name |> fold create_definition modes
bulwahn@32667
  1140
  end;
bulwahn@32667
  1141
bulwahn@33620
  1142
fun define_functions comp_modifiers compfuns options preds (name, modes) thy =
bulwahn@32667
  1143
  let
bulwahn@32667
  1144
    val T = AList.lookup (op =) preds name |> the
bulwahn@32667
  1145
    fun create_definition mode thy =
bulwahn@32667
  1146
      let
bulwahn@33485
  1147
        val function_name_prefix = Comp_Mod.function_name_prefix comp_modifiers
bulwahn@33620
  1148
        val mode_cname = create_constname_of_mode options thy function_name_prefix name T mode
bulwahn@34935
  1149
        val funT = Comp_Mod.funT_of comp_modifiers mode T
bulwahn@32667
  1150
      in
bulwahn@32667
  1151
        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
bulwahn@34935
  1152
        |> set_function_name (Comp_Mod.compilation comp_modifiers) name mode mode_cname
bulwahn@32667
  1153
      end;
bulwahn@32667
  1154
  in
bulwahn@34935
  1155
    thy
bulwahn@34935
  1156
    |> defined_function_of (Comp_Mod.compilation comp_modifiers) name
bulwahn@34935
  1157
    |> fold create_definition modes
bulwahn@32667
  1158
  end;
bulwahn@32672
  1159
bulwahn@32667
  1160
bulwahn@32667
  1161
(* composition of mode inference, definition, compilation and proof *)
bulwahn@32667
  1162
bulwahn@32667
  1163
(** auxillary combinators for table of preds and modes **)
bulwahn@32667
  1164
bulwahn@32667
  1165
fun map_preds_modes f preds_modes_table =
bulwahn@32667
  1166
  map (fn (pred, modes) =>
bulwahn@32667
  1167
    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
bulwahn@32667
  1168
bulwahn@32667
  1169
fun join_preds_modes table1 table2 =
bulwahn@32667
  1170
  map_preds_modes (fn pred => fn mode => fn value =>
bulwahn@32667
  1171
    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
bulwahn@36254
  1172
bulwahn@32667
  1173
fun maps_modes preds_modes_table =
bulwahn@32667
  1174
  map (fn (pred, modes) =>
bulwahn@34935
  1175
    (pred, map (fn (mode, value) => value) modes)) preds_modes_table
bulwahn@36254
  1176
bulwahn@36996
  1177
fun compile_preds options comp_modifiers ctxt all_vs param_vs preds moded_clauses =
bulwahn@36996
  1178
  map_preds_modes (fn pred => compile_pred options comp_modifiers ctxt all_vs param_vs pred
bulwahn@33143
  1179
      (the (AList.lookup (op =) preds pred))) moded_clauses
bulwahn@33143
  1180
bulwahn@35324
  1181
fun prove options thy clauses preds moded_clauses compiled_terms =
bulwahn@35324
  1182
  map_preds_modes (prove_pred options thy clauses preds)
bulwahn@32667
  1183
    (join_preds_modes moded_clauses compiled_terms)
bulwahn@32667
  1184
bulwahn@35324
  1185
fun prove_by_skip options thy _ _ _ compiled_terms =
wenzelm@35021
  1186
  map_preds_modes
wenzelm@35021
  1187
    (fn pred => fn mode => fn t => Drule.export_without_context (Skip_Proof.make_thm thy t))
bulwahn@32667
  1188
    compiled_terms
bulwahn@33106
  1189
bulwahn@33376
  1190
(* preparation of introduction rules into special datastructures *)
bulwahn@36994
  1191
fun dest_prem ctxt params t =
bulwahn@33106
  1192
  (case strip_comb t of
bulwahn@34935
  1193
    (v as Free _, ts) => if member (op =) params v then Prem t else Sidecond t
bulwahn@36994
  1194
  | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem ctxt params t of
bulwahn@34935
  1195
      Prem t => Negprem t
bulwahn@33482
  1196
    | Negprem _ => error ("Double negation not allowed in premise: " ^
bulwahn@36994
  1197
        Syntax.string_of_term ctxt (c $ t)) 
bulwahn@33106
  1198
    | Sidecond t => Sidecond (c $ t))
bulwahn@33106
  1199
  | (c as Const (s, _), ts) =>
bulwahn@36994
  1200
    if is_registered ctxt s then Prem t else Sidecond t
bulwahn@33106
  1201
  | _ => Sidecond t)
bulwahn@34935
  1202
bulwahn@39556
  1203
fun prepare_intrs options ctxt prednames intros =
bulwahn@32667
  1204
  let
wenzelm@43232
  1205
    val thy = Proof_Context.theory_of ctxt
bulwahn@33126
  1206
    val intrs = map prop_of intros
bulwahn@33146
  1207
    val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames
bulwahn@33146
  1208
    val (preds, intrs) = unify_consts thy preds intrs
bulwahn@36994
  1209
    val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] ctxt
bulwahn@33126
  1210
    val preds = map dest_Const preds
bulwahn@34935
  1211
    val all_vs = terms_vs intrs
bulwahn@39431
  1212
    fun generate_modes s T =
bulwahn@39431
  1213
      if member (op =) (no_higher_order_predicate options) s then
bulwahn@39431
  1214
        all_smodes_of_typ T
bulwahn@39431
  1215
      else
bulwahn@39431
  1216
        all_modes_of_typ T
bulwahn@35324
  1217
    val all_modes = 
bulwahn@35324
  1218
      map (fn (s, T) =>
bulwahn@39628
  1219
        (s, case proposed_modes options s of
bulwahn@39875
  1220
            SOME ms => check_matches_type ctxt s T ms
bulwahn@39431
  1221
          | NONE => generate_modes s T)) preds
bulwahn@34935
  1222
    val params =
bulwahn@34935
  1223
      case intrs of
bulwahn@33146
  1224
        [] =>
bulwahn@33146
  1225
          let
bulwahn@34935
  1226
            val T = snd (hd preds)
bulwahn@39997
  1227
            val one_mode = hd (the (AList.lookup (op =) all_modes (fst (hd preds))))
bulwahn@34935
  1228
            val paramTs =
bulwahn@39997
  1229
              ho_argsT_of one_mode (binder_types T)
bulwahn@33482
  1230
            val param_names = Name.variant_list [] (map (fn i => "p" ^ string_of_int i)
bulwahn@33482
  1231
              (1 upto length paramTs))
bulwahn@34935
  1232
          in
bulwahn@34935
  1233
            map2 (curry Free) param_names paramTs
bulwahn@34935
  1234
          end
bulwahn@35324
  1235
      | (intr :: _) =>
bulwahn@35324
  1236
        let
bulwahn@39997
  1237
          val (p, args) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))
bulwahn@39997
  1238
          val one_mode = hd (the (AList.lookup (op =) all_modes (fst (dest_Const p))))
bulwahn@35324
  1239
        in
bulwahn@39997
  1240
          ho_args_of one_mode args
bulwahn@35324
  1241
        end
bulwahn@34935
  1242
    val param_vs = map (fst o dest_Free) params
bulwahn@34935
  1243
    fun add_clause intr clauses =
bulwahn@32667
  1244
      let
bulwahn@34935
  1245
        val (Const (name, T), ts) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))
bulwahn@36994
  1246
        val prems = map (dest_prem ctxt params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr)
bulwahn@32667
  1247
      in
bulwahn@34935
  1248
        AList.update op = (name, these (AList.lookup op = clauses name) @
bulwahn@34935
  1249
          [(ts, prems)]) clauses
bulwahn@34935
  1250
      end;
bulwahn@34935
  1251
    val clauses = fold add_clause intrs []
bulwahn@34935
  1252
  in
bulwahn@35324
  1253
    (preds, all_vs, param_vs, all_modes, clauses)
bulwahn@34935
  1254
  end;
bulwahn@32667
  1255
bulwahn@33376
  1256
(* sanity check of introduction rules *)
bulwahn@34935
  1257
(* TODO: rethink check with new modes *)
bulwahn@34935
  1258
(*
bulwahn@33106
  1259
fun check_format_of_intro_rule thy intro =
bulwahn@33106
  1260
  let
bulwahn@33106
  1261
    val concl = Logic.strip_imp_concl (prop_of intro)
bulwahn@33106
  1262
    val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
bulwahn@33629
  1263
    val params = fst (chop (nparams_of thy (fst (dest_Const p))) args)
bulwahn@33106
  1264
    fun check_arg arg = case HOLogic.strip_tupleT (fastype_of arg) of
bulwahn@33106
  1265
      (Ts as _ :: _ :: _) =>
bulwahn@33629
  1266
        if length (HOLogic.strip_tuple arg) = length Ts then
bulwahn@33629
  1267
          true
bulwahn@33114
  1268
        else
bulwahn@33629
  1269
          error ("Format of introduction rule is invalid: tuples must be expanded:"
bulwahn@33629
  1270
          ^ (Syntax.string_of_term_global thy arg) ^ " in " ^
bulwahn@33629
  1271
          (Display.string_of_thm_global thy intro)) 
bulwahn@33106
  1272
      | _ => true
bulwahn@33106
  1273
    val prems = Logic.strip_imp_prems (prop_of intro)
bulwahn@34935
  1274
    fun check_prem (Prem t) = forall check_arg args
bulwahn@34935
  1275
      | check_prem (Negprem t) = forall check_arg args
bulwahn@33106
  1276
      | check_prem _ = true
bulwahn@33106
  1277
  in
bulwahn@33106
  1278
    forall check_arg args andalso
bulwahn@33106
  1279
    forall (check_prem o dest_prem thy params o HOLogic.dest_Trueprop) prems
bulwahn@33106
  1280
  end
bulwahn@34935
  1281
*)
bulwahn@33124
  1282
(*
bulwahn@33124
  1283
fun check_intros_elim_match thy prednames =
bulwahn@33124
  1284
  let
bulwahn@33124
  1285
    fun check predname =
bulwahn@33124
  1286
      let
bulwahn@33124
  1287
        val intros = intros_of thy predname
bulwahn@33124
  1288
        val elim = the_elim_of thy predname
bulwahn@33124
  1289
        val nparams = nparams_of thy predname
bulwahn@33124
  1290
        val elim' =
wenzelm@35021
  1291
          (Drule.export_without_context o Skip_Proof.make_thm thy)
wenzelm@43232
  1292
          (mk_casesrule (Proof_Context.init_global thy) nparams intros)
bulwahn@33124
  1293
      in
bulwahn@33124
  1294
        if not (Thm.equiv_thm (elim, elim')) then
bulwahn@33124
  1295
          error "Introduction and elimination rules do not match!"
bulwahn@33124
  1296
        else true
bulwahn@33124
  1297
      end
bulwahn@33124
  1298
  in forall check prednames end
bulwahn@33124
  1299
*)
bulwahn@33113
  1300
bulwahn@33376
  1301
(* create code equation *)
bulwahn@33376
  1302
bulwahn@36998
  1303
fun add_code_equations ctxt preds result_thmss =
bulwahn@33376
  1304
  let
bulwahn@33629
  1305
    fun add_code_equation (predname, T) (pred, result_thms) =
bulwahn@33376
  1306
      let
bulwahn@34935
  1307
        val full_mode = fold_rev (curry Fun) (map (K Input) (binder_types T)) Bool
bulwahn@33376
  1308
      in
bulwahn@36994
  1309
        if member (op =) (modes_of Pred ctxt predname) full_mode then
bulwahn@33376
  1310
          let
bulwahn@33376
  1311
            val Ts = binder_types T
bulwahn@33376
  1312
            val arg_names = Name.variant_list []
bulwahn@33376
  1313
              (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
bulwahn@33629
  1314
            val args = map2 (curry Free) arg_names Ts
bulwahn@36994
  1315
            val predfun = Const (function_name_of Pred ctxt predname full_mode,
bulwahn@33376
  1316
              Ts ---> PredicateCompFuns.mk_predT @{typ unit})
bulwahn@33754
  1317
            val rhs = @{term Predicate.holds} $ (list_comb (predfun, args))
bulwahn@33376
  1318
            val eq_term = HOLogic.mk_Trueprop
bulwahn@33376
  1319
              (HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs))
bulwahn@35888
  1320
            val def = predfun_definition_of ctxt predname full_mode
wenzelm@33448
  1321
            val tac = fn _ => Simplifier.simp_tac
bulwahn@33754
  1322
              (HOL_basic_ss addsimps [def, @{thm holds_eq}, @{thm eval_pred}]) 1
bulwahn@35888
  1323
            val eq = Goal.prove ctxt arg_names [] eq_term tac
bulwahn@33376
  1324
          in
bulwahn@33376
  1325
            (pred, result_thms @ [eq])
bulwahn@33376
  1326
          end
bulwahn@33376
  1327
        else
bulwahn@33376
  1328
          (pred, result_thms)
bulwahn@33376
  1329
      end
bulwahn@33376
  1330
  in
bulwahn@33629
  1331
    map2 add_code_equation preds result_thmss
bulwahn@33376
  1332
  end
bulwahn@33376
  1333
bulwahn@32667
  1334
(** main function of predicate compiler **)
bulwahn@32667
  1335
bulwahn@33330
  1336
datatype steps = Steps of
bulwahn@33330
  1337
  {
bulwahn@35324
  1338
  define_functions : options -> (string * typ) list -> string * (bool * mode) list -> theory -> theory,
bulwahn@35324
  1339
  prove : options -> theory -> (string * (term list * indprem list) list) list -> (string * typ) list
bulwahn@33330
  1340
    -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
bulwahn@36998
  1341
  add_code_equations : Proof.context -> (string * typ) list
bulwahn@33376
  1342
    -> (string * thm list) list -> (string * thm list) list,
bulwahn@34935
  1343
  comp_modifiers : Comp_Mod.comp_modifiers,
bulwahn@39994
  1344
  use_generators : bool,
bulwahn@33330
  1345
  qname : bstring
bulwahn@33330
  1346
  }
bulwahn@33330
  1347
bulwahn@35324
  1348
fun add_equations_of steps mode_analysis_options options prednames thy =
bulwahn@32667
  1349
  let
bulwahn@33330
  1350
    fun dest_steps (Steps s) = s
bulwahn@35879
  1351
    val compilation = Comp_Mod.compilation (#comp_modifiers (dest_steps steps))
wenzelm@43232
  1352
    val ctxt = Proof_Context.init_global thy
bulwahn@33482
  1353
    val _ = print_step options
bulwahn@35879
  1354
      ("Starting predicate compiler (compilation: " ^ string_of_compilation compilation
bulwahn@35879
  1355
        ^ ") for predicates " ^ commas prednames ^ "...")
bulwahn@33124
  1356
      (*val _ = check_intros_elim_match thy prednames*)
bulwahn@33114
  1357
      (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
bulwahn@35324
  1358
    val _ =
bulwahn@35324
  1359
      if show_intermediate_results options then
bulwahn@36998
  1360
        tracing (commas (map (Display.string_of_thm ctxt) (maps (intros_of ctxt) prednames)))
bulwahn@35324
  1361
      else ()
bulwahn@35324
  1362
    val (preds, all_vs, param_vs, all_modes, clauses) =
bulwahn@39556
  1363
      prepare_intrs options ctxt prednames (maps (intros_of ctxt) prednames)
bulwahn@33123
  1364
    val _ = print_step options "Infering modes..."
bulwahn@39500
  1365
    val (lookup_mode, lookup_neg_mode, needs_random) = (modes_of compilation ctxt,
bulwahn@39500
  1366
      modes_of (negative_compilation_of compilation) ctxt, needs_random ctxt)
bulwahn@39500
  1367
    val ((moded_clauses, needs_random), errors) =
bulwahn@42952
  1368
      cond_timeit (Config.get ctxt Quickcheck.timing) "Infering modes"
bulwahn@36251
  1369
      (fn _ => infer_modes mode_analysis_options
bulwahn@39500
  1370
        options (lookup_mode, lookup_neg_mode, needs_random) ctxt preds all_modes param_vs clauses)
bulwahn@39500
  1371
    val thy' = fold (fn (s, ms) => set_needs_random s ms) needs_random thy
bulwahn@33132
  1372
    val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
bulwahn@40380
  1373
    val _ = check_expected_modes options preds modes
bulwahn@40380
  1374
    val _ = check_proposed_modes options preds modes errors
bulwahn@36995
  1375
    val _ = print_modes options modes
bulwahn@33123
  1376
    val _ = print_step options "Defining executable functions..."
bulwahn@36252
  1377
    val thy'' =
bulwahn@42952
  1378
      cond_timeit (Config.get ctxt Quickcheck.timing) "Defining executable functions..."
bulwahn@36252
  1379
      (fn _ => fold (#define_functions (dest_steps steps) options preds) modes thy'
bulwahn@36252
  1380
      |> Theory.checkpoint)
wenzelm@43232
  1381
    val ctxt'' = Proof_Context.init_global thy''
bulwahn@33123
  1382
    val _ = print_step options "Compiling equations..."
bulwahn@32667
  1383
    val compiled_terms =
bulwahn@42952
  1384
      cond_timeit (Config.get ctxt Quickcheck.timing) "Compiling equations...." (fn _ =>
bulwahn@36254
  1385
        compile_preds options
bulwahn@36996
  1386
          (#comp_modifiers (dest_steps steps)) ctxt'' all_vs param_vs preds moded_clauses)
bulwahn@36996
  1387
    val _ = print_compiled_terms options ctxt'' compiled_terms
bulwahn@33123
  1388
    val _ = print_step options "Proving equations..."
bulwahn@35324
  1389
    val result_thms =
bulwahn@42952
  1390
      cond_timeit (Config.get ctxt Quickcheck.timing) "Proving equations...." (fn _ =>
bulwahn@36252
  1391
      #prove (dest_steps steps) options thy'' clauses preds moded_clauses compiled_terms)
bulwahn@36998
  1392
    val result_thms' = #add_code_equations (dest_steps steps) ctxt'' preds
bulwahn@33376
  1393
      (maps_modes result_thms)
bulwahn@33330
  1394
    val qname = #qname (dest_steps steps)
bulwahn@32667
  1395
    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
bulwahn@32667
  1396
      (fn thm => Context.mapping (Code.add_eqn thm) I))))
bulwahn@36252
  1397
    val thy''' =
bulwahn@42952
  1398
      cond_timeit (Config.get ctxt Quickcheck.timing) "Setting code equations...." (fn _ =>
wenzelm@39814
  1399
      fold (fn (name, result_thms) => fn thy => snd (Global_Theory.add_thmss
bulwahn@32667
  1400
      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
bulwahn@32667
  1401
        [attrib thy ])] thy))
bulwahn@36252
  1402
      result_thms' thy'' |> Theory.checkpoint)
bulwahn@32667
  1403
  in
bulwahn@34935
  1404
    thy'''
bulwahn@32667
  1405
  end
bulwahn@32667
  1406
  
bulwahn@33132
  1407
fun gen_add_equations steps options names thy =
bulwahn@32667
  1408
  let
bulwahn@33330
  1409
    fun dest_steps (Steps s) = s
bulwahn@34935
  1410
    val defined = defined_functions (Comp_Mod.compilation (#comp_modifiers (dest_steps steps)))
bulwahn@40234
  1411
    val thy' = extend_intro_graph names thy |> Theory.checkpoint;
bulwahn@32667
  1412
    fun strong_conn_of gr keys =
bulwahn@32667
  1413
      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
bulwahn@32667
  1414
    val scc = strong_conn_of (PredData.get thy') names
bulwahn@40234
  1415
    val thy'' = fold preprocess_intros (flat scc) thy'
bulwahn@39881
  1416
    val thy''' = fold_rev
bulwahn@32667
  1417
      (fn preds => fn thy =>
wenzelm@43232
  1418
        if not (forall (defined (Proof_Context.init_global thy)) preds) then
bulwahn@35324
  1419
          let
bulwahn@39994
  1420
            val mode_analysis_options = {use_generators = #use_generators (dest_steps steps),
bulwahn@35324
  1421
              reorder_premises =
bulwahn@35324
  1422
                not (no_topmost_reordering options andalso not (null (inter (op =) preds names))),
bulwahn@39994
  1423
              infer_pos_and_neg_modes = #use_generators (dest_steps steps)}
bulwahn@35324
  1424
          in
bulwahn@35324
  1425
            add_equations_of steps mode_analysis_options options preds thy
bulwahn@35324
  1426
          end
bulwahn@33483
  1427
        else thy)
bulwahn@39881
  1428
      scc thy'' |> Theory.checkpoint
bulwahn@39881
  1429
  in thy''' end
bulwahn@35879
  1430
bulwahn@34935
  1431
val add_equations = gen_add_equations
bulwahn@35324
  1432
  (Steps {
bulwahn@35324
  1433
  define_functions =
bulwahn@35324
  1434
    fn options => fn preds => fn (s, modes) =>
bulwahn@35324
  1435
      create_definitions
bulwahn@35324
  1436
      options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
bulwahn@34935
  1437
  prove = prove,
bulwahn@34935
  1438
  add_code_equations = add_code_equations,
bulwahn@34935
  1439
  comp_modifiers = predicate_comp_modifiers,
bulwahn@39994
  1440
  use_generators = false,
bulwahn@34935
  1441
  qname = "equation"})
bulwahn@33143
  1442
bulwahn@33134
  1443
val add_depth_limited_equations = gen_add_equations
bulwahn@35879
  1444
  (Steps {
bulwahn@35879
  1445
  define_functions =
bulwahn@35879
  1446
    fn options => fn preds => fn (s, modes) =>
bulwahn@35879
  1447
    define_functions depth_limited_comp_modifiers PredicateCompFuns.compfuns
bulwahn@35879
  1448
    options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
bulwahn@32667
  1449
  prove = prove_by_skip,
bulwahn@34935
  1450
  add_code_equations = K (K I),
bulwahn@35879
  1451
  comp_modifiers = depth_limited_comp_modifiers,
bulwahn@39994
  1452
  use_generators = false,
bulwahn@33330
  1453
  qname = "depth_limited_equation"})
bulwahn@35879
  1454
bulwahn@33473
  1455
val add_annotated_equations = gen_add_equations
bulwahn@35324
  1456
  (Steps {
bulwahn@35324
  1457
  define_functions =
bulwahn@35324
  1458
    fn options => fn preds => fn (s, modes) =>
bulwahn@35324
  1459
      define_functions annotated_comp_modifiers PredicateCompFuns.compfuns options preds
bulwahn@35324
  1460
      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
bulwahn@33473
  1461
  prove = prove_by_skip,
bulwahn@34935
  1462
  add_code_equations = K (K I),
bulwahn@34935
  1463
  comp_modifiers = annotated_comp_modifiers,
bulwahn@39994
  1464
  use_generators = false,
bulwahn@33473
  1465
  qname = "annotated_equation"})
bulwahn@35880
  1466
bulwahn@35880
  1467
val add_random_equations = gen_add_equations
bulwahn@35880
  1468
  (Steps {
bulwahn@35880
  1469
  define_functions =
bulwahn@35880
  1470
    fn options => fn preds => fn (s, modes) =>
bulwahn@35880
  1471
      define_functions random_comp_modifiers PredicateCompFuns.compfuns options preds
bulwahn@35880
  1472
      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
bulwahn@35880
  1473
  comp_modifiers = random_comp_modifiers,
bulwahn@32667
  1474
  prove = prove_by_skip,
bulwahn@34935
  1475
  add_code_equations = K (K I),
bulwahn@39994
  1476
  use_generators = true,
bulwahn@33375
  1477
  qname = "random_equation"})
bulwahn@35880
  1478
bulwahn@35881
  1479
val add_depth_limited_random_equations = gen_add_equations
bulwahn@35881
  1480
  (Steps {
bulwahn@35881
  1481
  define_functions =
bulwahn@35881
  1482
    fn options => fn preds => fn (s, modes) =>
bulwahn@35881
  1483
      define_functions depth_limited_random_comp_modifiers PredicateCompFuns.compfuns options preds
bulwahn@35881
  1484
      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
bulwahn@35881
  1485
  comp_modifiers = depth_limited_random_comp_modifiers,
bulwahn@35881
  1486
  prove = prove_by_skip,
bulwahn@35881
  1487
  add_code_equations = K (K I),
bulwahn@39994
  1488
  use_generators = true,
bulwahn@35881
  1489
  qname = "depth_limited_random_equation"})
bulwahn@35881
  1490
bulwahn@34935
  1491
val add_dseq_equations = gen_add_equations
bulwahn@35324
  1492
  (Steps {
bulwahn@35324
  1493
  define_functions =
bulwahn@35324
  1494
  fn options => fn preds => fn (s, modes) =>
bulwahn@35324
  1495
    define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns
bulwahn@35324
  1496
    options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
bulwahn@34935
  1497
  prove = prove_by_skip,
bulwahn@34935
  1498
  add_code_equations = K (K I),
bulwahn@34935
  1499
  comp_modifiers = dseq_comp_modifiers,
bulwahn@39994
  1500
  use_generators = false,
bulwahn@34935
  1501
  qname = "dseq_equation"})
bulwahn@34935
  1502
bulwahn@34935
  1503
val add_random_dseq_equations = gen_add_equations
bulwahn@35324
  1504
  (Steps {
bulwahn@35324
  1505
  define_functions =
bulwahn@35324
  1506
    fn options => fn preds => fn (s, modes) =>
bulwahn@35324
  1507
    let
bulwahn@35324
  1508
      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
bulwahn@35324
  1509
      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
bulwahn@35324
  1510
    in define_functions pos_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
bulwahn@35324
  1511
      options preds (s, pos_modes)
bulwahn@35324
  1512
      #> define_functions neg_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
bulwahn@35324
  1513
      options preds (s, neg_modes)
bulwahn@35324
  1514
    end,
bulwahn@34935
  1515
  prove = prove_by_skip,
bulwahn@34935
  1516
  add_code_equations = K (K I),
bulwahn@35324
  1517
  comp_modifiers = pos_random_dseq_comp_modifiers,
bulwahn@39994
  1518
  use_generators = true,
bulwahn@34935
  1519
  qname = "random_dseq_equation"})
bulwahn@34935
  1520
bulwahn@36012
  1521
val add_new_random_dseq_equations = gen_add_equations
bulwahn@36012
  1522
  (Steps {
bulwahn@36012
  1523
  define_functions =
bulwahn@36012
  1524
    fn options => fn preds => fn (s, modes) =>
bulwahn@36012
  1525
    let
bulwahn@36012
  1526
      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
bulwahn@36012
  1527
      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
bulwahn@40230
  1528
    in define_functions new_pos_random_dseq_comp_modifiers New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
bulwahn@36012
  1529
      options preds (s, pos_modes)
bulwahn@40230
  1530
      #> define_functions new_neg_random_dseq_comp_modifiers New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns
bulwahn@36012
  1531
      options preds (s, neg_modes)
bulwahn@36012
  1532
    end,
bulwahn@36012
  1533
  prove = prove_by_skip,
bulwahn@36012
  1534
  add_code_equations = K (K I),
bulwahn@36012
  1535
  comp_modifiers = new_pos_random_dseq_comp_modifiers,
bulwahn@39994
  1536
  use_generators = true,
bulwahn@36012
  1537
  qname = "new_random_dseq_equation"})
bulwahn@32667
  1538
bulwahn@40232
  1539
val add_generator_dseq_equations = gen_add_equations
bulwahn@40232
  1540
  (Steps {
bulwahn@40232
  1541
  define_functions =
bulwahn@40232
  1542
  fn options => fn preds => fn (s, modes) =>
bulwahn@40232
  1543
    let
bulwahn@40232
  1544
      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
bulwahn@40232
  1545
      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
bulwahn@40232
  1546
    in 
bulwahn@40232
  1547
      define_functions pos_generator_dseq_comp_modifiers New_Pos_DSequence_CompFuns.depth_limited_compfuns
bulwahn@40232
  1548
        options preds (s, pos_modes)
bulwahn@40232
  1549
      #> define_functions neg_generator_dseq_comp_modifiers New_Neg_DSequence_CompFuns.depth_limited_compfuns
bulwahn@40232
  1550
        options preds (s, neg_modes)
bulwahn@40232
  1551
    end,
bulwahn@40232
  1552
  prove = prove_by_skip,
bulwahn@40232
  1553
  add_code_equations = K (K I),
bulwahn@40232
  1554
  comp_modifiers = pos_generator_dseq_comp_modifiers,
bulwahn@40232
  1555
  use_generators = true,
bulwahn@40232
  1556
  qname = "generator_dseq_equation"})
bulwahn@40232
  1557
bulwahn@32667
  1558
(** user interface **)
bulwahn@32667
  1559
bulwahn@32667
  1560
(* code_pred_intro attribute *)
bulwahn@32667
  1561
bulwahn@39767
  1562
fun attrib' f opt_case_name =
bulwahn@39767
  1563
  Thm.declaration_attribute (fn thm => Context.mapping (f (opt_case_name, thm)) I);
bulwahn@32667
  1564
bulwahn@39767
  1565
val code_pred_intro_attrib = attrib' add_intro NONE;
bulwahn@32667
  1566
bulwahn@32668
  1567
(*FIXME
bulwahn@32668
  1568
- Naming of auxiliary rules necessary?
bulwahn@32668
  1569
*)
bulwahn@32668
  1570
bulwahn@43012
  1571
(* values_timeout configuration *)
bulwahn@43012
  1572
krauss@44767
  1573
val default_values_timeout =
krauss@44767
  1574
  if String.isPrefix "smlnj" (getenv "ML_SYSTEM") then 600.0 else 20.0
krauss@44767
  1575
krauss@44767
  1576
val values_timeout = Attrib.setup_config_real @{binding values_timeout} (K default_values_timeout)
bulwahn@43012
  1577
bulwahn@32668
  1578
val setup = PredData.put (Graph.empty) #>
bulwahn@39767
  1579
  Attrib.setup @{binding code_pred_intro} (Scan.lift (Scan.option Args.name) >> attrib' add_intro)
bulwahn@32668
  1580
    "adding alternative introduction rules for code generation of inductive predicates"
bulwahn@32667
  1581
wenzelm@33522
  1582
(* TODO: make Theory_Data to Generic_Data & remove duplication of local theory and theory *)
wenzelm@39032
  1583
(* FIXME ... this is important to avoid changing the background theory below *)
bulwahn@33132
  1584
fun generic_code_pred prep_const options raw_const lthy =
bulwahn@32667
  1585
  let
wenzelm@43232
  1586
    val thy = Proof_Context.theory_of lthy
bulwahn@32667
  1587
    val const = prep_const thy raw_const
bulwahn@40234
  1588
    val lthy' = Local_Theory.background_theory (extend_intro_graph [const]) lthy
wenzelm@43232
  1589
    val thy' = Proof_Context.theory_of lthy'
wenzelm@43232
  1590
    val ctxt' = Proof_Context.init_global thy'
bulwahn@36998
  1591
    val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim ctxt')
bulwahn@32667
  1592
    fun mk_cases const =
bulwahn@32667
  1593
      let
bulwahn@39532
  1594
        val T = Sign.the_const_type thy' const
bulwahn@33146
  1595
        val pred = Const (const, T)
bulwahn@36998
  1596
        val intros = intros_of ctxt' const
bulwahn@34935
  1597
      in mk_casesrule lthy' pred intros end  
bulwahn@32667
  1598
    val cases_rules = map mk_cases preds
bulwahn@32667
  1599
    val cases =
bulwahn@39767
  1600
      map2 (fn pred_name => fn case_rule => Rule_Cases.Case {fixes = [],
bulwahn@39876
  1601
        assumes = ("that", tl (Logic.strip_imp_prems case_rule))
bulwahn@39767
  1602
          :: (map_filter (fn (fact_name, fact) => Option.map (rpair [fact]) fact_name)
bulwahn@39876
  1603
            ((SOME (Long_Name.base_name pred_name ^ ".prems") :: names_of ctxt' pred_name) ~~ Logic.strip_imp_prems case_rule)),
bulwahn@39767
  1604
        binds = [], cases = []}) preds cases_rules
bulwahn@32667
  1605
    val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
bulwahn@32667
  1606
    val lthy'' = lthy'
bulwahn@39767
  1607
      |> fold Variable.auto_fixes cases_rules
wenzelm@43232
  1608
      |> Proof_Context.add_cases true case_env
bulwahn@32667
  1609
    fun after_qed thms goal_ctxt =
bulwahn@32667
  1610
      let
wenzelm@43232
  1611
        val global_thms = Proof_Context.export goal_ctxt
wenzelm@43232
  1612
          (Proof_Context.init_global (Proof_Context.theory_of goal_ctxt)) (map the_single thms)
bulwahn@32667
  1613
      in
wenzelm@39032
  1614
        goal_ctxt |> Local_Theory.background_theory (fold set_elim global_thms #>
bulwahn@34935
  1615
          ((case compilation options of
bulwahn@34935
  1616
             Pred => add_equations
bulwahn@34935
  1617
           | DSeq => add_dseq_equations
bulwahn@35879
  1618
           | Pos_Random_DSeq => add_random_dseq_equations
bulwahn@35879
  1619
           | Depth_Limited => add_depth_limited_equations
bulwahn@35880
  1620
           | Random => add_random_equations
bulwahn@35881
  1621
           | Depth_Limited_Random => add_depth_limited_random_equations
bulwahn@36012
  1622
           | New_Pos_Random_DSeq => add_new_random_dseq_equations
bulwahn@40232
  1623
           | Pos_Generator_DSeq => add_generator_dseq_equations
bulwahn@34935
  1624
           | compilation => error ("Compilation not supported")
bulwahn@34935
  1625
           ) options [const]))
bulwahn@33144
  1626
      end
bulwahn@32667
  1627
  in
wenzelm@36334
  1628
    Proof.theorem NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
bulwahn@32667
  1629
  end;
bulwahn@32667
  1630
bulwahn@32667
  1631
val code_pred = generic_code_pred (K I);
bulwahn@32667
  1632
val code_pred_cmd = generic_code_pred Code.read_const
bulwahn@32667
  1633
bulwahn@32667
  1634
(* transformation for code generation *)
bulwahn@32667
  1635
wenzelm@41720
  1636
(* FIXME just one data slot (record) per program unit *)
wenzelm@41720
  1637
wenzelm@41720
  1638
structure Pred_Result = Proof_Data
wenzelm@41720
  1639
(
haftmann@39634
  1640
  type T = unit -> term Predicate.pred
wenzelm@41720
  1641
  (* FIXME avoid user error with non-user text *)
haftmann@39634
  1642
  fun init _ () = error "Pred_Result"
haftmann@39634
  1643
);
haftmann@39634
  1644
val put_pred_result = Pred_Result.put;
haftmann@39634
  1645
wenzelm@41720
  1646
structure Pred_Random_Result = Proof_Data
wenzelm@41720
  1647
(
haftmann@39634
  1648
  type T = unit -> int * int -> term Predicate.pred * (int * int)
wenzelm@41720
  1649
  (* FIXME avoid user error with non-user text *)
haftmann@39634
  1650
  fun init _ () = error "Pred_Random_Result"
haftmann@39634
  1651
);
haftmann@39634
  1652
val put_pred_random_result = Pred_Random_Result.put;
haftmann@39634
  1653
wenzelm@41720
  1654
structure Dseq_Result = Proof_Data
wenzelm@41720
  1655
(
haftmann@39634
  1656
  type T = unit -> term DSequence.dseq
wenzelm@41720
  1657
  (* FIXME avoid user error with non-user text *)
haftmann@39634
  1658
  fun init _ () = error "Dseq_Result"
haftmann@39634
  1659
);
haftmann@39634
  1660
val put_dseq_result = Dseq_Result.put;
haftmann@39634
  1661
wenzelm@41720
  1662
structure Dseq_Random_Result = Proof_Data
wenzelm@41720
  1663
(
haftmann@39634
  1664
  type T = unit -> int -> int -> int * int -> term DSequence.dseq * (int * int)
wenzelm@41720
  1665
  (* FIXME avoid user error with non-user text *)
haftmann@39634
  1666
  fun init _ () = error "Dseq_Random_Result"
haftmann@39634
  1667
);
haftmann@39634
  1668
val put_dseq_random_result = Dseq_Random_Result.put;
haftmann@39634
  1669
wenzelm@41720
  1670
structure New_Dseq_Result = Proof_Data
wenzelm@41720
  1671
(
bulwahn@40348
  1672
  type T = unit -> int -> term Lazy_Sequence.lazy_sequence
wenzelm@41720
  1673
  (* FIXME avoid user error with non-user text *)
bulwahn@40348
  1674
  fun init _ () = error "New_Dseq_Random_Result"
bulwahn@40348
  1675
);
bulwahn@40348
  1676
val put_new_dseq_result = New_Dseq_Result.put;
bulwahn@40348
  1677
wenzelm@41720
  1678
structure Lseq_Random_Result = Proof_Data
wenzelm@41720
  1679
(
haftmann@39634
  1680
  type T = unit -> int -> int -> int * int -> int -> term Lazy_Sequence.lazy_sequence
wenzelm@41720
  1681
  (* FIXME avoid user error with non-user text *)
haftmann@39634
  1682
  fun init _ () = error "Lseq_Random_Result"
haftmann@39634
  1683
);
haftmann@39634
  1684
val put_lseq_random_result = Lseq_Random_Result.put;
haftmann@39634
  1685
wenzelm@41720
  1686
structure Lseq_Random_Stats_Result = Proof_Data
wenzelm@41720
  1687
(
haftmann@39634
  1688
  type T = unit -> int -> int -> int * int -> int -> (term * int) Lazy_Sequence.lazy_sequence
wenzelm@41720
  1689
  (* FIXME avoid user error with non-user text *)
haftmann@39634
  1690
  fun init _ () = error "Lseq_Random_Stats_Result"
haftmann@39634
  1691
);
haftmann@39634
  1692
val put_lseq_random_stats_result = Lseq_Random_Stats_Result.put;
bulwahn@32667
  1693
bulwahn@42958
  1694
fun dest_special_compr t =
bulwahn@42958
  1695
  let
bulwahn@42958
  1696
    val (inner_t, T_compr) = case t of (Const (@{const_name Collect}, _) $ Abs (x, T, t)) => (t, T)
bulwahn@42958
  1697
      | _ => raise TERM ("dest_special_compr", [t])
bulwahn@42958
  1698
    val (Ts, conj) = apfst (map snd) (Predicate_Compile_Aux.strip_ex inner_t)
bulwahn@42958
  1699
    val [eq, body] = HOLogic.dest_conj conj
bulwahn@42958
  1700
    val rhs = case HOLogic.dest_eq eq of
bulwahn@42958
  1701
        (Bound i, rhs) => if i = length Ts then rhs else raise TERM ("dest_special_compr", [t])
bulwahn@42958
  1702
      | _ => raise TERM ("dest_special_compr", [t])
bulwahn@42958
  1703
    val output_names = Name.variant_list (fold Term.add_free_names [rhs, body] [])
bulwahn@42958
  1704
      (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
bulwahn@42958
  1705
    val output_frees = map2 (curry Free) output_names (rev Ts)
bulwahn@42958
  1706
    val body = subst_bounds (output_frees, body)
bulwahn@42958
  1707
    val output = subst_bounds (output_frees, rhs)
bulwahn@42958
  1708
  in
bulwahn@42958
  1709
    (((body, output), T_compr), output_names)
bulwahn@42958
  1710
  end
bulwahn@42958
  1711
bulwahn@42958
  1712
fun dest_general_compr ctxt t_compr =
bulwahn@42958
  1713
  let      
bulwahn@42958
  1714
    val inner_t = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
bulwahn@42958
  1715
      | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);    
bulwahn@42958
  1716
    val (body, Ts, fp) = HOLogic.strip_psplits inner_t;
bulwahn@42958
  1717
    val output_names = Name.variant_list (Term.add_free_names body [])
bulwahn@42958
  1718
      (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
bulwahn@42958
  1719
    val output_frees = map2 (curry Free) output_names (rev Ts)
bulwahn@42958
  1720
    val body = subst_bounds (output_frees, body)
bulwahn@42958
  1721
    val T_compr = HOLogic.mk_ptupleT fp Ts
bulwahn@42958
  1722
    val output = HOLogic.mk_ptuple fp T_compr (rev output_frees)
bulwahn@42958
  1723
  in
bulwahn@42958
  1724
    (((body, output), T_compr), output_names)
bulwahn@42958
  1725
  end
bulwahn@42958
  1726
bulwahn@32667
  1727
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
bulwahn@40348
  1728
fun analyze_compr ctxt (comp_modifiers, additional_arguments) param_user_modes
bulwahn@40348
  1729
  (compilation, arguments) t_compr =
bulwahn@32667
  1730
  let
bulwahn@40348
  1731
    val compfuns = Comp_Mod.compfuns comp_modifiers
bulwahn@34935
  1732
    val all_modes_of = all_modes_of compilation
bulwahn@42958
  1733
    val (((body, output), T_compr), output_names) =
bulwahn@42958
  1734
      case try dest_special_compr t_compr of SOME r => r | NONE => dest_general_compr ctxt t_compr
bulwahn@36025
  1735
    val (pred as Const (name, T), all_args) =
bulwahn@36025
  1736
      case strip_comb body of
bulwahn@36025
  1737
        (Const (name, T), all_args) => (Const (name, T), all_args)
bulwahn@36994
  1738
      | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
bulwahn@34935
  1739
  in
bulwahn@36994
  1740
    if defined_functions compilation ctxt name then
bulwahn@32668
  1741
      let
haftmann@37366
  1742
        fun extract_mode (Const (@{const_name Pair}, _) $ t1 $ t2) = Pair (extract_mode t1, extract_mode t2)
bulwahn@34935
  1743
          | extract_mode (Free (x, _)) = if member (op =) output_names x then Output else Input
bulwahn@34935
  1744
          | extract_mode _ = Input
bulwahn@34935
  1745
        val user_mode = fold_rev (curry Fun) (map extract_mode all_args) Bool
bulwahn@34935
  1746
        fun valid modes1 modes2 =
bulwahn@34935
  1747
          case int_ord (length modes1, length modes2) of
bulwahn@34935
  1748
            GREATER => error "Not enough mode annotations"
bulwahn@34935
  1749
          | LESS => error "Too many mode annotations"
bulwahn@34935
  1750
          | EQUAL => forall (fn (m, NONE) => true | (m, SOME m2) => eq_mode (m, m2))
bulwahn@34935
  1751
            (modes1 ~~ modes2)
bulwahn@34935
  1752
        fun mode_instance_of (m1, m2) =
bulwahn@34935
  1753
          let
bulwahn@34935
  1754
            fun instance_of (Fun _, Input) = true
bulwahn@34935
  1755
              | instance_of (Input, Input) = true
bulwahn@34935
  1756
              | instance_of (Output, Output) = true
bulwahn@34935
  1757
              | instance_of (Pair (m1, m2), Pair (m1', m2')) =
bulwahn@34935
  1758
                  instance_of  (m1, m1') andalso instance_of (m2, m2')
bulwahn@34935
  1759
              | instance_of (Pair (m1, m2), Input) =
bulwahn@34935
  1760
                  instance_of (m1, Input) andalso instance_of (m2, Input)
bulwahn@34935
  1761
              | instance_of (Pair (m1, m2), Output) =
bulwahn@34935
  1762
                  instance_of (m1, Output) andalso instance_of (m2, Output)
bulwahn@36993
  1763
              | instance_of (Input, Pair (m1, m2)) =
bulwahn@36993
  1764
                  instance_of (Input, m1) andalso instance_of (Input, m2)
bulwahn@36993
  1765
              | instance_of (Output, Pair (m1, m2)) =
bulwahn@36993
  1766
                  instance_of (Output, m1) andalso instance_of (Output, m2)
bulwahn@34935
  1767
              | instance_of _ = false
bulwahn@34935
  1768
          in forall instance_of (strip_fun_mode m1 ~~ strip_fun_mode m2) end
bulwahn@36994
  1769
        val derivs = all_derivations_of ctxt (all_modes_of ctxt) [] body
bulwahn@34935
  1770
          |> filter (fn (d, missing_vars) =>
bulwahn@34935
  1771
            let
bulwahn@34935
  1772
              val (p_mode :: modes) = collect_context_modes d
bulwahn@34935
  1773
            in
bulwahn@34935
  1774
              null missing_vars andalso
bulwahn@34935
  1775
              mode_instance_of (p_mode, user_mode) andalso
bulwahn@34935
  1776
              the_default true (Option.map (valid modes) param_user_modes)
bulwahn@34935
  1777
            end)
bulwahn@34935
  1778
          |> map fst
bulwahn@34935
  1779
        val deriv = case derivs of
bulwahn@34935
  1780
            [] => error ("No mode possible for comprehension "
bulwahn@36994
  1781
                    ^ Syntax.string_of_term ctxt t_compr)
bulwahn@34935
  1782
          | [d] => d
bulwahn@34935
  1783
          | d :: _ :: _ => (warning ("Multiple modes possible for comprehension "
bulwahn@36994
  1784
                    ^ Syntax.string_of_term ctxt t_compr); d);
bulwahn@34935
  1785
        val (_, outargs) = split_mode (head_mode_of deriv) all_args
bulwahn@36994
  1786
        val t_pred = compile_expr comp_modifiers ctxt
bulwahn@39872
  1787
          (body, deriv) [] additional_arguments;
bulwahn@34935
  1788
        val T_pred = dest_predT compfuns (fastype_of t_pred)
bulwahn@42958
  1789
        val arrange = HOLogic.tupled_lambda (HOLogic.mk_tuple outargs) output
bulwahn@34935
  1790
      in
bulwahn@34935
  1791
        if null outargs then t_pred else mk_map compfuns T_pred T_compr arrange t_pred
bulwahn@34935
  1792
      end
bulwahn@34935
  1793
    else
bulwahn@34935
  1794
      error "Evaluation with values is not possible because compilation with code_pred was not invoked"
bulwahn@34935
  1795
  end
bulwahn@32667
  1796
bulwahn@36994
  1797
fun eval ctxt stats param_user_modes (options as (compilation, arguments)) k t_compr =
bulwahn@32667
  1798
  let
bulwahn@36021
  1799
    fun count xs x =
bulwahn@36021
  1800
      let
bulwahn@36021
  1801
        fun count' i [] = i
bulwahn@36021
  1802
          | count' i (x' :: xs) = if x = x' then count' (i + 1) xs else count' i xs
bulwahn@36021
  1803
      in count' 0 xs end
bulwahn@36021
  1804
    fun accumulate xs = map (fn x => (x, count xs x)) (sort int_ord (distinct (op =) xs))
bulwahn@40348
  1805
    val comp_modifiers =
bulwahn@34935
  1806
      case compilation of
bulwahn@40348
  1807
          Pred => predicate_comp_modifiers
bulwahn@40348
  1808
        | Random => random_comp_modifiers
bulwahn@40348
  1809
        | Depth_Limited => depth_limited_comp_modifiers
bulwahn@40348
  1810
        | Depth_Limited_Random => depth_limited_random_comp_modifiers
bulwahn@40348
  1811
        (*| Annotated => annotated_comp_modifiers*)
bulwahn@40348
  1812
        | DSeq => dseq_comp_modifiers
bulwahn@40348
  1813
        | Pos_Random_DSeq => pos_random_dseq_comp_modifiers
bulwahn@40348
  1814
        | New_Pos_Random_DSeq => new_pos_random_dseq_comp_modifiers
bulwahn@40348
  1815
        | Pos_Generator_DSeq => pos_generator_dseq_comp_modifiers
bulwahn@40348
  1816
    val compfuns = Comp_Mod.compfuns comp_modifiers
bulwahn@40348
  1817
    val additional_arguments =
bulwahn@40348
  1818
      case compilation of
bulwahn@40348
  1819
        Pred => []
bulwahn@40348
  1820
      | Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
bulwahn@40348
  1821
        [@{term "(1, 1) :: code_numeral * code_numeral"}]
bulwahn@40348
  1822
      | Annotated => []
bulwahn@40348
  1823
      | Depth_Limited => [HOLogic.mk_number @{typ "code_numeral"} (hd arguments)]
bulwahn@40348
  1824
      | Depth_Limited_Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
bulwahn@40348
  1825
        [@{term "(1, 1) :: code_numeral * code_numeral"}]
bulwahn@40348
  1826
      | DSeq => []
bulwahn@40348
  1827
      | Pos_Random_DSeq => []
bulwahn@40348
  1828
      | New_Pos_Random_DSeq => []
bulwahn@40348
  1829
      | Pos_Generator_DSeq => []
bulwahn@40348
  1830
    val t = analyze_compr ctxt (comp_modifiers, additional_arguments) param_user_modes options t_compr;
bulwahn@33137
  1831
    val T = dest_predT compfuns (fastype_of t);
bulwahn@36021
  1832
    val t' =
bulwahn@36021
  1833
      if stats andalso compilation = New_Pos_Random_DSeq then
bulwahn@36021
  1834
        mk_map compfuns T (HOLogic.mk_prodT (HOLogic.termT, @{typ code_numeral}))
bulwahn@36021
  1835
          (absdummy (T, HOLogic.mk_prod (HOLogic.term_of_const T $ Bound 0,
bulwahn@36021
  1836
            @{term Code_Numeral.of_nat} $ (HOLogic.size_const T $ Bound 0)))) t
bulwahn@36021
  1837
      else
bulwahn@36021
  1838
        mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t
wenzelm@43232
  1839
    val thy = Proof_Context.theory_of ctxt
bulwahn@43012
  1840
    val time_limit = seconds (Config.get ctxt values_timeout)
bulwahn@36021
  1841
    val (ts, statistics) =
bulwahn@40377
  1842
      (case compilation of
bulwahn@35880
  1843
       (* Random =>
bulwahn@34935
  1844
          fst (Predicate.yieldn k
bulwahn@34935
  1845
          (Code_Eval.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
bulwahn@33137
  1846
            (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
bulwahn@35880
  1847
            |> Random_Engine.run))*)
bulwahn@35880
  1848
        Pos_Random_DSeq =>
bulwahn@34935
  1849
          let
bulwahn@34935
  1850
            val [nrandom, size, depth] = arguments
bulwahn@34935
  1851
          in
bulwahn@43012
  1852
            rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (DSequence.yieldn k
haftmann@39711
  1853
              (Code_Runtime.dynamic_value_strict (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result")
haftmann@39711
  1854
                thy NONE (fn proc => fn g => fn nrandom => fn size => fn s => g nrandom size s |>> DSequence.map proc)
haftmann@39711
  1855
                  t' [] nrandom size
bulwahn@34935
  1856
                |> Random_Engine.run)
bulwahn@40377
  1857
              depth true)) ())
bulwahn@34935
  1858
          end
bulwahn@34935
  1859
      | DSeq =>
bulwahn@43012
  1860
          rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (DSequence.yieldn k
haftmann@39711
  1861
            (Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result")
bulwahn@40377
  1862
              thy NONE DSequence.map t' []) (the_single arguments) true)) ())
bulwahn@40348
  1863
      | Pos_Generator_DSeq =>
bulwahn@40348
  1864
          let
bulwahn@40348
  1865
            val depth = (the_single arguments)
bulwahn@40348
  1866
          in
bulwahn@43012
  1867
            rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Lazy_Sequence.yieldn k
bulwahn@40348
  1868
              (Code_Runtime.dynamic_value_strict (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Core.put_new_dseq_result")
bulwahn@40348
  1869
              thy NONE (fn proc => fn g => fn depth => g depth |> Lazy_Sequence.mapa proc)
bulwahn@40377
  1870
              t' [] depth))) ())
bulwahn@40348
  1871
          end
bulwahn@36014
  1872
      | New_Pos_Random_DSeq =>
bulwahn@36014
  1873
          let
bulwahn@36014
  1874
            val [nrandom, size, depth] = arguments
bulwahn@36014
  1875
            val seed = Random_Engine.next_seed ()
bulwahn@36014
  1876
          in
bulwahn@36021
  1877
            if stats then
bulwahn@43012
  1878
              apsnd (SOME o accumulate) (split_list (TimeLimit.timeLimit time_limit
bulwahn@40377
  1879
              (fn () => fst (Lazy_Sequence.yieldn k
haftmann@39711
  1880
                (Code_Runtime.dynamic_value_strict
haftmann@39634
  1881
                  (Lseq_Random_Stats_Result.get, put_lseq_random_stats_result, "Predicate_Compile_Core.put_lseq_random_stats_result")
haftmann@39711
  1882
                  thy NONE
bulwahn@36021
  1883
                  (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
haftmann@36533
  1884
                    |> Lazy_Sequence.mapa (apfst proc))
bulwahn@40377
  1885
                    t' [] nrandom size seed depth))) ()))
bulwahn@36021
  1886
            else rpair NONE
bulwahn@43012
  1887
              (TimeLimit.timeLimit time_limit (fn () => fst (Lazy_Sequence.yieldn k
haftmann@39711
  1888
                (Code_Runtime.dynamic_value_strict
haftmann@39634
  1889
                  (Lseq_Random_Result.get, put_lseq_random_result, "Predicate_Compile_Core.put_lseq_random_result")
haftmann@39711
  1890
                  thy NONE 
bulwahn@36021
  1891
                  (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
haftmann@36533
  1892
                    |> Lazy_Sequence.mapa proc)
bulwahn@40377
  1893
                    t' [] nrandom size seed depth))) ())
bulwahn@36014
  1894
          end
bulwahn@34935
  1895
      | _ =>
bulwahn@43012
  1896
          rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Predicate.yieldn k
haftmann@39711
  1897
            (Code_Runtime.dynamic_value_strict (Pred_Result.get, put_pred_result, "Predicate_Compile_Core.put_pred_result")
bulwahn@40377
  1898
              thy NONE Predicate.map t' []))) ()))
bulwahn@40377
  1899
     handle TimeLimit.TimeOut => error "Reached timeout during execution of values"
bulwahn@36021
  1900
  in ((T, ts), statistics) end;
bulwahn@36021
  1901
bulwahn@43964
  1902
fun values param_user_modes ((raw_expected, stats), comp_options) k t_compr ctxt =
bulwahn@32667
  1903
  let
bulwahn@36994
  1904
    val ((T, ts), statistics) = eval ctxt stats param_user_modes comp_options k t_compr
bulwahn@34935
  1905
    val setT = HOLogic.mk_setT T
bulwahn@34935
  1906
    val elems = HOLogic.mk_set T ts
bulwahn@43964
  1907
    val ([dots], ctxt') =
wenzelm@44500
  1908
      Proof_Context.add_fixes [(@{binding dots}, SOME setT, Mixfix ("...", [], 1000))] ctxt 
bulwahn@34935
  1909
    (* check expected values *)
bulwahn@43964
  1910
    val union = Const (@{const_abbrev Set.union}, setT --> setT --> setT)
bulwahn@34935
  1911
    val () =
bulwahn@34935
  1912
      case raw_expected of
bulwahn@34935
  1913
        NONE => ()
bulwahn@34935
  1914
      | SOME s =>
bulwahn@34935
  1915
        if eq_set (op =) (HOLogic.dest_set (Syntax.read_term ctxt s), ts) then ()
bulwahn@34935
  1916
        else
bulwahn@34935
  1917
          error ("expected and computed values do not match:\n" ^
bulwahn@34935
  1918
            "expected values: " ^ Syntax.string_of_term ctxt (Syntax.read_term ctxt s) ^ "\n" ^
bulwahn@34935
  1919
            "computed values: " ^ Syntax.string_of_term ctxt elems ^ "\n")
bulwahn@34935
  1920
  in
bulwahn@43964
  1921
    ((if k = ~1 orelse length ts < k then elems else union $ elems $ Free (dots, setT), statistics), ctxt')
bulwahn@32667
  1922
  end;
bulwahn@33623
  1923
bulwahn@33479
  1924
fun values_cmd print_modes param_user_modes options k raw_t state =
bulwahn@32667
  1925
  let
bulwahn@34935
  1926
    val ctxt = Toplevel.context_of state
bulwahn@34935
  1927
    val t = Syntax.read_term ctxt raw_t
bulwahn@43964
  1928
    val ((t', stats), ctxt') = values param_user_modes options k t ctxt
bulwahn@34935
  1929
    val ty' = Term.type_of t'
bulwahn@43964
  1930
    val ctxt'' = Variable.auto_fixes t' ctxt'
bulwahn@36021
  1931
    val pretty_stat =
bulwahn@36021
  1932
      case stats of
bulwahn@36021
  1933
          NONE => []
bulwahn@36021
  1934
        | SOME xs =>
bulwahn@36021
  1935
          let
bulwahn@36021
  1936
            val total = fold (curry (op +)) (map snd xs) 0
bulwahn@36021
  1937
            fun pretty_entry (s, n) =
bulwahn@36021
  1938
              [Pretty.str "size", Pretty.brk 1,
bulwahn@36021
  1939
               Pretty.str (string_of_int s), Pretty.str ":", Pretty.brk 1,
bulwahn@36021
  1940
               Pretty.str (string_of_int n), Pretty.fbrk]
bulwahn@36021
  1941
          in
bulwahn@36021
  1942
            [Pretty.fbrk, Pretty.str "Statistics:", Pretty.fbrk,
bulwahn@36021
  1943
             Pretty.str "total:", Pretty.brk 1, Pretty.str (string_of_int total), Pretty.fbrk]
bulwahn@36021
  1944
             @ maps pretty_entry xs
bulwahn@36021
  1945
          end
wenzelm@37154
  1946
    val p = Print_Mode.with_modes print_modes (fn () =>
bulwahn@43964
  1947
      Pretty.block ([Pretty.quote (Syntax.pretty_term ctxt'' t'), Pretty.fbrk,
bulwahn@43964
  1948
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt'' ty')]
bulwahn@36021
  1949
        @ pretty_stat)) ();
bulwahn@32667
  1950
  in Pretty.writeln p end;
bulwahn@32667
  1951
bulwahn@32667
  1952
end;