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