src/HOL/ex/predicate_compile.ML
changeset 33332 9b5286c0fb14
parent 33331 d8bfa9564a52
child 33333 78faaec3209f
equal deleted inserted replaced
33331:d8bfa9564a52 33332:9b5286c0fb14
     1 (* Author: Lukas Bulwahn, TU Muenchen
       
     2 
       
     3 (Prototype of) A compiler from predicates specified by intro/elim rules
       
     4 to equations.
       
     5 *)
       
     6 
       
     7 signature PREDICATE_COMPILE =
       
     8 sig
       
     9   type mode = int list option list * int list
       
    10   (*val add_equations_of: bool -> string list -> theory -> theory *)
       
    11   val register_predicate : (thm list * thm * int) -> theory -> theory
       
    12   val is_registered : theory -> string -> bool
       
    13  (* val fetch_pred_data : theory -> string -> (thm list * thm * int)  *)
       
    14   val predfun_intro_of: theory -> string -> mode -> thm
       
    15   val predfun_elim_of: theory -> string -> mode -> thm
       
    16   val strip_intro_concl: int -> term -> term * (term list * term list)
       
    17   val predfun_name_of: theory -> string -> mode -> string
       
    18   val all_preds_of : theory -> string list
       
    19   val modes_of: theory -> string -> mode list
       
    20   val string_of_mode : mode -> string
       
    21   val intros_of: theory -> string -> thm list
       
    22   val nparams_of: theory -> string -> int
       
    23   val add_intro: thm -> theory -> theory
       
    24   val set_elim: thm -> theory -> theory
       
    25   val setup: theory -> theory
       
    26   val code_pred: string -> Proof.context -> Proof.state
       
    27   val code_pred_cmd: string -> Proof.context -> Proof.state
       
    28   val print_stored_rules: theory -> unit
       
    29   val print_all_modes: theory -> unit
       
    30   val do_proofs: bool Unsynchronized.ref
       
    31   val mk_casesrule : Proof.context -> int -> thm list -> term
       
    32   val analyze_compr: theory -> term -> term
       
    33   val eval_ref: (unit -> term Predicate.pred) option Unsynchronized.ref
       
    34   val add_equations : string list -> theory -> theory
       
    35   val code_pred_intros_attrib : attribute
       
    36   (* used by Quickcheck_Generator *) 
       
    37   (*val funT_of : mode -> typ -> typ
       
    38   val mk_if_pred : term -> term
       
    39   val mk_Eval : term * term -> term*)
       
    40   val mk_tupleT : typ list -> typ
       
    41 (*  val mk_predT :  typ -> typ *)
       
    42   (* temporary for testing of the compilation *)
       
    43   datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
       
    44     GeneratorPrem of term list * term | Generator of (string * typ);
       
    45   val prepare_intrs: theory -> string list ->
       
    46     (string * typ) list * int * string list * string list * (string * mode list) list *
       
    47     (string * (term list * indprem list) list) list * (string * (int option list * int)) list
       
    48   datatype compilation_funs = CompilationFuns of {
       
    49     mk_predT : typ -> typ,
       
    50     dest_predT : typ -> typ,
       
    51     mk_bot : typ -> term,
       
    52     mk_single : term -> term,
       
    53     mk_bind : term * term -> term,
       
    54     mk_sup : term * term -> term,
       
    55     mk_if : term -> term,
       
    56     mk_not : term -> term,
       
    57     mk_map : typ -> typ -> term -> term -> term,
       
    58     lift_pred : term -> term
       
    59   };  
       
    60   datatype tmode = Mode of mode * int list * tmode option list;
       
    61   type moded_clause = term list * (indprem * tmode) list
       
    62   type 'a pred_mode_table = (string * (mode * 'a) list) list
       
    63   val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list
       
    64     -> (string * (int option list * int)) list -> string list
       
    65     -> (string * (term list * indprem list) list) list
       
    66     -> (moded_clause list) pred_mode_table
       
    67   val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list
       
    68     -> (string * (int option list * int)) list -> string list
       
    69     -> (string * (term list * indprem list) list) list
       
    70     -> (moded_clause list) pred_mode_table  
       
    71   (*val compile_preds : theory -> compilation_funs -> string list -> string list
       
    72     -> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table
       
    73   val rpred_create_definitions :(string * typ) list -> string * mode list
       
    74     -> theory -> theory 
       
    75   val split_smode : int list -> term list -> (term list * term list) *)
       
    76   val print_moded_clauses :
       
    77     theory -> (moded_clause list) pred_mode_table -> unit
       
    78   val print_compiled_terms : theory -> term pred_mode_table -> unit
       
    79   (*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*)
       
    80   val rpred_compfuns : compilation_funs
       
    81   val dest_funT : typ -> typ * typ
       
    82  (* val depending_preds_of : theory -> thm list -> string list *)
       
    83   val add_quickcheck_equations : string list -> theory -> theory
       
    84   val add_sizelim_equations : string list -> theory -> theory
       
    85   val is_inductive_predicate : theory -> string -> bool
       
    86   val terms_vs : term list -> string list
       
    87   val subsets : int -> int -> int list list
       
    88   val check_mode_clause : bool -> theory -> string list ->
       
    89     (string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list)
       
    90       -> (term list * (indprem * tmode) list) option
       
    91   val string_of_moded_prem : theory -> (indprem * tmode) -> string
       
    92   val all_modes_of : theory -> (string * mode list) list
       
    93   val all_generator_modes_of : theory -> (string * mode list) list
       
    94   val compile_clause : compilation_funs -> term option -> (term list -> term) ->
       
    95     theory -> string list -> string list -> mode -> term -> moded_clause -> term
       
    96   val preprocess_intro : theory -> thm -> thm
       
    97   val is_constrt : theory -> term -> bool
       
    98   val is_predT : typ -> bool
       
    99   val guess_nparams : typ -> int
       
   100 end;
       
   101 
       
   102 structure Predicate_Compile : PREDICATE_COMPILE =
       
   103 struct
       
   104 
       
   105 (** auxiliary **)
       
   106 
       
   107 (* debug stuff *)
       
   108 
       
   109 fun tracing s = (if ! Toplevel.debug then tracing s else ());
       
   110 
       
   111 fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
       
   112 fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *)
       
   113 
       
   114 val do_proofs = Unsynchronized.ref true;
       
   115 
       
   116 fun mycheat_tac thy i st =
       
   117   (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) i) st
       
   118 
       
   119 fun remove_last_goal thy st =
       
   120   (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st
       
   121 
       
   122 (* reference to preprocessing of InductiveSet package *)
       
   123 
       
   124 val ind_set_codegen_preproc = Inductive_Set.codegen_preproc;
       
   125 
       
   126 (** fundamentals **)
       
   127 
       
   128 (* syntactic operations *)
       
   129 
       
   130 fun mk_eq (x, xs) =
       
   131   let fun mk_eqs _ [] = []
       
   132         | mk_eqs a (b::cs) =
       
   133             HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
       
   134   in mk_eqs x xs end;
       
   135 
       
   136 fun mk_tupleT [] = HOLogic.unitT
       
   137   | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts;
       
   138 
       
   139 fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = []
       
   140   | dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2)
       
   141   | dest_tupleT t = [t]
       
   142 
       
   143 fun mk_tuple [] = HOLogic.unit
       
   144   | mk_tuple ts = foldr1 HOLogic.mk_prod ts;
       
   145 
       
   146 fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = []
       
   147   | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2)
       
   148   | dest_tuple t = [t]
       
   149 
       
   150 fun mk_scomp (t, u) =
       
   151   let
       
   152     val T = fastype_of t
       
   153     val U = fastype_of u
       
   154     val [A] = binder_types T
       
   155     val D = body_type U 
       
   156   in 
       
   157     Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
       
   158   end;
       
   159 
       
   160 fun dest_funT (Type ("fun",[S, T])) = (S, T)
       
   161   | dest_funT T = raise TYPE ("dest_funT", [T], [])
       
   162  
       
   163 fun mk_fun_comp (t, u) =
       
   164   let
       
   165     val (_, B) = dest_funT (fastype_of t)
       
   166     val (C, A) = dest_funT (fastype_of u)
       
   167   in
       
   168     Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
       
   169   end;
       
   170 
       
   171 fun dest_randomT (Type ("fun", [@{typ Random.seed},
       
   172   Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
       
   173   | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
       
   174 
       
   175 (* destruction of intro rules *)
       
   176 
       
   177 (* FIXME: look for other place where this functionality was used before *)
       
   178 fun strip_intro_concl nparams intro = let
       
   179   val _ $ u = Logic.strip_imp_concl intro
       
   180   val (pred, all_args) = strip_comb u
       
   181   val (params, args) = chop nparams all_args
       
   182 in (pred, (params, args)) end
       
   183 
       
   184 (** data structures **)
       
   185 
       
   186 type smode = int list;
       
   187 type mode = smode option list * smode;
       
   188 datatype tmode = Mode of mode * int list * tmode option list;
       
   189 
       
   190 fun split_smode is ts =
       
   191   let
       
   192     fun split_smode' _ _ [] = ([], [])
       
   193       | split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t)
       
   194           (split_smode' is (i+1) ts)
       
   195   in split_smode' is 1 ts end
       
   196 
       
   197 fun split_mode (iss, is) ts =
       
   198   let
       
   199     val (t1, t2) = chop (length iss) ts 
       
   200   in (t1, split_smode is t2) end
       
   201 
       
   202 fun string_of_mode (iss, is) = space_implode " -> " (map
       
   203   (fn NONE => "X"
       
   204     | SOME js => enclose "[" "]" (commas (map string_of_int js)))
       
   205        (iss @ [SOME is]));
       
   206 
       
   207 fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
       
   208   "predmode: " ^ (string_of_mode predmode) ^ 
       
   209   (if null param_modes then "" else
       
   210     "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
       
   211     
       
   212 datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
       
   213   GeneratorPrem of term list * term | Generator of (string * typ);
       
   214 
       
   215 type moded_clause = term list * (indprem * tmode) list
       
   216 type 'a pred_mode_table = (string * (mode * 'a) list) list
       
   217 
       
   218 datatype predfun_data = PredfunData of {
       
   219   name : string,
       
   220   definition : thm,
       
   221   intro : thm,
       
   222   elim : thm
       
   223 };
       
   224 
       
   225 fun rep_predfun_data (PredfunData data) = data;
       
   226 fun mk_predfun_data (name, definition, intro, elim) =
       
   227   PredfunData {name = name, definition = definition, intro = intro, elim = elim}
       
   228 
       
   229 datatype function_data = FunctionData of {
       
   230   name : string,
       
   231   equation : thm option (* is not used at all? *)
       
   232 };
       
   233 
       
   234 fun rep_function_data (FunctionData data) = data;
       
   235 fun mk_function_data (name, equation) =
       
   236   FunctionData {name = name, equation = equation}
       
   237 
       
   238 datatype pred_data = PredData of {
       
   239   intros : thm list,
       
   240   elim : thm option,
       
   241   nparams : int,
       
   242   functions : (mode * predfun_data) list,
       
   243   generators : (mode * function_data) list,
       
   244   sizelim_functions : (mode * function_data) list 
       
   245 };
       
   246 
       
   247 fun rep_pred_data (PredData data) = data;
       
   248 fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
       
   249   PredData {intros = intros, elim = elim, nparams = nparams,
       
   250     functions = functions, generators = generators, sizelim_functions = sizelim_functions}
       
   251 fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) =
       
   252   mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions)))
       
   253   
       
   254 fun eq_option eq (NONE, NONE) = true
       
   255   | eq_option eq (SOME x, SOME y) = eq (x, y)
       
   256   | eq_option eq _ = false
       
   257   
       
   258 fun eq_pred_data (PredData d1, PredData d2) = 
       
   259   eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso
       
   260   eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso
       
   261   #nparams d1 = #nparams d2
       
   262   
       
   263 structure PredData = TheoryDataFun
       
   264 (
       
   265   type T = pred_data Graph.T;
       
   266   val empty = Graph.empty;
       
   267   val copy = I;
       
   268   val extend = I;
       
   269   fun merge _ = Graph.merge eq_pred_data;
       
   270 );
       
   271 
       
   272 (* queries *)
       
   273 
       
   274 fun lookup_pred_data thy name =
       
   275   Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
       
   276 
       
   277 fun the_pred_data thy name = case lookup_pred_data thy name
       
   278  of NONE => error ("No such predicate " ^ quote name)  
       
   279   | SOME data => data;
       
   280 
       
   281 val is_registered = is_some oo lookup_pred_data 
       
   282 
       
   283 val all_preds_of = Graph.keys o PredData.get
       
   284 
       
   285 val intros_of = #intros oo the_pred_data
       
   286 
       
   287 fun the_elim_of thy name = case #elim (the_pred_data thy name)
       
   288  of NONE => error ("No elimination rule for predicate " ^ quote name)
       
   289   | SOME thm => thm 
       
   290   
       
   291 val has_elim = is_some o #elim oo the_pred_data;
       
   292 
       
   293 val nparams_of = #nparams oo the_pred_data
       
   294 
       
   295 val modes_of = (map fst) o #functions oo the_pred_data
       
   296 
       
   297 fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) 
       
   298 
       
   299 val is_compiled = not o null o #functions oo the_pred_data
       
   300 
       
   301 fun lookup_predfun_data thy name mode =
       
   302   Option.map rep_predfun_data (AList.lookup (op =)
       
   303   (#functions (the_pred_data thy name)) mode)
       
   304 
       
   305 fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
       
   306   of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
       
   307    | SOME data => data;
       
   308 
       
   309 val predfun_name_of = #name ooo the_predfun_data
       
   310 
       
   311 val predfun_definition_of = #definition ooo the_predfun_data
       
   312 
       
   313 val predfun_intro_of = #intro ooo the_predfun_data
       
   314 
       
   315 val predfun_elim_of = #elim ooo the_predfun_data
       
   316 
       
   317 fun lookup_generator_data thy name mode = 
       
   318   Option.map rep_function_data (AList.lookup (op =)
       
   319   (#generators (the_pred_data thy name)) mode)
       
   320   
       
   321 fun the_generator_data thy name mode = case lookup_generator_data thy name mode
       
   322   of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
       
   323    | SOME data => data
       
   324 
       
   325 val generator_name_of = #name ooo the_generator_data
       
   326 
       
   327 val generator_modes_of = (map fst) o #generators oo the_pred_data
       
   328 
       
   329 fun all_generator_modes_of thy =
       
   330   map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) 
       
   331 
       
   332 fun lookup_sizelim_function_data thy name mode =
       
   333   Option.map rep_function_data (AList.lookup (op =)
       
   334   (#sizelim_functions (the_pred_data thy name)) mode)
       
   335 
       
   336 fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode
       
   337   of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode
       
   338     ^ " of predicate " ^ name)
       
   339    | SOME data => data
       
   340 
       
   341 val sizelim_function_name_of = #name ooo the_sizelim_function_data
       
   342 
       
   343 (*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
       
   344      
       
   345 (* diagnostic display functions *)
       
   346 
       
   347 fun print_modes modes = tracing ("Inferred modes:\n" ^
       
   348   cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
       
   349     string_of_mode ms)) modes));
       
   350 
       
   351 fun print_pred_mode_table string_of_entry thy pred_mode_table =
       
   352   let
       
   353     fun print_mode pred (mode, entry) =  "mode : " ^ (string_of_mode mode)
       
   354       ^ (string_of_entry pred mode entry)  
       
   355     fun print_pred (pred, modes) =
       
   356       "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
       
   357     val _ = tracing (cat_lines (map print_pred pred_mode_table))
       
   358   in () end;
       
   359 
       
   360 fun string_of_moded_prem thy (Prem (ts, p), tmode) =
       
   361     (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
       
   362     "(" ^ (string_of_tmode tmode) ^ ")"
       
   363   | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) =
       
   364     (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
       
   365     "(generator_mode: " ^ (string_of_mode predmode) ^ ")"
       
   366   | string_of_moded_prem thy (Generator (v, T), _) =
       
   367     "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
       
   368   | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
       
   369     (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
       
   370     "(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"
       
   371   | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
       
   372     (Syntax.string_of_term_global thy t) ^
       
   373     "(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"    
       
   374   | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
       
   375      
       
   376 fun print_moded_clauses thy =
       
   377   let        
       
   378     fun string_of_clause pred mode clauses =
       
   379       cat_lines (map (fn (ts, prems) => (space_implode " --> "
       
   380         (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
       
   381         ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
       
   382   in print_pred_mode_table string_of_clause thy end;
       
   383 
       
   384 fun print_compiled_terms thy =
       
   385   print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
       
   386     
       
   387 fun print_stored_rules thy =
       
   388   let
       
   389     val preds = (Graph.keys o PredData.get) thy
       
   390     fun print pred () = let
       
   391       val _ = writeln ("predicate: " ^ pred)
       
   392       val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
       
   393       val _ = writeln ("introrules: ")
       
   394       val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
       
   395         (rev (intros_of thy pred)) ()
       
   396     in
       
   397       if (has_elim thy pred) then
       
   398         writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
       
   399       else
       
   400         writeln ("no elimrule defined")
       
   401     end
       
   402   in
       
   403     fold print preds ()
       
   404   end;
       
   405 
       
   406 fun print_all_modes thy =
       
   407   let
       
   408     val _ = writeln ("Inferred modes:")
       
   409     fun print (pred, modes) u =
       
   410       let
       
   411         val _ = writeln ("predicate: " ^ pred)
       
   412         val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
       
   413       in u end  
       
   414   in
       
   415     fold print (all_modes_of thy) ()
       
   416   end
       
   417   
       
   418 (** preprocessing rules **)  
       
   419 
       
   420 fun imp_prems_conv cv ct =
       
   421   case Thm.term_of ct of
       
   422     Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
       
   423   | _ => Conv.all_conv ct
       
   424 
       
   425 fun Trueprop_conv cv ct =
       
   426   case Thm.term_of ct of
       
   427     Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
       
   428   | _ => error "Trueprop_conv"
       
   429 
       
   430 fun preprocess_intro thy rule =
       
   431   Conv.fconv_rule
       
   432     (imp_prems_conv
       
   433       (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
       
   434     (Thm.transfer thy rule)
       
   435 
       
   436 fun preprocess_elim thy nparams elimrule =
       
   437   let
       
   438     fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
       
   439        HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
       
   440      | replace_eqs t = t
       
   441     val prems = Thm.prems_of elimrule
       
   442     val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
       
   443     fun preprocess_case t =
       
   444      let
       
   445        val params = Logic.strip_params t
       
   446        val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
       
   447        val assums_hyp' = assums1 @ (map replace_eqs assums2)
       
   448      in
       
   449        list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
       
   450      end 
       
   451     val cases' = map preprocess_case (tl prems)
       
   452     val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
       
   453   in
       
   454     Thm.equal_elim
       
   455       (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
       
   456          (cterm_of thy elimrule')))
       
   457       elimrule
       
   458   end;
       
   459 
       
   460 (* special case: predicate with no introduction rule *)
       
   461 fun noclause thy predname elim = let
       
   462   val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
       
   463   val Ts = binder_types T
       
   464   val names = Name.variant_list []
       
   465         (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
       
   466   val vs = map2 (curry Free) names Ts
       
   467   val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
       
   468   val intro_t = Logic.mk_implies (@{prop False}, clausehd)
       
   469   val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
       
   470   val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
       
   471   val intro = Goal.prove (ProofContext.init thy) names [] intro_t
       
   472         (fn {...} => etac @{thm FalseE} 1)
       
   473   val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
       
   474         (fn {...} => etac elim 1) 
       
   475 in
       
   476   ([intro], elim)
       
   477 end
       
   478 
       
   479 fun fetch_pred_data thy name =
       
   480   case try (Inductive.the_inductive (ProofContext.init thy)) name of
       
   481     SOME (info as (_, result)) => 
       
   482       let
       
   483         fun is_intro_of intro =
       
   484           let
       
   485             val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
       
   486           in (fst (dest_Const const) = name) end;      
       
   487         val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy))
       
   488           (filter is_intro_of (#intrs result)))
       
   489         val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
       
   490         val nparams = length (Inductive.params_of (#raw_induct result))
       
   491         val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
       
   492         val (intros, elim) = if null intros then noclause thy name elim else (intros, elim)
       
   493       in
       
   494         mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
       
   495       end                                                                    
       
   496   | NONE => error ("No such predicate: " ^ quote name)
       
   497   
       
   498 (* updaters *)
       
   499 
       
   500 fun apfst3 f (x, y, z) =  (f x, y, z)
       
   501 fun apsnd3 f (x, y, z) =  (x, f y, z)
       
   502 fun aptrd3 f (x, y, z) =  (x, y, f z)
       
   503 
       
   504 fun add_predfun name mode data =
       
   505   let
       
   506     val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data)
       
   507   in PredData.map (Graph.map_node name (map_pred_data add)) end
       
   508 
       
   509 fun is_inductive_predicate thy name =
       
   510   is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
       
   511 
       
   512 fun depending_preds_of thy (key, value) =
       
   513   let
       
   514     val intros = (#intros o rep_pred_data) value
       
   515   in
       
   516     fold Term.add_const_names (map Thm.prop_of intros) []
       
   517       |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c))
       
   518   end;
       
   519     
       
   520     
       
   521 (* code dependency graph *)    
       
   522 (*
       
   523 fun dependencies_of thy name =
       
   524   let
       
   525     val (intros, elim, nparams) = fetch_pred_data thy name 
       
   526     val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
       
   527     val keys = depending_preds_of thy intros
       
   528   in
       
   529     (data, keys)
       
   530   end;
       
   531 *)
       
   532 (* guessing number of parameters *)
       
   533 fun find_indexes pred xs =
       
   534   let
       
   535     fun find is n [] = is
       
   536       | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs;
       
   537   in rev (find [] 0 xs) end;
       
   538 
       
   539 fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
       
   540   | is_predT _ = false
       
   541   
       
   542 fun guess_nparams T =
       
   543   let
       
   544     val argTs = binder_types T
       
   545     val nparams = fold Integer.max
       
   546       (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
       
   547   in nparams end;
       
   548 
       
   549 fun add_intro thm thy = let
       
   550    val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
       
   551    fun cons_intro gr =
       
   552      case try (Graph.get_node gr) name of
       
   553        SOME pred_data => Graph.map_node name (map_pred_data
       
   554          (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
       
   555      | NONE =>
       
   556        let
       
   557          val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
       
   558        in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end;
       
   559   in PredData.map cons_intro thy end
       
   560 
       
   561 fun set_elim thm = let
       
   562     val (name, _) = dest_Const (fst 
       
   563       (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
       
   564     fun set (intros, _, nparams) = (intros, SOME thm, nparams)  
       
   565   in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
       
   566 
       
   567 fun set_nparams name nparams = let
       
   568     fun set (intros, elim, _ ) = (intros, elim, nparams) 
       
   569   in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
       
   570     
       
   571 fun register_predicate (pre_intros, pre_elim, nparams) thy = let
       
   572     val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros))))
       
   573     (* preprocessing *)
       
   574     val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
       
   575     val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
       
   576   in
       
   577     PredData.map
       
   578       (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
       
   579   end
       
   580 
       
   581 fun set_generator_name pred mode name = 
       
   582   let
       
   583     val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE))
       
   584   in
       
   585     PredData.map (Graph.map_node pred (map_pred_data set))
       
   586   end
       
   587 
       
   588 fun set_sizelim_function_name pred mode name = 
       
   589   let
       
   590     val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
       
   591   in
       
   592     PredData.map (Graph.map_node pred (map_pred_data set))
       
   593   end
       
   594 
       
   595 (** data structures for generic compilation for different monads **)
       
   596 
       
   597 (* maybe rename functions more generic:
       
   598   mk_predT -> mk_monadT; dest_predT -> dest_monadT
       
   599   mk_single -> mk_return (?)
       
   600 *)
       
   601 datatype compilation_funs = CompilationFuns of {
       
   602   mk_predT : typ -> typ,
       
   603   dest_predT : typ -> typ,
       
   604   mk_bot : typ -> term,
       
   605   mk_single : term -> term,
       
   606   mk_bind : term * term -> term,
       
   607   mk_sup : term * term -> term,
       
   608   mk_if : term -> term,
       
   609   mk_not : term -> term,
       
   610 (*  funT_of : mode -> typ -> typ, *)
       
   611 (*  mk_fun_of : theory -> (string * typ) -> mode -> term, *) 
       
   612   mk_map : typ -> typ -> term -> term -> term,
       
   613   lift_pred : term -> term
       
   614 };
       
   615 
       
   616 fun mk_predT (CompilationFuns funs) = #mk_predT funs
       
   617 fun dest_predT (CompilationFuns funs) = #dest_predT funs
       
   618 fun mk_bot (CompilationFuns funs) = #mk_bot funs
       
   619 fun mk_single (CompilationFuns funs) = #mk_single funs
       
   620 fun mk_bind (CompilationFuns funs) = #mk_bind funs
       
   621 fun mk_sup (CompilationFuns funs) = #mk_sup funs
       
   622 fun mk_if (CompilationFuns funs) = #mk_if funs
       
   623 fun mk_not (CompilationFuns funs) = #mk_not funs
       
   624 (*fun funT_of (CompilationFuns funs) = #funT_of funs*)
       
   625 (*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*)
       
   626 fun mk_map (CompilationFuns funs) = #mk_map funs
       
   627 fun lift_pred (CompilationFuns funs) = #lift_pred funs
       
   628 
       
   629 fun funT_of compfuns (iss, is) T =
       
   630   let
       
   631     val Ts = binder_types T
       
   632     val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
       
   633     val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs 
       
   634   in
       
   635     (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs))
       
   636   end;
       
   637 
       
   638 fun sizelim_funT_of compfuns (iss, is) T =
       
   639   let
       
   640     val Ts = binder_types T
       
   641     val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
       
   642     val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs 
       
   643   in
       
   644     (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
       
   645   end;  
       
   646 
       
   647 fun mk_fun_of compfuns thy (name, T) mode = 
       
   648   Const (predfun_name_of thy name mode, funT_of compfuns mode T)
       
   649 
       
   650 fun mk_sizelim_fun_of compfuns thy (name, T) mode =
       
   651   Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T)
       
   652   
       
   653 fun mk_generator_of compfuns thy (name, T) mode = 
       
   654   Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T)
       
   655 
       
   656 
       
   657 structure PredicateCompFuns =
       
   658 struct
       
   659 
       
   660 fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
       
   661 
       
   662 fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
       
   663   | dest_predT T = raise TYPE ("dest_predT", [T], []);
       
   664 
       
   665 fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
       
   666 
       
   667 fun mk_single t =
       
   668   let val T = fastype_of t
       
   669   in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
       
   670 
       
   671 fun mk_bind (x, f) =
       
   672   let val T as Type ("fun", [_, U]) = fastype_of f
       
   673   in
       
   674     Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
       
   675   end;
       
   676 
       
   677 val mk_sup = HOLogic.mk_binop @{const_name sup};
       
   678 
       
   679 fun mk_if cond = Const (@{const_name Predicate.if_pred},
       
   680   HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
       
   681 
       
   682 fun mk_not t = let val T = mk_predT HOLogic.unitT
       
   683   in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
       
   684 
       
   685 fun mk_Enum f =
       
   686   let val T as Type ("fun", [T', _]) = fastype_of f
       
   687   in
       
   688     Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
       
   689   end;
       
   690 
       
   691 fun mk_Eval (f, x) =
       
   692   let
       
   693     val T = fastype_of x
       
   694   in
       
   695     Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
       
   696   end;
       
   697 
       
   698 fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
       
   699   (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
       
   700 
       
   701 val lift_pred = I
       
   702 
       
   703 val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
       
   704   mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
       
   705   mk_map = mk_map, lift_pred = lift_pred};
       
   706 
       
   707 end;
       
   708 
       
   709 (* termify_code:
       
   710 val termT = Type ("Code_Evaluation.term", []);
       
   711 fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT)
       
   712 *)
       
   713 (*
       
   714 fun lift_random random =
       
   715   let
       
   716     val T = dest_randomT (fastype_of random)
       
   717   in
       
   718     mk_scomp (random,
       
   719       mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed},
       
   720         mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)),
       
   721           Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T)))) 
       
   722   end;
       
   723 *)
       
   724  
       
   725 structure RPredCompFuns =
       
   726 struct
       
   727 
       
   728 fun mk_rpredT T =
       
   729   @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
       
   730 
       
   731 fun dest_rpredT (Type ("fun", [_,
       
   732   Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
       
   733   | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); 
       
   734 
       
   735 fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
       
   736 
       
   737 fun mk_single t =
       
   738   let
       
   739     val T = fastype_of t
       
   740   in
       
   741     Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
       
   742   end;
       
   743 
       
   744 fun mk_bind (x, f) =
       
   745   let
       
   746     val T as (Type ("fun", [_, U])) = fastype_of f
       
   747   in
       
   748     Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
       
   749   end
       
   750 
       
   751 val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
       
   752 
       
   753 fun mk_if cond = Const (@{const_name RPred.if_rpred},
       
   754   HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
       
   755 
       
   756 fun mk_not t = error "Negation is not defined for RPred"
       
   757 
       
   758 fun mk_map t = error "FIXME" (*FIXME*)
       
   759 
       
   760 fun lift_pred t =
       
   761   let
       
   762     val T = PredicateCompFuns.dest_predT (fastype_of t)
       
   763     val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T 
       
   764   in
       
   765     Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t  
       
   766   end;
       
   767 
       
   768 val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
       
   769     mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
       
   770     mk_map = mk_map, lift_pred = lift_pred};
       
   771 
       
   772 end;
       
   773 (* for external use with interactive mode *)
       
   774 val rpred_compfuns = RPredCompFuns.compfuns;
       
   775 
       
   776 fun lift_random random =
       
   777   let
       
   778     val T = dest_randomT (fastype_of random)
       
   779   in
       
   780     Const (@{const_name lift_random}, (@{typ Random.seed} -->
       
   781       HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> 
       
   782       RPredCompFuns.mk_rpredT T) $ random
       
   783   end;
       
   784  
       
   785 (* Mode analysis *)
       
   786 
       
   787 (*** check if a term contains only constructor functions ***)
       
   788 fun is_constrt thy =
       
   789   let
       
   790     val cnstrs = flat (maps
       
   791       (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
       
   792       (Symtab.dest (Datatype.get_all thy)));
       
   793     fun check t = (case strip_comb t of
       
   794         (Free _, []) => true
       
   795       | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
       
   796             (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
       
   797           | _ => false)
       
   798       | _ => false)
       
   799   in check end;
       
   800 
       
   801 (*** check if a type is an equality type (i.e. doesn't contain fun)
       
   802   FIXME this is only an approximation ***)
       
   803 fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
       
   804   | is_eqT _ = true;
       
   805 
       
   806 fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
       
   807 val terms_vs = distinct (op =) o maps term_vs;
       
   808 
       
   809 (** collect all Frees in a term (with duplicates!) **)
       
   810 fun term_vTs tm =
       
   811   fold_aterms (fn Free xT => cons xT | _ => I) tm [];
       
   812 
       
   813 (*FIXME this function should not be named merge... make it local instead*)
       
   814 fun merge xs [] = xs
       
   815   | merge [] ys = ys
       
   816   | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
       
   817       else y::merge (x::xs) ys;
       
   818 
       
   819 fun subsets i j = if i <= j then
       
   820        let val is = subsets (i+1) j
       
   821        in merge (map (fn ks => i::ks) is) is end
       
   822      else [[]];
       
   823      
       
   824 (* FIXME: should be in library - map_prod *)
       
   825 fun cprod ([], ys) = []
       
   826   | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
       
   827 
       
   828 fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
       
   829 
       
   830 
       
   831   
       
   832 (*TODO: cleanup function and put together with modes_of_term *)
       
   833 (*
       
   834 fun modes_of_param default modes t = let
       
   835     val (vs, t') = strip_abs t
       
   836     val b = length vs
       
   837     fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
       
   838         let
       
   839           val (args1, args2) =
       
   840             if length args < length iss then
       
   841               error ("Too few arguments for inductive predicate " ^ name)
       
   842             else chop (length iss) args;
       
   843           val k = length args2;
       
   844           val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1)
       
   845             (1 upto b)  
       
   846           val partial_mode = (1 upto k) \\ perm
       
   847         in
       
   848           if not (partial_mode subset is) then [] else
       
   849           let
       
   850             val is' = 
       
   851             (fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm [])
       
   852             |> fold (fn i => if i > k then cons (i - k + b) else I) is
       
   853               
       
   854            val res = map (fn x => Mode (m, is', x)) (cprods (map
       
   855             (fn (NONE, _) => [NONE]
       
   856               | (SOME js, arg) => map SOME (filter
       
   857                   (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
       
   858                     (iss ~~ args1)))
       
   859           in res end
       
   860         end)) (AList.lookup op = modes name)
       
   861   in case strip_comb t' of
       
   862     (Const (name, _), args) => the_default default (mk_modes name args)
       
   863     | (Var ((name, _), _), args) => the (mk_modes name args)
       
   864     | (Free (name, _), args) => the (mk_modes name args)
       
   865     | _ => default end
       
   866   
       
   867 and
       
   868 *)
       
   869 fun modes_of_term modes t =
       
   870   let
       
   871     val ks = 1 upto length (binder_types (fastype_of t));
       
   872     val default = [Mode (([], ks), ks, [])];
       
   873     fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
       
   874         let
       
   875           val (args1, args2) =
       
   876             if length args < length iss then
       
   877               error ("Too few arguments for inductive predicate " ^ name)
       
   878             else chop (length iss) args;
       
   879           val k = length args2;
       
   880           val prfx = 1 upto k
       
   881         in
       
   882           if not (is_prefix op = prfx is) then [] else
       
   883           let val is' = map (fn i => i - k) (List.drop (is, k))
       
   884           in map (fn x => Mode (m, is', x)) (cprods (map
       
   885             (fn (NONE, _) => [NONE]
       
   886               | (SOME js, arg) => map SOME (filter
       
   887                   (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
       
   888                     (iss ~~ args1)))
       
   889           end
       
   890         end)) (AList.lookup op = modes name)
       
   891 
       
   892   in
       
   893     case strip_comb (Envir.eta_contract t) of
       
   894       (Const (name, _), args) => the_default default (mk_modes name args)
       
   895     | (Var ((name, _), _), args) => the (mk_modes name args)
       
   896     | (Free (name, _), args) => the (mk_modes name args)
       
   897     | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
       
   898     | _ => default
       
   899   end
       
   900   
       
   901 fun select_mode_prem thy modes vs ps =
       
   902   find_first (is_some o snd) (ps ~~ map
       
   903     (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
       
   904           let
       
   905             val (in_ts, out_ts) = split_smode is us;
       
   906             val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
       
   907             val vTs = maps term_vTs out_ts';
       
   908             val dupTs = map snd (duplicates (op =) vTs) @
       
   909               map_filter (AList.lookup (op =) vTs) vs;
       
   910           in
       
   911             subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso
       
   912             forall (is_eqT o fastype_of) in_ts' andalso
       
   913             subset (op =) (term_vs t, vs) andalso
       
   914             forall is_eqT dupTs
       
   915           end)
       
   916             (modes_of_term modes t handle Option =>
       
   917                error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
       
   918       | Negprem (us, t) => find_first (fn Mode (_, is, _) =>
       
   919             length us = length is andalso
       
   920             subset (op =) (terms_vs us, vs) andalso
       
   921             subset (op =) (term_vs t, vs)
       
   922             (modes_of_term modes t handle Option =>
       
   923                error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
       
   924       | Sidecond t => if subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], []))
       
   925           else NONE
       
   926       ) ps);
       
   927 
       
   928 fun fold_prem f (Prem (args, _)) = fold f args
       
   929   | fold_prem f (Negprem (args, _)) = fold f args
       
   930   | fold_prem f (Sidecond t) = f t
       
   931 
       
   932 fun all_subsets [] = [[]]
       
   933   | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
       
   934 
       
   935 fun generator vTs v = 
       
   936   let
       
   937     val T = the (AList.lookup (op =) vTs v)
       
   938   in
       
   939     (Generator (v, T), Mode (([], []), [], []))
       
   940   end;
       
   941 
       
   942 fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t) 
       
   943   | gen_prem _ = error "gen_prem : invalid input for gen_prem"
       
   944 
       
   945 fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) =
       
   946   if member (op =) param_vs v then
       
   947     GeneratorPrem (us, t)
       
   948   else p  
       
   949   | param_gen_prem param_vs p = p
       
   950   
       
   951 fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) =
       
   952   let
       
   953     val modes' = modes @ map_filter
       
   954       (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
       
   955         (param_vs ~~ iss);
       
   956     val gen_modes' = gen_modes @ map_filter
       
   957       (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
       
   958         (param_vs ~~ iss);  
       
   959     val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
       
   960     val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
       
   961     fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
       
   962       | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
       
   963           NONE =>
       
   964             (if with_generator then
       
   965               (case select_mode_prem thy gen_modes' vs ps of
       
   966                   SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) 
       
   967                   (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
       
   968                   (filter_out (equal p) ps)
       
   969                 | NONE =>
       
   970                   let 
       
   971                     val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length))
       
   972                   in
       
   973                     case (find_first (fn generator_vs => is_some
       
   974                       (select_mode_prem thy modes' (union (op =) vs generator_vs) ps)) all_generator_vs) of
       
   975                       SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
       
   976                         (union (op =) vs generator_vs) ps
       
   977                     | NONE => NONE
       
   978                   end)
       
   979             else
       
   980               NONE)
       
   981         | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) 
       
   982             (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
       
   983             (filter_out (equal p) ps))
       
   984     val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
       
   985     val in_vs = terms_vs in_ts;
       
   986     val concl_vs = terms_vs ts
       
   987   in
       
   988     if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
       
   989     forall (is_eqT o fastype_of) in_ts' then
       
   990       case check_mode_prems [] (union (op =) param_vs in_vs) ps of
       
   991          NONE => NONE
       
   992        | SOME (acc_ps, vs) =>
       
   993          if with_generator then
       
   994            SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) 
       
   995          else
       
   996            if subset (op =) (concl_vs, vs) then SOME (ts, rev acc_ps) else NONE
       
   997     else NONE
       
   998   end;
       
   999 
       
  1000 fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
       
  1001   let val SOME rs = AList.lookup (op =) preds p
       
  1002   in (p, List.filter (fn m => case find_index
       
  1003     (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
       
  1004       ~1 => true
       
  1005     | i => (tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
       
  1006       p ^ " violates mode " ^ string_of_mode m); false)) ms)
       
  1007   end;
       
  1008 
       
  1009 fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
       
  1010   let
       
  1011     val SOME rs = AList.lookup (op =) preds p 
       
  1012   in
       
  1013     (p, map (fn m =>
       
  1014       (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms)
       
  1015   end;
       
  1016   
       
  1017 fun fixp f (x : (string * mode list) list) =
       
  1018   let val y = f x
       
  1019   in if x = y then x else fixp f y end;
       
  1020 
       
  1021 fun modes_of_arities arities =
       
  1022   (map (fn (s, (ks, k)) => (s, cprod (cprods (map
       
  1023             (fn NONE => [NONE]
       
  1024               | SOME k' => map SOME (subsets 1 k')) ks),
       
  1025             subsets 1 k))) arities)
       
  1026   
       
  1027 fun infer_modes with_generator thy extra_modes arities param_vs preds =
       
  1028   let
       
  1029     val modes =
       
  1030       fixp (fn modes =>
       
  1031         map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes)
       
  1032           (modes_of_arities arities)
       
  1033   in
       
  1034     map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes
       
  1035   end;
       
  1036 
       
  1037 fun remove_from rem [] = []
       
  1038   | remove_from rem ((k, vs) :: xs) =
       
  1039     (case AList.lookup (op =) rem k of
       
  1040       NONE => (k, vs)
       
  1041     | SOME vs' => (k, vs \\ vs'))
       
  1042     :: remove_from rem xs
       
  1043     
       
  1044 fun infer_modes_with_generator thy extra_modes arities param_vs preds =
       
  1045   let
       
  1046     val prednames = map fst preds
       
  1047     val extra_modes = all_modes_of thy
       
  1048     val gen_modes = all_generator_modes_of thy
       
  1049       |> filter_out (fn (name, _) => member (op =) prednames name)
       
  1050     val starting_modes = remove_from extra_modes (modes_of_arities arities) 
       
  1051     val modes =
       
  1052       fixp (fn modes =>
       
  1053         map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes)
       
  1054          starting_modes 
       
  1055   in
       
  1056     map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes
       
  1057   end;
       
  1058 
       
  1059 (* term construction *)
       
  1060 
       
  1061 fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
       
  1062       NONE => (Free (s, T), (names, (s, [])::vs))
       
  1063     | SOME xs =>
       
  1064         let
       
  1065           val s' = Name.variant names s;
       
  1066           val v = Free (s', T)
       
  1067         in
       
  1068           (v, (s'::names, AList.update (op =) (s, v::xs) vs))
       
  1069         end);
       
  1070 
       
  1071 fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
       
  1072   | distinct_v (t $ u) nvs =
       
  1073       let
       
  1074         val (t', nvs') = distinct_v t nvs;
       
  1075         val (u', nvs'') = distinct_v u nvs';
       
  1076       in (t' $ u', nvs'') end
       
  1077   | distinct_v x nvs = (x, nvs);
       
  1078 
       
  1079 fun compile_match thy compfuns eqs eqs' out_ts success_t =
       
  1080   let
       
  1081     val eqs'' = maps mk_eq eqs @ eqs'
       
  1082     val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
       
  1083     val name = Name.variant names "x";
       
  1084     val name' = Name.variant (name :: names) "y";
       
  1085     val T = mk_tupleT (map fastype_of out_ts);
       
  1086     val U = fastype_of success_t;
       
  1087     val U' = dest_predT compfuns U;
       
  1088     val v = Free (name, T);
       
  1089     val v' = Free (name', T);
       
  1090   in
       
  1091     lambda v (fst (Datatype.make_case
       
  1092       (ProofContext.init thy) false [] v
       
  1093       [(mk_tuple out_ts,
       
  1094         if null eqs'' then success_t
       
  1095         else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
       
  1096           foldr1 HOLogic.mk_conj eqs'' $ success_t $
       
  1097             mk_bot compfuns U'),
       
  1098        (v', mk_bot compfuns U')]))
       
  1099   end;
       
  1100 
       
  1101 (*FIXME function can be removed*)
       
  1102 fun mk_funcomp f t =
       
  1103   let
       
  1104     val names = Term.add_free_names t [];
       
  1105     val Ts = binder_types (fastype_of t);
       
  1106     val vs = map Free
       
  1107       (Name.variant_list names (replicate (length Ts) "x") ~~ Ts)
       
  1108   in
       
  1109     fold_rev lambda vs (f (list_comb (t, vs)))
       
  1110   end;
       
  1111 (*
       
  1112 fun compile_param_ext thy compfuns modes (NONE, t) = t
       
  1113   | compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
       
  1114       let
       
  1115         val (vs, u) = strip_abs t
       
  1116         val (ivs, ovs) = split_mode is vs    
       
  1117         val (f, args) = strip_comb u
       
  1118         val (params, args') = chop (length ms) args
       
  1119         val (inargs, outargs) = split_mode is' args'
       
  1120         val b = length vs
       
  1121         val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b)
       
  1122         val outp_perm =
       
  1123           snd (split_mode is perm)
       
  1124           |> map (fn i => i - length (filter (fn x => x < i) is'))
       
  1125         val names = [] -- TODO
       
  1126         val out_names = Name.variant_list names (replicate (length outargs) "x")
       
  1127         val f' = case f of
       
  1128             Const (name, T) =>
       
  1129               if AList.defined op = modes name then
       
  1130                 mk_predfun_of thy compfuns (name, T) (iss, is')
       
  1131               else error "compile param: Not an inductive predicate with correct mode"
       
  1132           | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is'))
       
  1133         val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f')))
       
  1134         val out_vs = map Free (out_names ~~ outTs)
       
  1135         val params' = map (compile_param thy modes) (ms ~~ params)
       
  1136         val f_app = list_comb (f', params' @ inargs)
       
  1137         val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm)))
       
  1138         val match_t = compile_match thy compfuns [] [] out_vs single_t
       
  1139       in list_abs (ivs,
       
  1140         mk_bind compfuns (f_app, match_t))
       
  1141       end
       
  1142   | compile_param_ext _ _ _ _ = error "compile params"
       
  1143 *)
       
  1144 
       
  1145 fun compile_param size thy compfuns (NONE, t) = t
       
  1146   | compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) =
       
  1147    let
       
  1148      val (f, args) = strip_comb (Envir.eta_contract t)
       
  1149      val (params, args') = chop (length ms) args
       
  1150      val params' = map (compile_param size thy compfuns) (ms ~~ params)
       
  1151      val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
       
  1152      val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
       
  1153      val f' =
       
  1154        case f of
       
  1155          Const (name, T) =>
       
  1156            mk_fun_of compfuns thy (name, T) (iss, is')
       
  1157        | Free (name, T) => Free (name, funT_of compfuns (iss, is') T)
       
  1158        | _ => error ("PredicateCompiler: illegal parameter term")
       
  1159    in list_comb (f', params' @ args') end
       
  1160    
       
  1161 fun compile_expr size thy ((Mode (mode, is, ms)), t) =
       
  1162   case strip_comb t of
       
  1163     (Const (name, T), params) =>
       
  1164        let
       
  1165          val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params)
       
  1166          val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
       
  1167        in
       
  1168          list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params')
       
  1169        end
       
  1170   | (Free (name, T), args) =>
       
  1171        let 
       
  1172          val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of 
       
  1173        in
       
  1174          list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args)
       
  1175        end;
       
  1176        
       
  1177 fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) =
       
  1178   case strip_comb t of
       
  1179     (Const (name, T), params) =>
       
  1180       let
       
  1181         val params' = map (compile_param size thy compfuns) (ms ~~ params)
       
  1182       in
       
  1183         list_comb (mk_generator_of compfuns thy (name, T) mode, params')
       
  1184       end
       
  1185     | (Free (name, T), args) =>
       
  1186       list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args)
       
  1187           
       
  1188 (** specific rpred functions -- move them to the correct place in this file *)
       
  1189 
       
  1190 (* uncommented termify code; causes more trouble than expected at first *) 
       
  1191 (*
       
  1192 fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t))
       
  1193   | mk_valtermify_term (Free (x, T)) = Free (x, termifyT T) 
       
  1194   | mk_valtermify_term (t1 $ t2) =
       
  1195     let
       
  1196       val T = fastype_of t1
       
  1197       val (T1, T2) = dest_funT T
       
  1198       val t1' = mk_valtermify_term t1
       
  1199       val t2' = mk_valtermify_term t2
       
  1200     in
       
  1201       Const ("Code_Evaluation.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
       
  1202     end
       
  1203   | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term"
       
  1204 *)
       
  1205 
       
  1206 fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
       
  1207   let
       
  1208     fun check_constrt t (names, eqs) =
       
  1209       if is_constrt thy t then (t, (names, eqs)) else
       
  1210         let
       
  1211           val s = Name.variant names "x";
       
  1212           val v = Free (s, fastype_of t)
       
  1213         in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
       
  1214 
       
  1215     val (in_ts, out_ts) = split_smode is ts;
       
  1216     val (in_ts', (all_vs', eqs)) =
       
  1217       fold_map check_constrt in_ts (all_vs, []);
       
  1218 
       
  1219     fun compile_prems out_ts' vs names [] =
       
  1220           let
       
  1221             val (out_ts'', (names', eqs')) =
       
  1222               fold_map check_constrt out_ts' (names, []);
       
  1223             val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
       
  1224               out_ts'' (names', map (rpair []) vs);
       
  1225           in
       
  1226           (* termify code:
       
  1227             compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
       
  1228               (mk_single compfuns (mk_tuple (map mk_valtermify_term out_ts)))
       
  1229            *)
       
  1230             compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
       
  1231               (final_term out_ts)
       
  1232           end
       
  1233       | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
       
  1234           let
       
  1235             val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
       
  1236             val (out_ts', (names', eqs)) =
       
  1237               fold_map check_constrt out_ts (names, [])
       
  1238             val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
       
  1239               out_ts' ((names', map (rpair []) vs))
       
  1240             val (compiled_clause, rest) = case p of
       
  1241                Prem (us, t) =>
       
  1242                  let
       
  1243                    val (in_ts, out_ts''') = split_smode is us;
       
  1244                    val args = case size of
       
  1245                      NONE => in_ts
       
  1246                    | SOME size_t => in_ts @ [size_t]
       
  1247                    val u = lift_pred compfuns
       
  1248                      (list_comb (compile_expr size thy (mode, t), args))                     
       
  1249                    val rest = compile_prems out_ts''' vs' names'' ps
       
  1250                  in
       
  1251                    (u, rest)
       
  1252                  end
       
  1253              | Negprem (us, t) =>
       
  1254                  let
       
  1255                    val (in_ts, out_ts''') = split_smode is us
       
  1256                    val u = lift_pred compfuns
       
  1257                      (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts)))
       
  1258                    val rest = compile_prems out_ts''' vs' names'' ps
       
  1259                  in
       
  1260                    (u, rest)
       
  1261                  end
       
  1262              | Sidecond t =>
       
  1263                  let
       
  1264                    val rest = compile_prems [] vs' names'' ps;
       
  1265                  in
       
  1266                    (mk_if compfuns t, rest)
       
  1267                  end
       
  1268              | GeneratorPrem (us, t) =>
       
  1269                  let
       
  1270                    val (in_ts, out_ts''') = split_smode is us;
       
  1271                    val args = case size of
       
  1272                      NONE => in_ts
       
  1273                    | SOME size_t => in_ts @ [size_t]
       
  1274                    val u = list_comb (compile_gen_expr size thy compfuns (mode, t), args)
       
  1275                    val rest = compile_prems out_ts''' vs' names'' ps
       
  1276                  in
       
  1277                    (u, rest)
       
  1278                  end
       
  1279              | Generator (v, T) =>
       
  1280                  let
       
  1281                    val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"})
       
  1282                    val rest = compile_prems [Free (v, T)]  vs' names'' ps;
       
  1283                  in
       
  1284                    (u, rest)
       
  1285                  end
       
  1286           in
       
  1287             compile_match thy compfuns constr_vs' eqs out_ts'' 
       
  1288               (mk_bind compfuns (compiled_clause, rest))
       
  1289           end
       
  1290     val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
       
  1291   in
       
  1292     mk_bind compfuns (mk_single compfuns inp, prem_t)
       
  1293   end
       
  1294 
       
  1295 fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls =
       
  1296   let
       
  1297     val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T)
       
  1298     val funT_of = if use_size then sizelim_funT_of else funT_of 
       
  1299     val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1
       
  1300     val xnames = Name.variant_list (all_vs @ param_vs)
       
  1301       (map (fn i => "x" ^ string_of_int i) (snd mode));
       
  1302     val size_name = Name.variant (all_vs @ param_vs @ xnames) "size"
       
  1303     (* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *)
       
  1304     val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1;
       
  1305     val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
       
  1306     val size = Free (size_name, @{typ "code_numeral"})
       
  1307     val decr_size =
       
  1308       if use_size then
       
  1309         SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
       
  1310           $ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"}))
       
  1311       else
       
  1312         NONE
       
  1313     val cl_ts =
       
  1314       map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts))
       
  1315         thy all_vs param_vs mode (mk_tuple xs)) moded_cls;
       
  1316     val t = foldr1 (mk_sup compfuns) cl_ts
       
  1317     val T' = mk_predT compfuns (mk_tupleT Us2)
       
  1318     val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
       
  1319       $ HOLogic.mk_eq (size, @{term "0 :: code_numeral"})
       
  1320       $ mk_bot compfuns (dest_predT compfuns T') $ t
       
  1321     val fun_const = mk_fun_of compfuns thy (s, T) mode
       
  1322     val eq = if use_size then
       
  1323       (list_comb (fun_const, params @ xs @ [size]), size_t)
       
  1324     else
       
  1325       (list_comb (fun_const, params @ xs), t)
       
  1326   in
       
  1327     HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
       
  1328   end;
       
  1329   
       
  1330 (* special setup for simpset *)                  
       
  1331 val HOL_basic_ss' = HOL_basic_ss setSolver 
       
  1332   (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
       
  1333 
       
  1334 (* Definition of executable functions and their intro and elim rules *)
       
  1335 
       
  1336 fun print_arities arities = tracing ("Arities:\n" ^
       
  1337   cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
       
  1338     space_implode " -> " (map
       
  1339       (fn NONE => "X" | SOME k' => string_of_int k')
       
  1340         (ks @ [SOME k]))) arities));
       
  1341 
       
  1342 fun mk_Eval_of ((x, T), NONE) names = (x, names)
       
  1343   | mk_Eval_of ((x, T), SOME mode) names = let
       
  1344   val Ts = binder_types T
       
  1345   val argnames = Name.variant_list names
       
  1346         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
       
  1347   val args = map Free (argnames ~~ Ts)
       
  1348   val (inargs, outargs) = split_smode mode args
       
  1349   val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs)
       
  1350   val t = fold_rev lambda args r 
       
  1351 in
       
  1352   (t, argnames @ names)
       
  1353 end;
       
  1354 
       
  1355 fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
       
  1356 let
       
  1357   val Ts = binder_types (fastype_of pred)
       
  1358   val funtrm = Const (mode_id, funT)
       
  1359   val argnames = Name.variant_list []
       
  1360         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
       
  1361   val (Ts1, Ts2) = chop (length iss) Ts;
       
  1362   val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
       
  1363   val args = map Free (argnames ~~ (Ts1' @ Ts2))
       
  1364   val (params, ioargs) = chop (length iss) args
       
  1365   val (inargs, outargs) = split_smode is ioargs
       
  1366   val param_names = Name.variant_list argnames
       
  1367     (map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
       
  1368   val param_vs = map Free (param_names ~~ Ts1)
       
  1369   val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) []
       
  1370   val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs))
       
  1371   val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs))
       
  1372   val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
       
  1373   val funargs = params @ inargs
       
  1374   val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
       
  1375                   if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
       
  1376   val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
       
  1377                    mk_tuple outargs))
       
  1378   val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
       
  1379   val simprules = [defthm, @{thm eval_pred},
       
  1380                    @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}]
       
  1381   val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
       
  1382   val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
       
  1383   val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
       
  1384   val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
       
  1385   val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac)
       
  1386 in 
       
  1387   (introthm, elimthm)
       
  1388 end;
       
  1389 
       
  1390 fun create_constname_of_mode thy prefix name mode = 
       
  1391   let
       
  1392     fun string_of_mode mode = if null mode then "0"
       
  1393       else space_implode "_" (map string_of_int mode)
       
  1394     val HOmode = space_implode "_and_"
       
  1395       (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) [])
       
  1396   in
       
  1397     (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^
       
  1398       (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode))
       
  1399   end;
       
  1400   
       
  1401 fun create_definitions preds (name, modes) thy =
       
  1402   let
       
  1403     val compfuns = PredicateCompFuns.compfuns
       
  1404     val T = AList.lookup (op =) preds name |> the
       
  1405     fun create_definition (mode as (iss, is)) thy = let
       
  1406       val mode_cname = create_constname_of_mode thy "" name mode
       
  1407       val mode_cbasename = Long_Name.base_name mode_cname
       
  1408       val Ts = binder_types T
       
  1409       val (Ts1, Ts2) = chop (length iss) Ts
       
  1410       val (Us1, Us2) =  split_smode is Ts2
       
  1411       val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
       
  1412       val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2))
       
  1413       val names = Name.variant_list []
       
  1414         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
       
  1415       val xs = map Free (names ~~ (Ts1' @ Ts2));                   
       
  1416       val (xparams, xargs) = chop (length iss) xs;
       
  1417       val (xins, xouts) = split_smode is xargs 
       
  1418       val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names
       
  1419       fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
       
  1420         | mk_split_lambda [x] t = lambda x t
       
  1421         | mk_split_lambda xs t =
       
  1422         let
       
  1423           fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
       
  1424             | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
       
  1425         in
       
  1426           mk_split_lambda' xs t
       
  1427         end;
       
  1428       val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
       
  1429         (list_comb (Const (name, T), xparams' @ xargs)))
       
  1430       val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
       
  1431       val def = Logic.mk_equals (lhs, predterm)
       
  1432       val ([definition], thy') = thy |>
       
  1433         Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
       
  1434         PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
       
  1435       val (intro, elim) =
       
  1436         create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
       
  1437       in thy' |> add_predfun name mode (mode_cname, definition, intro, elim)
       
  1438         |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
       
  1439         |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
       
  1440         |> Theory.checkpoint
       
  1441       end;
       
  1442   in
       
  1443     fold create_definition modes thy
       
  1444   end;
       
  1445 
       
  1446 fun sizelim_create_definitions preds (name, modes) thy =
       
  1447   let
       
  1448     val T = AList.lookup (op =) preds name |> the
       
  1449     fun create_definition mode thy =
       
  1450       let
       
  1451         val mode_cname = create_constname_of_mode thy "sizelim_" name mode
       
  1452         val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T
       
  1453       in
       
  1454         thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
       
  1455         |> set_sizelim_function_name name mode mode_cname 
       
  1456       end;
       
  1457   in
       
  1458     fold create_definition modes thy
       
  1459   end;
       
  1460     
       
  1461 fun rpred_create_definitions preds (name, modes) thy =
       
  1462   let
       
  1463     val T = AList.lookup (op =) preds name |> the
       
  1464     fun create_definition mode thy =
       
  1465       let
       
  1466         val mode_cname = create_constname_of_mode thy "gen_" name mode
       
  1467         val funT = sizelim_funT_of RPredCompFuns.compfuns mode T
       
  1468       in
       
  1469         thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
       
  1470         |> set_generator_name name mode mode_cname 
       
  1471       end;
       
  1472   in
       
  1473     fold create_definition modes thy
       
  1474   end;
       
  1475   
       
  1476 (* Proving equivalence of term *)
       
  1477 
       
  1478 fun is_Type (Type _) = true
       
  1479   | is_Type _ = false
       
  1480 
       
  1481 (* returns true if t is an application of an datatype constructor *)
       
  1482 (* which then consequently would be splitted *)
       
  1483 (* else false *)
       
  1484 fun is_constructor thy t =
       
  1485   if (is_Type (fastype_of t)) then
       
  1486     (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
       
  1487       NONE => false
       
  1488     | SOME info => (let
       
  1489       val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
       
  1490       val (c, _) = strip_comb t
       
  1491       in (case c of
       
  1492         Const (name, _) => name mem_string constr_consts
       
  1493         | _ => false) end))
       
  1494   else false
       
  1495 
       
  1496 (* MAJOR FIXME:  prove_params should be simple
       
  1497  - different form of introrule for parameters ? *)
       
  1498 fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
       
  1499   | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
       
  1500   let
       
  1501     val  (f, args) = strip_comb (Envir.eta_contract t)
       
  1502     val (params, _) = chop (length ms) args
       
  1503     val f_tac = case f of
       
  1504       Const (name, T) => simp_tac (HOL_basic_ss addsimps 
       
  1505          (@{thm eval_pred}::(predfun_definition_of thy name mode)::
       
  1506          @{thm "Product_Type.split_conv"}::[])) 1
       
  1507     | Free _ => TRY (rtac @{thm refl} 1)
       
  1508     | Abs _ => error "prove_param: No valid parameter term"
       
  1509   in
       
  1510     REPEAT_DETERM (etac @{thm thin_rl} 1)
       
  1511     THEN REPEAT_DETERM (rtac @{thm ext} 1)
       
  1512     THEN print_tac "prove_param"
       
  1513     THEN f_tac
       
  1514     THEN print_tac "after simplification in prove_args"
       
  1515     THEN (EVERY (map (prove_param thy) (ms ~~ params)))
       
  1516     THEN (REPEAT_DETERM (atac 1))
       
  1517   end
       
  1518 
       
  1519 fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
       
  1520   case strip_comb t of
       
  1521     (Const (name, T), args) =>  
       
  1522       let
       
  1523         val introrule = predfun_intro_of thy name mode
       
  1524         val (args1, args2) = chop (length ms) args
       
  1525       in
       
  1526         rtac @{thm bindI} 1
       
  1527         THEN print_tac "before intro rule:"
       
  1528         (* for the right assumption in first position *)
       
  1529         THEN rotate_tac premposition 1
       
  1530         THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
       
  1531         THEN rtac introrule 1
       
  1532         THEN print_tac "after intro rule"
       
  1533         (* work with parameter arguments *)
       
  1534         THEN (atac 1)
       
  1535         THEN (print_tac "parameter goal")
       
  1536         THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
       
  1537         THEN (REPEAT_DETERM (atac 1))
       
  1538       end
       
  1539   | _ => rtac @{thm bindI} 1 THEN atac 1
       
  1540 
       
  1541 fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
       
  1542 
       
  1543 fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
       
  1544 
       
  1545 fun prove_match thy (out_ts : term list) = let
       
  1546   fun get_case_rewrite t =
       
  1547     if (is_constructor thy t) then let
       
  1548       val case_rewrites = (#case_rewrites (Datatype.the_info thy
       
  1549         ((fst o dest_Type o fastype_of) t)))
       
  1550       in case_rewrites @ maps get_case_rewrite (snd (strip_comb t)) end
       
  1551     else []
       
  1552   val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: maps get_case_rewrite out_ts
       
  1553 (* replace TRY by determining if it necessary - are there equations when calling compile match? *)
       
  1554 in
       
  1555    (* make this simpset better! *)
       
  1556   asm_simp_tac (HOL_basic_ss' addsimps simprules) 1
       
  1557   THEN print_tac "after prove_match:"
       
  1558   THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
       
  1559          THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
       
  1560          THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
       
  1561   THEN print_tac "after if simplification"
       
  1562 end;
       
  1563 
       
  1564 (* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
       
  1565 
       
  1566 fun prove_sidecond thy modes t =
       
  1567   let
       
  1568     fun preds_of t nameTs = case strip_comb t of 
       
  1569       (f as Const (name, T), args) =>
       
  1570         if AList.defined (op =) modes name then (name, T) :: nameTs
       
  1571           else fold preds_of args nameTs
       
  1572       | _ => nameTs
       
  1573     val preds = preds_of t []
       
  1574     val defs = map
       
  1575       (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
       
  1576         preds
       
  1577   in 
       
  1578     (* remove not_False_eq_True when simpset in prove_match is better *)
       
  1579     simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1 
       
  1580     (* need better control here! *)
       
  1581   end
       
  1582 
       
  1583 fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
       
  1584   let
       
  1585     val (in_ts, clause_out_ts) = split_smode is ts;
       
  1586     fun prove_prems out_ts [] =
       
  1587       (prove_match thy out_ts)
       
  1588       THEN asm_simp_tac HOL_basic_ss' 1
       
  1589       THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
       
  1590     | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
       
  1591       let
       
  1592         val premposition = (find_index (equal p) clauses) + nargs
       
  1593         val rest_tac = (case p of Prem (us, t) =>
       
  1594             let
       
  1595               val (_, out_ts''') = split_smode is us
       
  1596               val rec_tac = prove_prems out_ts''' ps
       
  1597             in
       
  1598               print_tac "before clause:"
       
  1599               THEN asm_simp_tac HOL_basic_ss 1
       
  1600               THEN print_tac "before prove_expr:"
       
  1601               THEN prove_expr thy (mode, t, us) premposition
       
  1602               THEN print_tac "after prove_expr:"
       
  1603               THEN rec_tac
       
  1604             end
       
  1605           | Negprem (us, t) =>
       
  1606             let
       
  1607               val (_, out_ts''') = split_smode is us
       
  1608               val rec_tac = prove_prems out_ts''' ps
       
  1609               val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
       
  1610               val (_, params) = strip_comb t
       
  1611             in
       
  1612               rtac @{thm bindI} 1
       
  1613               THEN (if (is_some name) then
       
  1614                   simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
       
  1615                   THEN rtac @{thm not_predI} 1
       
  1616                   THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
       
  1617                   THEN (REPEAT_DETERM (atac 1))
       
  1618                   (* FIXME: work with parameter arguments *)
       
  1619                   THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
       
  1620                 else
       
  1621                   rtac @{thm not_predI'} 1)
       
  1622                   THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
       
  1623               THEN rec_tac
       
  1624             end
       
  1625           | Sidecond t =>
       
  1626            rtac @{thm bindI} 1
       
  1627            THEN rtac @{thm if_predI} 1
       
  1628            THEN print_tac "before sidecond:"
       
  1629            THEN prove_sidecond thy modes t
       
  1630            THEN print_tac "after sidecond:"
       
  1631            THEN prove_prems [] ps)
       
  1632       in (prove_match thy out_ts)
       
  1633           THEN rest_tac
       
  1634       end;
       
  1635     val prems_tac = prove_prems in_ts moded_ps
       
  1636   in
       
  1637     rtac @{thm bindI} 1
       
  1638     THEN rtac @{thm singleI} 1
       
  1639     THEN prems_tac
       
  1640   end;
       
  1641 
       
  1642 fun select_sup 1 1 = []
       
  1643   | select_sup _ 1 = [rtac @{thm supI1}]
       
  1644   | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
       
  1645 
       
  1646 fun prove_one_direction thy clauses preds modes pred mode moded_clauses =
       
  1647   let
       
  1648     val T = the (AList.lookup (op =) preds pred)
       
  1649     val nargs = length (binder_types T) - nparams_of thy pred
       
  1650     val pred_case_rule = the_elim_of thy pred
       
  1651   in
       
  1652     REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
       
  1653     THEN etac (predfun_elim_of thy pred mode) 1
       
  1654     THEN etac pred_case_rule 1
       
  1655     THEN (EVERY (map
       
  1656            (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
       
  1657              (1 upto (length moded_clauses))))
       
  1658     THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses))
       
  1659     THEN print_tac "proved one direction"
       
  1660   end;
       
  1661 
       
  1662 (** Proof in the other direction **)
       
  1663 
       
  1664 fun prove_match2 thy out_ts = let
       
  1665   fun split_term_tac (Free _) = all_tac
       
  1666     | split_term_tac t =
       
  1667       if (is_constructor thy t) then let
       
  1668         val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
       
  1669         val num_of_constrs = length (#case_rewrites info)
       
  1670         (* special treatment of pairs -- because of fishing *)
       
  1671         val split_rules = case (fst o dest_Type o fastype_of) t of
       
  1672           "*" => [@{thm prod.split_asm}] 
       
  1673           | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm")
       
  1674         val (_, ts) = strip_comb t
       
  1675       in
       
  1676         (Splitter.split_asm_tac split_rules 1)
       
  1677 (*        THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
       
  1678           THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *)
       
  1679         THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2))
       
  1680         THEN (EVERY (map split_term_tac ts))
       
  1681       end
       
  1682     else all_tac
       
  1683   in
       
  1684     split_term_tac (mk_tuple out_ts)
       
  1685     THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
       
  1686   end
       
  1687 
       
  1688 (* VERY LARGE SIMILIRATIY to function prove_param 
       
  1689 -- join both functions
       
  1690 *)
       
  1691 (* TODO: remove function *)
       
  1692 
       
  1693 fun prove_param2 thy (NONE, t) = all_tac 
       
  1694   | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
       
  1695     val  (f, args) = strip_comb (Envir.eta_contract t)
       
  1696     val (params, _) = chop (length ms) args
       
  1697     val f_tac = case f of
       
  1698         Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
       
  1699            (@{thm eval_pred}::(predfun_definition_of thy name mode)
       
  1700            :: @{thm "Product_Type.split_conv"}::[])) 1
       
  1701       | Free _ => all_tac
       
  1702       | _ => error "prove_param2: illegal parameter term"
       
  1703   in  
       
  1704     print_tac "before simplification in prove_args:"
       
  1705     THEN f_tac
       
  1706     THEN print_tac "after simplification in prove_args"
       
  1707     THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
       
  1708   end
       
  1709 
       
  1710 
       
  1711 fun prove_expr2 thy (Mode (mode, is, ms), t) = 
       
  1712   (case strip_comb t of
       
  1713     (Const (name, T), args) =>
       
  1714       etac @{thm bindE} 1
       
  1715       THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
       
  1716       THEN print_tac "prove_expr2-before"
       
  1717       THEN (debug_tac (Syntax.string_of_term_global thy
       
  1718         (prop_of (predfun_elim_of thy name mode))))
       
  1719       THEN (etac (predfun_elim_of thy name mode) 1)
       
  1720       THEN print_tac "prove_expr2"
       
  1721       THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
       
  1722       THEN print_tac "finished prove_expr2"      
       
  1723     | _ => etac @{thm bindE} 1)
       
  1724     
       
  1725 (* FIXME: what is this for? *)
       
  1726 (* replace defined by has_mode thy pred *)
       
  1727 (* TODO: rewrite function *)
       
  1728 fun prove_sidecond2 thy modes t = let
       
  1729   fun preds_of t nameTs = case strip_comb t of 
       
  1730     (f as Const (name, T), args) =>
       
  1731       if AList.defined (op =) modes name then (name, T) :: nameTs
       
  1732         else fold preds_of args nameTs
       
  1733     | _ => nameTs
       
  1734   val preds = preds_of t []
       
  1735   val defs = map
       
  1736     (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
       
  1737       preds
       
  1738   in
       
  1739    (* only simplify the one assumption *)
       
  1740    full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
       
  1741    (* need better control here! *)
       
  1742    THEN print_tac "after sidecond2 simplification"
       
  1743    end
       
  1744   
       
  1745 fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
       
  1746   let
       
  1747     val pred_intro_rule = nth (intros_of thy pred) (i - 1)
       
  1748     val (in_ts, clause_out_ts) = split_smode is ts;
       
  1749     fun prove_prems2 out_ts [] =
       
  1750       print_tac "before prove_match2 - last call:"
       
  1751       THEN prove_match2 thy out_ts
       
  1752       THEN print_tac "after prove_match2 - last call:"
       
  1753       THEN (etac @{thm singleE} 1)
       
  1754       THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
       
  1755       THEN (asm_full_simp_tac HOL_basic_ss' 1)
       
  1756       THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
       
  1757       THEN (asm_full_simp_tac HOL_basic_ss' 1)
       
  1758       THEN SOLVED (print_tac "state before applying intro rule:"
       
  1759       THEN (rtac pred_intro_rule 1)
       
  1760       (* How to handle equality correctly? *)
       
  1761       THEN (print_tac "state before assumption matching")
       
  1762       THEN (REPEAT (atac 1 ORELSE 
       
  1763          (CHANGED (asm_full_simp_tac HOL_basic_ss' 1)
       
  1764           THEN print_tac "state after simp_tac:"))))
       
  1765     | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
       
  1766       let
       
  1767         val rest_tac = (case p of
       
  1768           Prem (us, t) =>
       
  1769           let
       
  1770             val (_, out_ts''') = split_smode is us
       
  1771             val rec_tac = prove_prems2 out_ts''' ps
       
  1772           in
       
  1773             (prove_expr2 thy (mode, t)) THEN rec_tac
       
  1774           end
       
  1775         | Negprem (us, t) =>
       
  1776           let
       
  1777             val (_, out_ts''') = split_smode is us
       
  1778             val rec_tac = prove_prems2 out_ts''' ps
       
  1779             val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
       
  1780             val (_, params) = strip_comb t
       
  1781           in
       
  1782             print_tac "before neg prem 2"
       
  1783             THEN etac @{thm bindE} 1
       
  1784             THEN (if is_some name then
       
  1785                 full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 
       
  1786                 THEN etac @{thm not_predE} 1
       
  1787                 THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
       
  1788                 THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
       
  1789               else
       
  1790                 etac @{thm not_predE'} 1)
       
  1791             THEN rec_tac
       
  1792           end 
       
  1793         | Sidecond t =>
       
  1794           etac @{thm bindE} 1
       
  1795           THEN etac @{thm if_predE} 1
       
  1796           THEN prove_sidecond2 thy modes t 
       
  1797           THEN prove_prems2 [] ps)
       
  1798       in print_tac "before prove_match2:"
       
  1799          THEN prove_match2 thy out_ts
       
  1800          THEN print_tac "after prove_match2:"
       
  1801          THEN rest_tac
       
  1802       end;
       
  1803     val prems_tac = prove_prems2 in_ts ps 
       
  1804   in
       
  1805     print_tac "starting prove_clause2"
       
  1806     THEN etac @{thm bindE} 1
       
  1807     THEN (etac @{thm singleE'} 1)
       
  1808     THEN (TRY (etac @{thm Pair_inject} 1))
       
  1809     THEN print_tac "after singleE':"
       
  1810     THEN prems_tac
       
  1811   end;
       
  1812  
       
  1813 fun prove_other_direction thy modes pred mode moded_clauses =
       
  1814   let
       
  1815     fun prove_clause clause i =
       
  1816       (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
       
  1817       THEN (prove_clause2 thy modes pred mode clause i)
       
  1818   in
       
  1819     (DETERM (TRY (rtac @{thm unit.induct} 1)))
       
  1820      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
       
  1821      THEN (rtac (predfun_intro_of thy pred mode) 1)
       
  1822      THEN (REPEAT_DETERM (rtac @{thm refl} 2))
       
  1823      THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
       
  1824   end;
       
  1825 
       
  1826 (** proof procedure **)
       
  1827 
       
  1828 fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
       
  1829   let
       
  1830     val ctxt = ProofContext.init thy
       
  1831     val clauses = the (AList.lookup (op =) clauses pred)
       
  1832   in
       
  1833     Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
       
  1834       (if !do_proofs then
       
  1835         (fn _ =>
       
  1836         rtac @{thm pred_iffI} 1
       
  1837         THEN prove_one_direction thy clauses preds modes pred mode moded_clauses
       
  1838         THEN print_tac "proved one direction"
       
  1839         THEN prove_other_direction thy modes pred mode moded_clauses
       
  1840         THEN print_tac "proved other direction")
       
  1841        else (fn _ => mycheat_tac thy 1))
       
  1842   end;
       
  1843 
       
  1844 (* composition of mode inference, definition, compilation and proof *)
       
  1845 
       
  1846 (** auxillary combinators for table of preds and modes **)
       
  1847 
       
  1848 fun map_preds_modes f preds_modes_table =
       
  1849   map (fn (pred, modes) =>
       
  1850     (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
       
  1851 
       
  1852 fun join_preds_modes table1 table2 =
       
  1853   map_preds_modes (fn pred => fn mode => fn value =>
       
  1854     (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
       
  1855     
       
  1856 fun maps_modes preds_modes_table =
       
  1857   map (fn (pred, modes) =>
       
  1858     (pred, map (fn (mode, value) => value) modes)) preds_modes_table  
       
  1859     
       
  1860 fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses =
       
  1861   map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred
       
  1862       (the (AList.lookup (op =) preds pred))) moded_clauses  
       
  1863   
       
  1864 fun prove thy clauses preds modes moded_clauses compiled_terms =
       
  1865   map_preds_modes (prove_pred thy clauses preds modes)
       
  1866     (join_preds_modes moded_clauses compiled_terms)
       
  1867 
       
  1868 fun prove_by_skip thy _ _ _ _ compiled_terms =
       
  1869   map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
       
  1870     compiled_terms
       
  1871     
       
  1872 fun prepare_intrs thy prednames =
       
  1873   let
       
  1874     val intrs = maps (intros_of thy) prednames
       
  1875       |> map (Logic.unvarify o prop_of)
       
  1876     val nparams = nparams_of thy (hd prednames)
       
  1877     val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
       
  1878     val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
       
  1879     val _ $ u = Logic.strip_imp_concl (hd intrs);
       
  1880     val params = List.take (snd (strip_comb u), nparams);
       
  1881     val param_vs = maps term_vs params
       
  1882     val all_vs = terms_vs intrs
       
  1883     fun dest_prem t =
       
  1884       (case strip_comb t of
       
  1885         (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t
       
  1886       | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of          
       
  1887           Prem (ts, t) => Negprem (ts, t)
       
  1888         | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) 
       
  1889         | Sidecond t => Sidecond (c $ t))
       
  1890       | (c as Const (s, _), ts) =>
       
  1891         if is_registered thy s then
       
  1892           let val (ts1, ts2) = chop (nparams_of thy s) ts
       
  1893           in Prem (ts2, list_comb (c, ts1)) end
       
  1894         else Sidecond t
       
  1895       | _ => Sidecond t)
       
  1896     fun add_clause intr (clauses, arities) =
       
  1897     let
       
  1898       val _ $ t = Logic.strip_imp_concl intr;
       
  1899       val (Const (name, T), ts) = strip_comb t;
       
  1900       val (ts1, ts2) = chop nparams ts;
       
  1901       val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
       
  1902       val (Ts, Us) = chop nparams (binder_types T)
       
  1903     in
       
  1904       (AList.update op = (name, these (AList.lookup op = clauses name) @
       
  1905         [(ts2, prems)]) clauses,
       
  1906        AList.update op = (name, (map (fn U => (case strip_type U of
       
  1907                  (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
       
  1908                | _ => NONE)) Ts,
       
  1909              length Us)) arities)
       
  1910     end;
       
  1911     val (clauses, arities) = fold add_clause intrs ([], []);
       
  1912   in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end;
       
  1913 
       
  1914 (** main function of predicate compiler **)
       
  1915 
       
  1916 fun add_equations_of steps prednames thy =
       
  1917   let
       
  1918     val _ = tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
       
  1919     val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
       
  1920       prepare_intrs thy prednames
       
  1921     val _ = tracing "Infering modes..."
       
  1922     val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses 
       
  1923     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
       
  1924     val _ = print_modes modes
       
  1925     val _ = print_moded_clauses thy moded_clauses
       
  1926     val _ = tracing "Defining executable functions..."
       
  1927     val thy' = fold (#create_definitions steps preds) modes thy
       
  1928       |> Theory.checkpoint
       
  1929     val _ = tracing "Compiling equations..."
       
  1930     val compiled_terms =
       
  1931       (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
       
  1932     val _ = print_compiled_terms thy' compiled_terms
       
  1933     val _ = tracing "Proving equations..."
       
  1934     val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
       
  1935       moded_clauses compiled_terms
       
  1936     val qname = #qname steps
       
  1937     (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *)
       
  1938     val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
       
  1939       (fn thm => Context.mapping (Code.add_eqn thm) I))))
       
  1940     val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
       
  1941       [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
       
  1942         [attrib thy ])] thy))
       
  1943       (maps_modes result_thms) thy'
       
  1944       |> Theory.checkpoint
       
  1945   in
       
  1946     thy''
       
  1947   end
       
  1948 
       
  1949 fun extend' value_of edges_of key (G, visited) =
       
  1950   let
       
  1951     val (G', v) = case try (Graph.get_node G) key of
       
  1952         SOME v => (G, v)
       
  1953       | NONE => (Graph.new_node (key, value_of key) G, value_of key)
       
  1954     val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited)
       
  1955       (G', key :: visited) 
       
  1956   in
       
  1957     (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
       
  1958   end;
       
  1959 
       
  1960 fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
       
  1961   
       
  1962 fun gen_add_equations steps names thy =
       
  1963   let
       
  1964     val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
       
  1965       |> Theory.checkpoint;
       
  1966     fun strong_conn_of gr keys =
       
  1967       Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
       
  1968     val scc = strong_conn_of (PredData.get thy') names
       
  1969     val thy'' = fold_rev
       
  1970       (fn preds => fn thy =>
       
  1971         if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy)
       
  1972       scc thy' |> Theory.checkpoint
       
  1973   in thy'' end
       
  1974 
       
  1975 (* different instantiantions of the predicate compiler *)
       
  1976 
       
  1977 val add_equations = gen_add_equations
       
  1978   {infer_modes = infer_modes false,
       
  1979   create_definitions = create_definitions,
       
  1980   compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
       
  1981   prove = prove,
       
  1982   are_not_defined = (fn thy => forall (null o modes_of thy)),
       
  1983   qname = "equation"}
       
  1984 
       
  1985 val add_sizelim_equations = gen_add_equations
       
  1986   {infer_modes = infer_modes false,
       
  1987   create_definitions = sizelim_create_definitions,
       
  1988   compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
       
  1989   prove = prove_by_skip,
       
  1990   are_not_defined = (fn thy => fn preds => true), (* TODO *)
       
  1991   qname = "sizelim_equation"
       
  1992   }
       
  1993   
       
  1994 val add_quickcheck_equations = gen_add_equations
       
  1995   {infer_modes = infer_modes_with_generator,
       
  1996   create_definitions = rpred_create_definitions,
       
  1997   compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true,
       
  1998   prove = prove_by_skip,
       
  1999   are_not_defined = (fn thy => fn preds => true), (* TODO *)
       
  2000   qname = "rpred_equation"}
       
  2001 
       
  2002 (** user interface **)
       
  2003 
       
  2004 (* generation of case rules from user-given introduction rules *)
       
  2005 
       
  2006 fun mk_casesrule ctxt nparams introrules =
       
  2007   let
       
  2008     val intros = map (Logic.unvarify o prop_of) introrules
       
  2009     val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
       
  2010     val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt
       
  2011     val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
       
  2012     val (argnames, ctxt2) = Variable.variant_fixes
       
  2013       (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
       
  2014     val argvs = map2 (curry Free) argnames (map fastype_of args)
       
  2015     fun mk_case intro =
       
  2016       let
       
  2017         val (_, (_, args)) = strip_intro_concl nparams intro
       
  2018         val prems = Logic.strip_imp_prems intro
       
  2019         val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args)
       
  2020         val frees = (fold o fold_aterms)
       
  2021           (fn t as Free _ =>
       
  2022               if member (op aconv) params t then I else insert (op aconv) t
       
  2023            | _ => I) (args @ prems) []
       
  2024       in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
       
  2025     val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
       
  2026     val cases = map mk_case intros
       
  2027   in Logic.list_implies (assm :: cases, prop) end;
       
  2028 
       
  2029 (* code_pred_intro attribute *)
       
  2030 
       
  2031 fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
       
  2032 
       
  2033 val code_pred_intros_attrib = attrib add_intro;
       
  2034 
       
  2035 local
       
  2036 
       
  2037 (* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
       
  2038 (* TODO: must create state to prove multiple cases *)
       
  2039 fun generic_code_pred prep_const raw_const lthy =
       
  2040   let
       
  2041     val thy = ProofContext.theory_of lthy
       
  2042     val const = prep_const thy raw_const
       
  2043     val lthy' = LocalTheory.theory (PredData.map
       
  2044         (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
       
  2045       |> LocalTheory.checkpoint
       
  2046     val thy' = ProofContext.theory_of lthy'
       
  2047     val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
       
  2048     fun mk_cases const =
       
  2049       let
       
  2050         val nparams = nparams_of thy' const
       
  2051         val intros = intros_of thy' const
       
  2052       in mk_casesrule lthy' nparams intros end  
       
  2053     val cases_rules = map mk_cases preds
       
  2054     val cases =
       
  2055       map (fn case_rule => RuleCases.Case {fixes = [],
       
  2056         assumes = [("", Logic.strip_imp_prems case_rule)],
       
  2057         binds = [], cases = []}) cases_rules
       
  2058     val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
       
  2059     val lthy'' = lthy'
       
  2060       |> fold Variable.auto_fixes cases_rules 
       
  2061       |> ProofContext.add_cases true case_env
       
  2062     fun after_qed thms goal_ctxt =
       
  2063       let
       
  2064         val global_thms = ProofContext.export goal_ctxt
       
  2065           (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
       
  2066       in
       
  2067         goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const])
       
  2068       end  
       
  2069   in
       
  2070     Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
       
  2071   end;
       
  2072 
       
  2073 structure P = OuterParse
       
  2074 
       
  2075 in
       
  2076 
       
  2077 val code_pred = generic_code_pred (K I);
       
  2078 val code_pred_cmd = generic_code_pred Code.read_const
       
  2079 
       
  2080 val setup = PredData.put (Graph.empty) #>
       
  2081   Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro))
       
  2082     "adding alternative introduction rules for code generation of inductive predicates"
       
  2083 (*  Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib)
       
  2084     "adding alternative elimination rules for code generation of inductive predicates";
       
  2085     *)
       
  2086   (*FIXME name discrepancy in attribs and ML code*)
       
  2087   (*FIXME intros should be better named intro*)
       
  2088   (*FIXME why distinguished attribute for cases?*)
       
  2089 
       
  2090 val _ = OuterSyntax.local_theory_to_proof "code_pred"
       
  2091   "prove equations for predicate specified by intro/elim rules"
       
  2092   OuterKeyword.thy_goal (P.term_group >> code_pred_cmd)
       
  2093 
       
  2094 end
       
  2095 
       
  2096 (*FIXME
       
  2097 - Naming of auxiliary rules necessary?
       
  2098 - add default code equations P x y z = P_i_i_i x y z
       
  2099 *)
       
  2100 
       
  2101 (* transformation for code generation *)
       
  2102 
       
  2103 val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option);
       
  2104 
       
  2105 (*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
       
  2106 fun analyze_compr thy t_compr =
       
  2107   let
       
  2108     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
       
  2109       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
       
  2110     val (body, Ts, fp) = HOLogic.strip_psplits split;
       
  2111     val (pred as Const (name, T), all_args) = strip_comb body;
       
  2112     val (params, args) = chop (nparams_of thy name) all_args;
       
  2113     val user_mode = map_filter I (map_index
       
  2114       (fn (i, t) => case t of Bound j => if j < length Ts then NONE
       
  2115         else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*)
       
  2116     val modes = filter (fn Mode (_, is, _) => is = user_mode)
       
  2117       (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
       
  2118     val m = case modes
       
  2119      of [] => error ("No mode possible for comprehension "
       
  2120                 ^ Syntax.string_of_term_global thy t_compr)
       
  2121       | [m] => m
       
  2122       | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
       
  2123                 ^ Syntax.string_of_term_global thy t_compr); m);
       
  2124     val (inargs, outargs) = split_smode user_mode args;
       
  2125     val t_pred = list_comb (compile_expr NONE thy (m, list_comb (pred, params)), inargs);
       
  2126     val t_eval = if null outargs then t_pred else let
       
  2127         val outargs_bounds = map (fn Bound i => i) outargs;
       
  2128         val outargsTs = map (nth Ts) outargs_bounds;
       
  2129         val T_pred = HOLogic.mk_tupleT outargsTs;
       
  2130         val T_compr = HOLogic.mk_ptupleT fp Ts;
       
  2131         val arrange_bounds = map_index I outargs_bounds
       
  2132           |> sort (prod_ord (K EQUAL) int_ord)
       
  2133           |> map fst;
       
  2134         val arrange = funpow (length outargs_bounds - 1) HOLogic.mk_split
       
  2135           (Term.list_abs (map (pair "") outargsTs,
       
  2136             HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds)))
       
  2137       in mk_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end
       
  2138   in t_eval end;
       
  2139 
       
  2140 fun eval thy t_compr =
       
  2141   let
       
  2142     val t = analyze_compr thy t_compr;
       
  2143     val T = dest_predT PredicateCompFuns.compfuns (fastype_of t);
       
  2144     val t' = mk_map PredicateCompFuns.compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
       
  2145   in (T, Code_ML.eval NONE ("Predicate_Compile.eval_ref", eval_ref) Predicate.map thy t' []) end;
       
  2146 
       
  2147 fun values ctxt k t_compr =
       
  2148   let
       
  2149     val thy = ProofContext.theory_of ctxt;
       
  2150     val (T, t) = eval thy t_compr;
       
  2151     val setT = HOLogic.mk_setT T;
       
  2152     val (ts, _) = Predicate.yieldn k t;
       
  2153     val elemsT = HOLogic.mk_set T ts;
       
  2154   in if k = ~1 orelse length ts < k then elemsT
       
  2155     else Const (@{const_name Lattices.sup}, setT --> setT --> setT) $ elemsT $ t_compr
       
  2156   end;
       
  2157 
       
  2158 fun values_cmd modes k raw_t state =
       
  2159   let
       
  2160     val ctxt = Toplevel.context_of state;
       
  2161     val t = Syntax.read_term ctxt raw_t;
       
  2162     val t' = values ctxt k t;
       
  2163     val ty' = Term.type_of t';
       
  2164     val ctxt' = Variable.auto_fixes t' ctxt;
       
  2165     val p = PrintMode.with_modes modes (fn () =>
       
  2166       Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
       
  2167         Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
       
  2168   in Pretty.writeln p end;
       
  2169 
       
  2170 local structure P = OuterParse in
       
  2171 
       
  2172 val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
       
  2173 
       
  2174 val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag
       
  2175   (opt_modes -- Scan.optional P.nat ~1 -- P.term
       
  2176     >> (fn ((modes, k), t) => Toplevel.no_timing o Toplevel.keep
       
  2177         (values_cmd modes k t)));
       
  2178 
       
  2179 end;
       
  2180 
       
  2181 end;
       
  2182