src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
author panny
Fri, 30 Aug 2013 21:43:15 +0200
changeset 54472 585b2fee55e5
parent 54466 c31c0c311cf0
child 54478 63015d035301
permissions -rw-r--r--
fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
     1 (*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar.ML
     2     Author:     Lorenz Panny, TU Muenchen
     3     Copyright   2013
     4 
     5 Recursor and corecursor sugar.
     6 *)
     7 
     8 signature BNF_FP_REC_SUGAR =
     9 sig
    10   val add_primrec_cmd: (binding * string option * mixfix) list ->
    11     (Attrib.binding * string) list -> local_theory -> local_theory;
    12   val add_primcorec_cmd: bool ->
    13     (binding * string option * mixfix) list * (Attrib.binding * string) list -> Proof.context ->
    14     Proof.state
    15 end;
    16 
    17 structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
    18 struct
    19 
    20 open BNF_Util
    21 open BNF_FP_Util
    22 open BNF_FP_Rec_Sugar_Util
    23 open BNF_FP_Rec_Sugar_Tactics
    24 
    25 exception Primrec_Error of string * term list;
    26 
    27 fun primrec_error str = raise Primrec_Error (str, []);
    28 fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
    29 fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
    30 
    31 fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
    32 
    33 val simp_attrs = @{attributes [simp]};
    34 
    35 
    36 
    37 (* Primrec *)
    38 
    39 type eqn_data = {
    40   fun_name: string,
    41   rec_type: typ,
    42   ctr: term,
    43   ctr_args: term list,
    44   left_args: term list,
    45   right_args: term list,
    46   res_type: typ,
    47   rhs_term: term,
    48   user_eqn: term
    49 };
    50 
    51 fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
    52   |> fold (K (fn u => Abs (Name.uu, dummyT, u))) (0 upto n);
    53 
    54 fun dissect_eqn lthy fun_names eqn' =
    55   let
    56     val eqn = subst_bounds (strip_qnt_vars @{const_name all} eqn' |> map Free |> rev,
    57         strip_qnt_body @{const_name all} eqn') |> HOLogic.dest_Trueprop
    58         handle TERM _ =>
    59           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    60     val (lhs, rhs) = HOLogic.dest_eq eqn
    61         handle TERM _ =>
    62           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    63     val (fun_name, args) = strip_comb lhs
    64       |>> (fn x => if is_Free x then fst (dest_Free x)
    65           else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
    66     val (left_args, rest) = take_prefix is_Free args;
    67     val (nonfrees, right_args) = take_suffix is_Free rest;
    68     val _ = length nonfrees = 1 orelse if length nonfrees = 0 then
    69       primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
    70       primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
    71     val _ = member (op =) fun_names fun_name orelse
    72       primrec_error_eqn "malformed function equation (does not start with function name)" eqn
    73 
    74     val (ctr, ctr_args) = strip_comb (the_single nonfrees);
    75     val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
    76       primrec_error_eqn "partially applied constructor in pattern" eqn;
    77     val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
    78       primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
    79         "\" in left-hand side") eqn end;
    80     val _ = forall is_Free ctr_args orelse
    81       primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
    82     val _ =
    83       let val b = fold_aterms (fn x as Free (v, _) =>
    84         if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
    85         not (member (op =) fun_names v) andalso
    86         not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
    87       in
    88         null b orelse
    89         primrec_error_eqn ("extra variable(s) in right-hand side: " ^
    90           commas (map (Syntax.string_of_term lthy) b)) eqn
    91       end;
    92   in
    93     {fun_name = fun_name,
    94      rec_type = body_type (type_of ctr),
    95      ctr = ctr,
    96      ctr_args = ctr_args,
    97      left_args = left_args,
    98      right_args = right_args,
    99      res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
   100      rhs_term = rhs,
   101      user_eqn = eqn'}
   102   end;
   103 
   104 (* substitutes (f ls x rs) by (y ls rs) for all f: get_idx f \<ge> 0, (x,y) \<in> substs *)
   105 fun subst_direct_calls get_idx get_ctr_pos substs = 
   106   let
   107     fun subst (Abs (v, T, b)) = Abs (v, T, subst b)
   108       | subst t =
   109         let
   110           val (f, args) = strip_comb t;
   111           val idx = get_idx f;
   112           val ctr_pos  = if idx >= 0 then get_ctr_pos idx else ~1;
   113         in
   114           if idx < 0 then
   115             list_comb (f, map subst args)
   116           else if ctr_pos >= length args then
   117             primrec_error_eqn "too few arguments in recursive call" t
   118           else
   119             let
   120               val (key, repl) = the (find_first (equal (nth args ctr_pos) o fst) substs)
   121                 handle Option.Option => primrec_error_eqn
   122                   "recursive call not directly applied to constructor argument" t;
   123             in
   124               remove (op =) key args |> map subst |> curry list_comb repl
   125             end
   126         end
   127   in subst end;
   128 
   129 (* FIXME get rid of funs_data or get_indices *)
   130 fun rewrite_map_arg funs_data get_indices y rec_type res_type =
   131   let
   132     val pT = HOLogic.mk_prodT (rec_type, res_type);
   133     val fstx = fst_const pT;
   134     val sndx = snd_const pT;
   135 
   136     val SOME ({fun_name, left_args, ...} :: _) =
   137       find_first (equal rec_type o #rec_type o hd) funs_data;
   138     val ctr_pos = length left_args;
   139 
   140     fun subst _ d (t as Bound d') = t |> d = d' ? curry (op $) fstx
   141       | subst l d (Abs (v, T, b)) = Abs (v, if d < 0 then pT else T, subst l (d + 1) b)
   142       | subst l d t =
   143         let val (u, vs) = strip_comb t in
   144           if try (fst o dest_Free) u = SOME fun_name then
   145             if l andalso length vs = ctr_pos then
   146               list_comb (sndx |> permute_args ctr_pos, vs)
   147             else if length vs <= ctr_pos then
   148               primrec_error_eqn "too few arguments in recursive call" t
   149             else if nth vs ctr_pos |> member (op =) [y, Bound d] then
   150               list_comb (sndx $ nth vs ctr_pos, nth_drop ctr_pos vs |> map (subst false d))
   151             else
   152               primrec_error_eqn "recursive call not directly applied to constructor argument" t
   153           else if try (fst o dest_Const) u = SOME @{const_name comp} then
   154             (hd vs |> get_indices |> null orelse
   155               primrec_error_eqn "recursive call not directly applied to constructor argument" t;
   156             list_comb
   157               (u |> map_types (strip_type #>> (fn Ts => Ts
   158                    |> nth_map (length Ts - 1) (K pT)
   159                    |> nth_map (length Ts - 2) (strip_type #>> nth_map 0 (K pT) #> (op --->)))
   160                  #> (op --->)),
   161               nth_map 1 (subst l d) vs))
   162           else
   163             list_comb (u, map (subst false d) vs)
   164         end
   165   in
   166     subst true ~1
   167   end;
   168 
   169 (* FIXME get rid of funs_data or get_indices *)
   170 fun subst_indirect_call lthy funs_data get_indices (y, y') =
   171   let
   172     fun massage massage_map_arg bound_Ts =
   173       massage_indirect_rec_call lthy (not o null o get_indices) massage_map_arg bound_Ts y y';
   174     fun subst bound_Ts (t as _ $ _) =
   175         let
   176           val ctr_args = fold_aterms (curry (op @) o get_indices) t []
   177             |> maps (maps #ctr_args o nth funs_data);
   178           val (f', args') = strip_comb t;
   179           val fun_arg_idx = find_index (exists_subterm (not o null o get_indices)) args';
   180           val arg_idx = find_index (exists_subterm (equal y)) args';
   181           val (f, args) = chop (arg_idx + 1) args' |>> curry list_comb f';
   182           val _ = fun_arg_idx < 0 orelse arg_idx >= 0 orelse
   183             exists (exists_subterm (member (op =) ctr_args)) args' orelse
   184             primrec_error_eqn "recursive call not applied to constructor argument" t;
   185         in
   186           if fun_arg_idx <> arg_idx andalso fun_arg_idx >= 0 andalso arg_idx >= 0 then
   187             if nth args' arg_idx = y then
   188               list_comb (massage (rewrite_map_arg funs_data get_indices y) bound_Ts f, args)
   189             else
   190               primrec_error_eqn "recursive call not directly applied to constructor argument" f
   191           else
   192             list_comb (f', map (subst bound_Ts) args')
   193         end
   194       | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
   195       | subst bound_Ts t = t |> t = y ? massage (K I |> K) bound_Ts;
   196   in subst [] end;
   197 
   198 fun build_rec_arg lthy get_indices funs_data ctr_spec maybe_eqn_data =
   199   if is_some maybe_eqn_data then
   200     let
   201       val eqn_data = the maybe_eqn_data;
   202       val t = #rhs_term eqn_data;
   203       val ctr_args = #ctr_args eqn_data;
   204 
   205       val calls = #calls ctr_spec;
   206       val n_args = fold (curry (op +) o (fn Direct_Rec _ => 2 | _ => 1)) calls 0;
   207 
   208       val no_calls' = tag_list 0 calls
   209         |> map_filter (try (apsnd (fn No_Rec n => n | Direct_Rec (n, _) => n)));
   210       val direct_calls' = tag_list 0 calls
   211         |> map_filter (try (apsnd (fn Direct_Rec (_, n) => n)));
   212       val indirect_calls' = tag_list 0 calls
   213         |> map_filter (try (apsnd (fn Indirect_Rec n => n)));
   214 
   215       fun make_direct_type T = dummyT; (* FIXME? *)
   216 
   217       val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data;
   218 
   219       fun make_indirect_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T =>
   220         let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in
   221           if is_some maybe_res_type
   222           then HOLogic.mk_prodT (T, the maybe_res_type)
   223           else make_indirect_type T end))
   224         | make_indirect_type T = T;
   225 
   226       val args = replicate n_args ("", dummyT)
   227         |> Term.rename_wrt_term t
   228         |> map Free
   229         |> fold (fn (ctr_arg_idx, arg_idx) =>
   230             nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
   231           no_calls'
   232         |> fold (fn (ctr_arg_idx, arg_idx) =>
   233             nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_direct_type)))
   234           direct_calls'
   235         |> fold (fn (ctr_arg_idx, arg_idx) =>
   236             nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_indirect_type)))
   237           indirect_calls';
   238 
   239       val direct_calls = map (apfst (nth ctr_args) o apsnd (nth args)) direct_calls';
   240       val indirect_calls = map (apfst (nth ctr_args) o apsnd (nth args)) indirect_calls';
   241 
   242       val get_idx = (fn Free (v, _) => find_index (equal v o #fun_name o hd) funs_data | _ => ~1);
   243 
   244       val t' = t
   245         |> fold (subst_indirect_call lthy funs_data get_indices) indirect_calls
   246         |> subst_direct_calls get_idx (length o #left_args o hd o nth funs_data) direct_calls;
   247 
   248       val abstractions = map dest_Free (args @ #left_args eqn_data @ #right_args eqn_data);
   249     in
   250       t' |> fold_rev absfree abstractions
   251     end
   252   else Const (@{const_name undefined}, dummyT)
   253 
   254 fun build_defs lthy bs funs_data rec_specs get_indices =
   255   let
   256     val n_funs = length funs_data;
   257 
   258     val ctr_spec_eqn_data_list' =
   259       (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
   260       |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
   261           ##> (fn x => null x orelse
   262             primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
   263     val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
   264       primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
   265 
   266     val ctr_spec_eqn_data_list =
   267       ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
   268 
   269     val recs = take n_funs rec_specs |> map #recx;
   270     val rec_args = ctr_spec_eqn_data_list
   271       |> sort ((op <) o pairself (#offset o fst) |> make_ord)
   272       |> map (uncurry (build_rec_arg lthy get_indices funs_data) o apsnd (try the_single));
   273     val ctr_poss = map (fn x =>
   274       if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
   275         primrec_error ("inconstant constructor pattern position for function " ^
   276           quote (#fun_name (hd x)))
   277       else
   278         hd x |> #left_args |> length) funs_data;
   279   in
   280     (recs, ctr_poss)
   281     |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
   282     |> Syntax.check_terms lthy
   283     |> map2 (fn b => fn t => ((b, NoSyn), ((Binding.map_name Thm.def_name b, []), t))) bs
   284   end;
   285 
   286 fun find_rec_calls get_indices eqn_data =
   287   let
   288     fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg
   289       | find (t as _ $ _) ctr_arg =
   290         let
   291           val (f', args') = strip_comb t;
   292           val n = find_index (equal ctr_arg) args';
   293         in
   294           if n < 0 then
   295             find f' ctr_arg @ maps (fn x => find x ctr_arg) args'
   296           else
   297             let val (f, args) = chop n args' |>> curry list_comb f' in
   298               if exists_subterm (not o null o get_indices) f then
   299                 f :: maps (fn x => find x ctr_arg) args
   300               else
   301                 find f ctr_arg @ maps (fn x => find x ctr_arg) args
   302             end
   303         end
   304       | find _ _ = [];
   305   in
   306     map (find (#rhs_term eqn_data)) (#ctr_args eqn_data)
   307     |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
   308   end;
   309 
   310 fun add_primrec fixes specs lthy =
   311   let
   312     val bs = map (fst o fst) fixes;
   313     val fun_names = map Binding.name_of bs;
   314     val eqns_data = map (snd #> dissect_eqn lthy fun_names) specs;
   315     val funs_data = eqns_data
   316       |> partition_eq ((op =) o pairself #fun_name)
   317       |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
   318       |> map (fn (x, y) => the_single y handle List.Empty =>
   319           primrec_error ("missing equations for function " ^ quote x));
   320 
   321     fun get_indices t = map (fst #>> Binding.name_of #> Free) fixes
   322       |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   323       |> map_filter I;
   324 
   325     val arg_Ts = map (#rec_type o hd) funs_data;
   326     val res_Ts = map (#res_type o hd) funs_data;
   327     val callssss = funs_data
   328       |> map (partition_eq ((op =) o pairself #ctr))
   329       |> map (maps (map_filter (find_rec_calls get_indices)));
   330 
   331     val ((nontriv, rec_specs, _, induct_thm, induct_thms), lthy') =
   332       rec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
   333 
   334     val actual_nn = length funs_data;
   335 
   336     val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
   337       map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
   338         primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
   339           " is not a constructor in left-hand side") user_eqn) eqns_data end;
   340 
   341     val defs = build_defs lthy' bs funs_data rec_specs get_indices;
   342 
   343     fun prove def_thms' {ctr_specs, nested_map_idents, nested_map_comps, ...} induct_thm fun_data
   344         lthy =
   345       let
   346         val fun_name = #fun_name (hd fun_data);
   347         val def_thms = map (snd o snd) def_thms';
   348         val simp_thms = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
   349           |> fst
   350           |> map_filter (try (fn (x, [y]) =>
   351             (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
   352           |> map (fn (user_eqn, num_extra_args, rec_thm) =>
   353             mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
   354             |> K |> Goal.prove lthy [] [] user_eqn)
   355 
   356         val notes =
   357           [(inductN, if actual_nn > 1 then [induct_thm] else [], []),
   358            (simpsN, simp_thms, simp_attrs)]
   359           |> filter_out (null o #2)
   360           |> map (fn (thmN, thms, attrs) =>
   361             ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]));
   362       in
   363         lthy |> Local_Theory.notes notes
   364       end;
   365 
   366     val common_name = mk_common_name fun_names;
   367 
   368     val common_notes =
   369       [(inductN, if nontriv andalso actual_nn > 1 then [induct_thm] else [], [])]
   370       |> filter_out (null o #2)
   371       |> map (fn (thmN, thms, attrs) =>
   372         ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
   373   in
   374     lthy'
   375     |> fold_map Local_Theory.define defs
   376     |-> (fn def_thms' => fold_map3 (prove def_thms') (take actual_nn rec_specs)
   377       (take actual_nn induct_thms) funs_data)
   378     |> snd
   379     |> Local_Theory.notes common_notes |> snd
   380   end;
   381 
   382 fun add_primrec_cmd raw_fixes raw_specs lthy =
   383   let
   384     val _ = let val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes) in null d orelse
   385       primrec_error ("duplicate function name(s): " ^ commas d) end;
   386     val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy);
   387   in
   388     add_primrec fixes specs lthy
   389       handle ERROR str => primrec_error str
   390   end
   391   handle Primrec_Error (str, eqns) =>
   392     if null eqns
   393     then error ("primrec_new error:\n  " ^ str)
   394     else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
   395       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns))
   396 
   397 
   398 
   399 (* Primcorec *)
   400 
   401 type co_eqn_data_dtr_disc = {
   402   fun_name: string,
   403   ctr_no: int,
   404   cond: term,
   405   user_eqn: term
   406 };
   407 type co_eqn_data_dtr_sel = {
   408   fun_name: string,
   409   ctr_no: int,
   410   sel_no: int,
   411   fun_args: term list,
   412   rhs_term: term,
   413   user_eqn: term
   414 };
   415 datatype co_eqn_data =
   416   Dtr_Disc of co_eqn_data_dtr_disc |
   417   Dtr_Sel of co_eqn_data_dtr_sel
   418 
   419 fun co_dissect_eqn_disc sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps =
   420   let
   421     fun find_subterm p = let (* FIXME \<exists>? *)
   422       fun f (t as u $ v) =
   423         fold_rev (curry merge_options) [if p t then SOME t else NONE, f u, f v] NONE
   424         | f t = if p t then SOME t else NONE
   425       in f end;
   426 
   427     val fun_name = imp_rhs
   428       |> perhaps (try HOLogic.dest_not)
   429       |> `(try (fst o dest_Free o head_of o snd o dest_comb))
   430       ||> (try (fst o dest_Free o head_of o fst o HOLogic.dest_eq))
   431       |> the o merge_options;
   432     val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name)
   433       handle Option.Option => primrec_error_eqn "malformed discriminator equation" imp_rhs;
   434 
   435     val discs = #ctr_specs corec_spec |> map #disc;
   436     val ctrs = #ctr_specs corec_spec |> map #ctr;
   437     val n_ctrs = length ctrs;
   438     val not_disc = head_of imp_rhs = @{term Not};
   439     val _ = not_disc andalso n_ctrs <> 2 andalso
   440       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" imp_rhs;
   441     val disc = find_subterm (member (op =) discs o head_of) imp_rhs;
   442     val eq_ctr0 = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   443         |> (fn SOME t => let val n = find_index (equal t) ctrs in
   444           if n >= 0 then SOME n else NONE end | _ => NONE);
   445     val _ = is_some disc orelse is_some eq_ctr0 orelse
   446       primrec_error_eqn "no discriminator in equation" imp_rhs;
   447     val ctr_no' =
   448       if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
   449     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
   450     val fun_args = if is_none disc
   451       then imp_rhs |> perhaps (try HOLogic.dest_not) |> HOLogic.dest_eq |> fst |> strip_comb |> snd
   452       else the disc |> the_single o snd o strip_comb
   453         |> (fn t => if try (fst o dest_Free o head_of) t = SOME fun_name
   454           then snd (strip_comb t) else []);
   455 
   456     val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
   457     val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
   458     val catch_all = try (fst o dest_Free o the_single) imp_lhs' = SOME Name.uu_;
   459     val matched_conds = filter (equal fun_name o fst) matched_conds_ps |> map snd;
   460     val imp_lhs = mk_conjs imp_lhs';
   461     val cond =
   462       if catch_all then
   463         if null matched_conds then fold_rev absfree (map dest_Free fun_args) @{const True} else
   464           (strip_abs_vars (hd matched_conds),
   465             mk_disjs (map strip_abs_body matched_conds) |> HOLogic.mk_not)
   466           |-> fold_rev (fn (v, T) => fn u => Abs (v, T, u))
   467       else if sequential then
   468         HOLogic.mk_conj (HOLogic.mk_not (mk_disjs (map strip_abs_body matched_conds)), imp_lhs)
   469         |> fold_rev absfree (map dest_Free fun_args)
   470       else
   471         imp_lhs |> fold_rev absfree (map dest_Free fun_args);
   472     val matched_cond =
   473       if sequential then fold_rev absfree (map dest_Free fun_args) imp_lhs else cond;
   474 
   475     val matched_conds_ps' = if catch_all
   476       then (fun_name, cond) :: filter (not_equal fun_name o fst) matched_conds_ps
   477       else (fun_name, matched_cond) :: matched_conds_ps;
   478   in
   479     (Dtr_Disc {
   480       fun_name = fun_name,
   481       ctr_no = ctr_no,
   482       cond = cond,
   483       user_eqn = eqn'
   484     }, matched_conds_ps')
   485   end;
   486 
   487 fun co_dissect_eqn_sel fun_name_corec_spec_list eqn' eqn =
   488   let
   489     val (lhs, rhs) = HOLogic.dest_eq eqn
   490       handle TERM _ =>
   491         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   492     val sel = head_of lhs;
   493     val (fun_name, fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst (fst o dest_Free)
   494       handle TERM _ =>
   495         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   496     val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name)
   497       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   498     val ((ctr_spec, ctr_no), sel) = #ctr_specs corec_spec
   499       |> the o get_index (try (the o find_first (equal sel) o #sels))
   500       |>> `(nth (#ctr_specs corec_spec));
   501     val sel_no = find_index (equal sel) (#sels ctr_spec);
   502   in
   503     Dtr_Sel {
   504       fun_name = fun_name,
   505       ctr_no = ctr_no,
   506       sel_no = sel_no,
   507       fun_args = fun_args,
   508       rhs_term = rhs,
   509       user_eqn = eqn'
   510     }
   511   end;
   512 
   513 fun co_dissect_eqn_ctr sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps =
   514   let 
   515     val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
   516     val fun_name = head_of lhs |> fst o dest_Free;
   517     val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name);
   518     val (ctr, ctr_args) = strip_comb rhs;
   519     val ctr_spec = the (find_first (equal ctr o #ctr) (#ctr_specs corec_spec))
   520       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   521     val disc_imp_rhs = betapply (#disc ctr_spec, lhs);
   522     val (eqn_data_disc, matched_conds_ps') = co_dissect_eqn_disc
   523         sequential fun_name_corec_spec_list eqn' imp_lhs' disc_imp_rhs matched_conds_ps;
   524 
   525     val sel_imp_rhss = (#sels ctr_spec ~~ ctr_args)
   526       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   527 
   528 val _ = warning ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
   529  Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> " ^
   530  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
   531 
   532     val eqns_data_sel =
   533       map (co_dissect_eqn_sel fun_name_corec_spec_list @{const True}(*###*)) sel_imp_rhss;
   534   in
   535     (eqn_data_disc :: eqns_data_sel, matched_conds_ps')
   536   end;
   537 
   538 fun co_dissect_eqn sequential fun_name_corec_spec_list eqn' matched_conds_ps =
   539   let
   540     val eqn = subst_bounds (strip_qnt_vars @{const_name all} eqn' |> map Free |> rev,
   541         strip_qnt_body @{const_name all} eqn')
   542         handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   543     val (imp_lhs', imp_rhs) =
   544       (map HOLogic.dest_Trueprop (Logic.strip_imp_prems eqn),
   545        HOLogic.dest_Trueprop (Logic.strip_imp_concl eqn));
   546 
   547     val head = imp_rhs
   548       |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
   549       |> head_of;
   550 
   551     val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
   552 
   553     val fun_names = map fst fun_name_corec_spec_list;
   554     val discs = maps (#ctr_specs o snd) fun_name_corec_spec_list |> map #disc;
   555     val sels = maps (#ctr_specs o snd) fun_name_corec_spec_list |> maps #sels;
   556     val ctrs = maps (#ctr_specs o snd) fun_name_corec_spec_list |> map #ctr;
   557   in
   558     if member (op =) discs head orelse
   559       is_some maybe_rhs andalso
   560         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   561       co_dissect_eqn_disc sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps
   562       |>> single
   563     else if member (op =) sels head then
   564       ([co_dissect_eqn_sel fun_name_corec_spec_list eqn' imp_rhs], matched_conds_ps)
   565     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
   566       co_dissect_eqn_ctr sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps
   567     else
   568       primrec_error_eqn "malformed function equation" eqn
   569   end;
   570 
   571 fun build_corec_args_discs ctr_specs disc_eqns =
   572   let
   573     val conds = map #cond disc_eqns;
   574     val args =
   575       if length ctr_specs = 1 then []
   576       else if length disc_eqns = length ctr_specs then
   577         fst (split_last conds)
   578       else if length disc_eqns = length ctr_specs - 1 then
   579         let val n = 0 upto length ctr_specs - 1
   580             |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns)) (*###*) in
   581           if n = length ctr_specs - 1 then
   582             conds
   583           else
   584             split_last conds
   585             ||> (fn t => fold_rev absfree (strip_abs_vars t) (strip_abs_body t |> HOLogic.mk_not))
   586             |>> chop n
   587             |> (fn ((l, r), x) => l @ (x :: r))
   588         end
   589       else
   590         0 upto length ctr_specs - 1
   591         |> map (fn idx => find_first (equal idx o #ctr_no) disc_eqns
   592           |> Option.map #cond
   593           |> the_default (Const (@{const_name undefined}, dummyT)))
   594         |> fst o split_last;
   595     fun finish t =
   596       let val n = length (fastype_of t |> binder_types) in
   597         if t = Const (@{const_name undefined}, dummyT) then t
   598         else if n = 0 then Abs (Name.uu_, @{typ unit}, t)
   599         else if n = 1 then t
   600         else Const (@{const_name prod_case}, dummyT) $ t
   601       end;
   602   in
   603     map finish args
   604   end;
   605 
   606 fun build_corec_args_sel sel_eqns ctr_spec =
   607   let
   608     (* FIXME *)
   609     val n_args = fold (curry (op +)) (map (fn Direct_Corec _ => 3 | _ => 1) (#calls ctr_spec)) 0;
   610   in
   611     replicate n_args (Const (@{const_name undefined}, dummyT))
   612   end;
   613 
   614 fun co_build_defs lthy sequential bs arg_Tss fun_name_corec_spec_list eqns_data =
   615   let
   616     val fun_names = map Binding.name_of bs;
   617 
   618 (*    fun group _ [] = [] (* FIXME \<exists>? *)
   619       | group eq (x :: xs) =
   620         let val (xs', ys) = List.partition (eq x) xs in (x :: xs') :: group eq ys end;*)
   621     val disc_eqnss = map_filter (try (fn Dtr_Disc x => x)) eqns_data
   622       |> partition_eq ((op =) o pairself #fun_name)
   623       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
   624       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
   625 
   626     val _ = disc_eqnss |> map (fn x =>
   627       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
   628         primrec_error_eqns "excess discriminator equations in definition"
   629           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
   630 
   631 val _ = warning ("disc_eqnss:\n    \<cdot> " ^ space_implode "\n    \<cdot> " (map @{make_string} disc_eqnss));
   632 
   633     val sel_eqnss = map_filter (try (fn Dtr_Sel x => x)) eqns_data
   634       |> partition_eq ((op =) o pairself #fun_name)
   635       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
   636       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
   637 
   638 val _ = warning ("sel_eqnss:\n    \<cdot> " ^ space_implode "\n    \<cdot> " (map @{make_string} sel_eqnss));
   639 
   640     fun splice (xs :: xss) (ys :: yss) = xs @ ys @ splice xss yss (* FIXME \<exists>? *)
   641       | splice xss yss = flat xss @ flat yss;
   642     val corecs = map (#corec o snd) fun_name_corec_spec_list;
   643     val corec_args = (map snd fun_name_corec_spec_list ~~ disc_eqnss ~~ sel_eqnss)
   644       |> maps (fn (({ctr_specs, ...}, disc_eqns), sel_eqns) =>
   645         splice (build_corec_args_discs ctr_specs disc_eqns |> map single)
   646           (map (build_corec_args_sel sel_eqns) ctr_specs));
   647 
   648 val _ = warning ("corecursor arguments:\n    \<cdot> " ^
   649  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) corec_args));
   650 
   651     fun uneq_pairs_rev xs = xs (* FIXME \<exists>? *)
   652       |> these o try (split_last #> (fn (ys, y) => uneq_pairs_rev ys @ map (pair y) ys));
   653     val proof_obligations = if sequential then [] else
   654       maps (uneq_pairs_rev o map #cond) disc_eqnss
   655       |> map (fn (x, y) => ((strip_abs_body x, strip_abs_body y), strip_abs_vars x))
   656       |> map (apfst (apsnd HOLogic.mk_not #> pairself HOLogic.mk_Trueprop
   657         #> apfst (curry (op $) @{const ==>}) #> (op $)))
   658       |> map (fn (t, abs_vars) => fold_rev (fn (v, T) => fn u =>
   659           Const (@{const_name all}, (T --> @{typ prop}) --> @{typ prop}) $
   660             Abs (v, T, u)) abs_vars t);
   661 
   662     fun currys Ts t = if length Ts <= 1 then t else
   663       t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
   664         (length Ts - 1 downto 0 |> map Bound)
   665       |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts;
   666   in
   667     map (list_comb o rpair corec_args) corecs
   668     |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
   669     |> map2 currys arg_Tss
   670     |> Syntax.check_terms lthy
   671     |> map2 (fn b => fn t => ((b, NoSyn), ((Binding.map_name Thm.def_name b, []), t))) bs
   672     |> rpair proof_obligations
   673   end;
   674 
   675 fun add_primcorec sequential fixes specs lthy =
   676   let
   677     val bs = map (fst o fst) fixes;
   678     val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
   679 
   680     (* copied from primrec_new *)
   681     fun get_indices t = map (fst #>> Binding.name_of #> Free) fixes
   682       |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   683       |> map_filter I;
   684 
   685     val callssss = []; (* FIXME *)
   686 
   687     val ((nontriv, corec_specs, _, coinduct_thm, strong_co_induct_thm, coinduct_thmss,
   688           strong_coinduct_thmss), lthy') =
   689       corec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
   690 
   691     val fun_names = map Binding.name_of bs;
   692 
   693     val fun_name_corec_spec_list = (fun_names ~~ res_Ts, corec_specs)
   694       |> uncurry (finds (fn ((v, T), {corec, ...}) => T = body_type (fastype_of corec))) |> fst
   695       |> map (apfst fst #> apsnd the_single); (*###*)
   696 
   697     val (eqns_data, _) =
   698       fold_map (co_dissect_eqn sequential fun_name_corec_spec_list) (map snd specs) []
   699       |>> flat;
   700 
   701     val (defs, proof_obligations) =
   702       co_build_defs lthy' sequential bs (map (binder_types o snd o fst) fixes)
   703         fun_name_corec_spec_list eqns_data;
   704   in
   705     lthy'
   706     |> fold_map Local_Theory.define defs |> snd
   707     |> Proof.theorem NONE (K I) [map (rpair []) proof_obligations]
   708     |> Proof.refine (Method.primitive_text I)
   709     |> Seq.hd
   710   end
   711 
   712 fun add_primcorec_cmd seq (raw_fixes, raw_specs) lthy =
   713   let
   714     val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy);
   715   in
   716     add_primcorec seq fixes specs lthy
   717     handle ERROR str => primrec_error str
   718   end
   719   handle Primrec_Error (str, eqns) =>
   720     if null eqns
   721     then error ("primcorec error:\n  " ^ str)
   722     else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
   723       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns))
   724 
   725 end;