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