src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
author blanchet
Thu, 24 Oct 2013 19:43:21 +0200
changeset 55654 0a06b51ffa56
parent 55653 334a29265b2d
child 55656 5151b84d0668
permissions -rw-r--r--
handle applied ctor arguments gracefully when computing 'callssss' (for recursion through functions)
     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: (binding * typ option * mixfix) list ->
    11     (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
    12   val add_primrec_cmd: (binding * string option * mixfix) list ->
    13     (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
    14   val add_primrec_global: (binding * typ option * mixfix) list ->
    15     (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
    16   val add_primrec_overloaded: (string * (string * typ) * bool) list ->
    17     (binding * typ option * mixfix) list ->
    18     (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
    19   val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
    20     local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
    21   val add_primcorecursive_cmd: bool ->
    22     (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
    23     Proof.context -> Proof.state
    24   val add_primcorec_cmd: bool ->
    25     (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
    26     local_theory -> local_theory
    27 end;
    28 
    29 structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
    30 struct
    31 
    32 open BNF_Util
    33 open BNF_FP_Util
    34 open BNF_FP_Rec_Sugar_Util
    35 open BNF_FP_Rec_Sugar_Tactics
    36 
    37 val codeN = "code";
    38 val ctrN = "ctr";
    39 val discN = "disc";
    40 val selN = "sel";
    41 
    42 val nitpicksimp_attrs = @{attributes [nitpick_simp]};
    43 val simp_attrs = @{attributes [simp]};
    44 val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
    45 val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
    46 
    47 exception Primrec_Error of string * term list;
    48 
    49 fun primrec_error str = raise Primrec_Error (str, []);
    50 fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
    51 fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
    52 
    53 fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
    54 
    55 val free_name = try (fn Free (v, _) => v);
    56 val const_name = try (fn Const (v, _) => v);
    57 val undef_const = Const (@{const_name undefined}, dummyT);
    58 
    59 fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
    60   |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
    61 val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
    62 fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
    63   strip_qnt_body @{const_name all} t)
    64 fun abstract vs =
    65   let fun a n (t $ u) = a n t $ a n u
    66         | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
    67         | a n t = let val idx = find_index (equal t) vs in
    68             if idx < 0 then t else Bound (n + idx) end
    69   in a 0 end;
    70 fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u;
    71 fun mk_tuple1 Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 Ts));
    72 
    73 fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
    74   |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
    75   |> map_filter I;
    76 
    77 
    78 (* Primrec *)
    79 
    80 type eqn_data = {
    81   fun_name: string,
    82   rec_type: typ,
    83   ctr: term,
    84   ctr_args: term list,
    85   left_args: term list,
    86   right_args: term list,
    87   res_type: typ,
    88   rhs_term: term,
    89   user_eqn: term
    90 };
    91 
    92 fun dissect_eqn lthy fun_names eqn' =
    93   let
    94     val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
    95       handle TERM _ =>
    96         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    97     val (lhs, rhs) = HOLogic.dest_eq eqn
    98         handle TERM _ =>
    99           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
   100     val (fun_name, args) = strip_comb lhs
   101       |>> (fn x => if is_Free x then fst (dest_Free x)
   102           else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
   103     val (left_args, rest) = take_prefix is_Free args;
   104     val (nonfrees, right_args) = take_suffix is_Free rest;
   105     val num_nonfrees = length nonfrees;
   106     val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
   107       primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
   108       primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
   109     val _ = member (op =) fun_names fun_name orelse
   110       primrec_error_eqn "malformed function equation (does not start with function name)" eqn
   111 
   112     val (ctr, ctr_args) = strip_comb (the_single nonfrees);
   113     val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
   114       primrec_error_eqn "partially applied constructor in pattern" eqn;
   115     val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
   116       primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
   117         "\" in left-hand side") eqn end;
   118     val _ = forall is_Free ctr_args orelse
   119       primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
   120     val _ =
   121       let val b = fold_aterms (fn x as Free (v, _) =>
   122         if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
   123         not (member (op =) fun_names v) andalso
   124         not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
   125       in
   126         null b orelse
   127         primrec_error_eqn ("extra variable(s) in right-hand side: " ^
   128           commas (map (Syntax.string_of_term lthy) b)) eqn
   129       end;
   130   in
   131     {fun_name = fun_name,
   132      rec_type = body_type (type_of ctr),
   133      ctr = ctr,
   134      ctr_args = ctr_args,
   135      left_args = left_args,
   136      right_args = right_args,
   137      res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
   138      rhs_term = rhs,
   139      user_eqn = eqn'}
   140   end;
   141 
   142 fun rewrite_map_arg get_ctr_pos rec_type res_type =
   143   let
   144     val pT = HOLogic.mk_prodT (rec_type, res_type);
   145 
   146     val maybe_suc = Option.map (fn x => x + 1);
   147     fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
   148       | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
   149       | subst d t =
   150         let
   151           val (u, vs) = strip_comb t;
   152           val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1;
   153         in
   154           if ctr_pos >= 0 then
   155             if d = SOME ~1 andalso length vs = ctr_pos then
   156               list_comb (permute_args ctr_pos (snd_const pT), vs)
   157             else if length vs > ctr_pos andalso is_some d
   158                 andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
   159               list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
   160             else
   161               primrec_error_eqn ("recursive call not directly applied to constructor argument") t
   162           else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then
   163             list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs)
   164           else
   165             list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
   166         end
   167   in
   168     subst (SOME ~1)
   169   end;
   170 
   171 fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
   172   let
   173     fun try_nested_rec bound_Ts y t =
   174       AList.lookup (op =) nested_calls y
   175       |> Option.map (fn y' =>
   176         massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
   177 
   178     fun subst bound_Ts (t as g' $ y) =
   179         let
   180           fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
   181           val y_head = head_of y;
   182         in
   183           if not (member (op =) ctr_args y_head) then
   184             subst_rec ()
   185           else
   186             (case try_nested_rec bound_Ts y_head t of
   187               SOME t' => t'
   188             | NONE =>
   189               let val (g, g_args) = strip_comb g' in
   190                 (case try (get_ctr_pos o the) (free_name g) of
   191                   SOME ctr_pos =>
   192                   (length g_args >= ctr_pos orelse
   193                    primrec_error_eqn "too few arguments in recursive call" t;
   194                    (case AList.lookup (op =) mutual_calls y of
   195                      SOME y' => list_comb (y', g_args)
   196                    | NONE => subst_rec ()))
   197                 | NONE => subst_rec ())
   198               end)
   199         end
   200       | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
   201       | subst _ t = t
   202 
   203     fun subst' t =
   204       if has_call t then
   205         (* FIXME detect this case earlier? *)
   206         primrec_error_eqn "recursive call not directly applied to constructor argument" t
   207       else
   208         try_nested_rec [] (head_of t) t |> the_default t
   209   in
   210     subst' o subst []
   211   end;
   212 
   213 fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
   214     (maybe_eqn_data : eqn_data option) =
   215   if is_none maybe_eqn_data then undef_const else
   216     let
   217       val eqn_data = the maybe_eqn_data;
   218       val t = #rhs_term eqn_data;
   219       val ctr_args = #ctr_args eqn_data;
   220 
   221       val calls = #calls ctr_spec;
   222       val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
   223 
   224       val no_calls' = tag_list 0 calls
   225         |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
   226       val mutual_calls' = tag_list 0 calls
   227         |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
   228       val nested_calls' = tag_list 0 calls
   229         |> map_filter (try (apsnd (fn Nested_Rec p => p)));
   230 
   231       val args = replicate n_args ("", dummyT)
   232         |> Term.rename_wrt_term t
   233         |> map Free
   234         |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
   235             nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
   236           no_calls'
   237         |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
   238             nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
   239           mutual_calls'
   240         |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
   241             nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
   242           nested_calls';
   243 
   244       val fun_name_ctr_pos_list =
   245         map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
   246       val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
   247       val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
   248       val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
   249 
   250       val abstractions = args @ #left_args eqn_data @ #right_args eqn_data;
   251     in
   252       t
   253       |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
   254       |> fold_rev lambda abstractions
   255     end;
   256 
   257 fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
   258   let
   259     val n_funs = length funs_data;
   260 
   261     val ctr_spec_eqn_data_list' =
   262       (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
   263       |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
   264           ##> (fn x => null x orelse
   265             primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
   266     val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
   267       primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
   268 
   269     val ctr_spec_eqn_data_list =
   270       ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
   271 
   272     val recs = take n_funs rec_specs |> map #recx;
   273     val rec_args = ctr_spec_eqn_data_list
   274       |> sort ((op <) o pairself (#offset o fst) |> make_ord)
   275       |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
   276     val ctr_poss = map (fn x =>
   277       if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
   278         primrec_error ("inconstant constructor pattern position for function " ^
   279           quote (#fun_name (hd x)))
   280       else
   281         hd x |> #left_args |> length) funs_data;
   282   in
   283     (recs, ctr_poss)
   284     |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
   285     |> Syntax.check_terms lthy
   286     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
   287       bs mxs
   288   end;
   289 
   290 fun massage_comp ctxt has_call bound_Ts t =
   291   massage_nested_corec_call ctxt has_call (K (K (K I))) bound_Ts (fastype_of1 (bound_Ts, t)) t;
   292 
   293 fun find_rec_calls ctxt has_call (eqn_data : eqn_data) =
   294   let
   295     fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
   296       | find bound_Ts (t as _ $ _) ctr_arg =
   297         let
   298           val typof = curry fastype_of1 bound_Ts;
   299           val (f', args') = strip_comb t;
   300           val n = find_index (equal ctr_arg o head_of) args';
   301         in
   302           if n < 0 then
   303             find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
   304           else
   305             let
   306               val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
   307               val (arg_head, arg_args) = Term.strip_comb arg;
   308             in
   309               if has_call f then
   310                 mk_partial_compN (length arg_args) (typof f) (typof arg_head) f ::
   311                 maps (fn x => find bound_Ts x ctr_arg) args
   312               else
   313                 find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
   314             end
   315         end
   316       | find _ _ _ = [];
   317   in
   318     map (find [] (#rhs_term eqn_data)) (#ctr_args eqn_data)
   319     |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
   320   end;
   321 
   322 fun prepare_primrec fixes specs lthy =
   323   let
   324     val (bs, mxs) = map_split (apfst fst) fixes;
   325     val fun_names = map Binding.name_of bs;
   326     val eqns_data = map (dissect_eqn lthy fun_names) specs;
   327     val funs_data = eqns_data
   328       |> partition_eq ((op =) o pairself #fun_name)
   329       |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
   330       |> map (fn (x, y) => the_single y handle List.Empty =>
   331           primrec_error ("missing equations for function " ^ quote x));
   332 
   333     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   334     val arg_Ts = map (#rec_type o hd) funs_data;
   335     val res_Ts = map (#res_type o hd) funs_data;
   336     val callssss = funs_data
   337       |> map (partition_eq ((op =) o pairself #ctr))
   338       |> map (maps (map_filter (find_rec_calls lthy has_call)));
   339 
   340     val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
   341       rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
   342 
   343     val actual_nn = length funs_data;
   344 
   345     val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
   346       map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
   347         primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
   348           " is not a constructor in left-hand side") user_eqn) eqns_data end;
   349 
   350     val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
   351 
   352     fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
   353         (fun_data : eqn_data list) =
   354       let
   355         val def_thms = map (snd o snd) def_thms';
   356         val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
   357           |> fst
   358           |> map_filter (try (fn (x, [y]) =>
   359             (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
   360           |> map (fn (user_eqn, num_extra_args, rec_thm) =>
   361             mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
   362             |> K |> Goal.prove lthy [] [] user_eqn
   363             |> Thm.close_derivation);
   364         val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
   365       in
   366         (poss, simp_thmss)
   367       end;
   368 
   369     val notes =
   370       (if n2m then map2 (fn name => fn thm =>
   371         (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
   372       |> map (fn (prefix, thmN, thms, attrs) =>
   373         ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
   374 
   375     val common_name = mk_common_name fun_names;
   376 
   377     val common_notes =
   378       (if n2m then [(inductN, [induct_thm], [])] else [])
   379       |> map (fn (thmN, thms, attrs) =>
   380         ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
   381   in
   382     (((fun_names, defs),
   383       fn lthy => fn defs =>
   384         split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
   385       lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
   386   end;
   387 
   388 (* primrec definition *)
   389 
   390 fun add_primrec_simple fixes ts lthy =
   391   let
   392     val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
   393       handle ERROR str => primrec_error str;
   394   in
   395     lthy
   396     |> fold_map Local_Theory.define defs
   397     |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
   398   end
   399   handle Primrec_Error (str, eqns) =>
   400     if null eqns
   401     then error ("primrec_new error:\n  " ^ str)
   402     else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
   403       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
   404 
   405 local
   406 
   407 fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
   408   let
   409     val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
   410     val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
   411 
   412     val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
   413 
   414     val mk_notes =
   415       flat ooo map3 (fn poss => fn prefix => fn thms =>
   416         let
   417           val (bs, attrss) = map_split (fst o nth specs) poss;
   418           val notes =
   419             map3 (fn b => fn attrs => fn thm =>
   420               ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
   421             bs attrss thms;
   422         in
   423           ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
   424         end);
   425   in
   426     lthy
   427     |> add_primrec_simple fixes (map snd specs)
   428     |-> (fn (names, (ts, (posss, simpss))) =>
   429       Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
   430       #> Local_Theory.notes (mk_notes posss names simpss)
   431       #>> pair ts o map snd)
   432   end;
   433 
   434 in
   435 
   436 val add_primrec = gen_primrec Specification.check_spec;
   437 val add_primrec_cmd = gen_primrec Specification.read_spec;
   438 
   439 end;
   440 
   441 fun add_primrec_global fixes specs thy =
   442   let
   443     val lthy = Named_Target.theory_init thy;
   444     val ((ts, simps), lthy') = add_primrec fixes specs lthy;
   445     val simps' = burrow (Proof_Context.export lthy' lthy) simps;
   446   in ((ts, simps'), Local_Theory.exit_global lthy') end;
   447 
   448 fun add_primrec_overloaded ops fixes specs thy =
   449   let
   450     val lthy = Overloading.overloading ops thy;
   451     val ((ts, simps), lthy') = add_primrec fixes specs lthy;
   452     val simps' = burrow (Proof_Context.export lthy' lthy) simps;
   453   in ((ts, simps'), Local_Theory.exit_global lthy') end;
   454 
   455 
   456 
   457 (* Primcorec *)
   458 
   459 type coeqn_data_disc = {
   460   fun_name: string,
   461   fun_T: typ,
   462   fun_args: term list,
   463   ctr: term,
   464   ctr_no: int, (*###*)
   465   disc: term,
   466   prems: term list,
   467   auto_gen: bool,
   468   maybe_ctr_rhs: term option,
   469   maybe_code_rhs: term option,
   470   user_eqn: term
   471 };
   472 
   473 type coeqn_data_sel = {
   474   fun_name: string,
   475   fun_T: typ,
   476   fun_args: term list,
   477   ctr: term,
   478   sel: term,
   479   rhs_term: term,
   480   user_eqn: term
   481 };
   482 
   483 datatype coeqn_data =
   484   Disc of coeqn_data_disc |
   485   Sel of coeqn_data_sel;
   486 
   487 fun dissect_coeqn_disc seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
   488     maybe_ctr_rhs maybe_code_rhs prems' concl matchedsss =
   489   let
   490     fun find_subterm p = let (* FIXME \<exists>? *)
   491       fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
   492         | f t = if p t then SOME t else NONE
   493       in f end;
   494 
   495     val applied_fun = concl
   496       |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
   497       |> the
   498       handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
   499     val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
   500     val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
   501 
   502     val discs = map #disc basic_ctr_specs;
   503     val ctrs = map #ctr basic_ctr_specs;
   504     val not_disc = head_of concl = @{term Not};
   505     val _ = not_disc andalso length ctrs <> 2 andalso
   506       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
   507     val disc' = find_subterm (member (op =) discs o head_of) concl;
   508     val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   509         |> (fn SOME t => let val n = find_index (equal t) ctrs in
   510           if n >= 0 then SOME n else NONE end | _ => NONE);
   511     val _ = is_some disc' orelse is_some eq_ctr0 orelse
   512       primrec_error_eqn "no discriminator in equation" concl;
   513     val ctr_no' =
   514       if is_none disc' then the eq_ctr0 else find_index (equal (head_of (the disc'))) discs;
   515     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
   516     val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
   517 
   518     val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
   519     val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
   520     val prems = map (abstract (List.rev fun_args)) prems';
   521     val real_prems =
   522       (if catch_all orelse seq then maps s_not_conj matchedss else []) @
   523       (if catch_all then [] else prems);
   524 
   525     val matchedsss' = AList.delete (op =) fun_name matchedsss
   526       |> cons (fun_name, if seq then matchedss @ [prems] else matchedss @ [real_prems]);
   527 
   528     val user_eqn =
   529       (real_prems, concl)
   530       |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
   531       |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
   532   in
   533     (Disc {
   534       fun_name = fun_name,
   535       fun_T = fun_T,
   536       fun_args = fun_args,
   537       ctr = ctr,
   538       ctr_no = ctr_no,
   539       disc = disc,
   540       prems = real_prems,
   541       auto_gen = catch_all,
   542       maybe_ctr_rhs = maybe_ctr_rhs,
   543       maybe_code_rhs = maybe_code_rhs,
   544       user_eqn = user_eqn
   545     }, matchedsss')
   546   end;
   547 
   548 fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn' of_spec
   549     eqn =
   550   let
   551     val (lhs, rhs) = HOLogic.dest_eq eqn
   552       handle TERM _ =>
   553         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   554     val sel = head_of lhs;
   555     val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
   556       handle TERM _ =>
   557         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   558     val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
   559       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   560     val {ctr, ...} =
   561       if is_some of_spec
   562       then the (find_first (equal (the of_spec) o #ctr) basic_ctr_specs)
   563       else filter (exists (equal sel) o #sels) basic_ctr_specs |> the_single
   564         handle List.Empty => primrec_error_eqn "ambiguous selector - use \"of\"" eqn;
   565     val user_eqn = drop_All eqn';
   566   in
   567     Sel {
   568       fun_name = fun_name,
   569       fun_T = fun_T,
   570       fun_args = fun_args,
   571       ctr = ctr,
   572       sel = sel,
   573       rhs_term = rhs,
   574       user_eqn = user_eqn
   575     }
   576   end;
   577 
   578 fun dissect_coeqn_ctr seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
   579     maybe_code_rhs prems concl matchedsss =
   580   let
   581     val (lhs, rhs) = HOLogic.dest_eq concl;
   582     val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
   583     val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
   584     val (ctr, ctr_args) = strip_comb (unfold_let rhs);
   585     val {disc, sels, ...} = the (find_first (equal ctr o #ctr) basic_ctr_specs)
   586       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   587 
   588     val disc_concl = betapply (disc, lhs);
   589     val (maybe_eqn_data_disc, matchedsss') = if length basic_ctr_specs = 1
   590       then (NONE, matchedsss)
   591       else apfst SOME (dissect_coeqn_disc seq fun_names basic_ctr_specss
   592           (SOME (abstract (List.rev fun_args) rhs)) maybe_code_rhs prems disc_concl matchedsss);
   593 
   594     val sel_concls = sels ~~ ctr_args
   595       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   596 
   597 (*
   598 val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
   599  (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
   600  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
   601  "\nfor premise(s)\n    \<cdot> " ^
   602  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
   603 *)
   604 
   605     val eqns_data_sel =
   606       map (dissect_coeqn_sel fun_names basic_ctr_specss eqn' (SOME ctr)) sel_concls;
   607   in
   608     (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
   609   end;
   610 
   611 fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss =
   612   let
   613     val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
   614     val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
   615     val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
   616 
   617     val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
   618         if member ((op =) o apsnd #ctr) basic_ctr_specs ctr
   619         then cons (ctr, cs)
   620         else primrec_error_eqn "not a constructor" ctr) [] rhs' []
   621       |> AList.group (op =);
   622 
   623     val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
   624     val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
   625         binder_types (fastype_of ctr)
   626         |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
   627           if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
   628         |> curry list_comb ctr
   629         |> curry HOLogic.mk_eq lhs);
   630   in
   631     fold_map2 (dissect_coeqn_ctr false fun_names basic_ctr_specss eqn'
   632         (SOME (abstract (List.rev fun_args) rhs)))
   633       ctr_premss ctr_concls matchedsss
   634   end;
   635 
   636 fun dissect_coeqn lthy seq has_call fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
   637     eqn' of_spec matchedsss =
   638   let
   639     val eqn = drop_All eqn'
   640       handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   641     val (prems, concl) = Logic.strip_horn eqn
   642       |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
   643 
   644     val head = concl
   645       |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
   646       |> head_of;
   647 
   648     val maybe_rhs = concl |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
   649 
   650     val discs = maps (map #disc) basic_ctr_specss;
   651     val sels = maps (maps #sels) basic_ctr_specss;
   652     val ctrs = maps (map #ctr) basic_ctr_specss;
   653   in
   654     if member (op =) discs head orelse
   655       is_some maybe_rhs andalso
   656         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   657       dissect_coeqn_disc seq fun_names basic_ctr_specss NONE NONE prems concl matchedsss
   658       |>> single
   659     else if member (op =) sels head then
   660       ([dissect_coeqn_sel fun_names basic_ctr_specss eqn' of_spec concl], matchedsss)
   661     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
   662       member (op =) ctrs (head_of (unfold_let (the maybe_rhs))) then
   663       dissect_coeqn_ctr seq fun_names basic_ctr_specss eqn' NONE prems concl matchedsss
   664     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
   665       null prems then
   666       dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss
   667       |>> flat
   668     else
   669       primrec_error_eqn "malformed function equation" eqn
   670   end;
   671 
   672 fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
   673     ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
   674   if is_none (#pred (nth ctr_specs ctr_no)) then I else
   675     s_conjs prems
   676     |> curry subst_bounds (List.rev fun_args)
   677     |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
   678     |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
   679 
   680 fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
   681   find_first (equal sel o #sel) sel_eqns
   682   |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
   683   |> the_default undef_const
   684   |> K;
   685 
   686 fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
   687   let
   688     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   689   in
   690     if is_none maybe_sel_eqn then (I, I, I) else
   691     let
   692       val {fun_args, rhs_term, ... } = the maybe_sel_eqn;
   693       val bound_Ts = List.rev (map fastype_of fun_args);
   694       fun rewrite_q _ t = if has_call t then @{term False} else @{term True};
   695       fun rewrite_g _ t = if has_call t then undef_const else t;
   696       fun rewrite_h bound_Ts t =
   697         if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
   698       fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
   699         |> abs_tuple fun_args;
   700     in
   701       (massage rewrite_q,
   702        massage rewrite_g,
   703        massage rewrite_h)
   704     end
   705   end;
   706 
   707 fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
   708   let
   709     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   710   in
   711     if is_none maybe_sel_eqn then I else
   712     let
   713       val {fun_args, rhs_term, ...} = the maybe_sel_eqn;
   714       val bound_Ts = List.rev (map fastype_of fun_args);
   715       fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
   716         | rewrite bound_Ts U T (t as _ $ _) =
   717           let val (u, vs) = strip_comb t in
   718             if is_Free u andalso has_call u then
   719               Inr_const U T $ mk_tuple1 bound_Ts vs
   720             else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
   721               map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
   722             else
   723               list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
   724           end
   725         | rewrite _ U T t =
   726           if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
   727       fun massage t =
   728         rhs_term
   729         |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
   730         |> abs_tuple fun_args;
   731     in
   732       massage
   733     end
   734   end;
   735 
   736 fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
   737     (ctr_spec : corec_ctr_spec) =
   738   let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
   739     if null sel_eqns then I else
   740       let
   741         val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
   742 
   743         val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
   744         val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
   745         val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
   746       in
   747         I
   748         #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
   749         #> fold (fn (sel, (q, g, h)) =>
   750           let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
   751             nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
   752         #> fold (fn (sel, n) => nth_map n
   753           (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
   754       end
   755   end;
   756 
   757 fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
   758     (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
   759   let
   760     val corecs = map #corec corec_specs;
   761     val ctr_specss = map #ctr_specs corec_specs;
   762     val corec_args = hd corecs
   763       |> fst o split_last o binder_types o fastype_of
   764       |> map (Const o pair @{const_name undefined})
   765       |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
   766       |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
   767     fun currys [] t = t
   768       | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
   769           |> fold_rev (Term.abs o pair Name.uu) Ts;
   770 
   771 (*
   772 val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
   773  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
   774 *)
   775 
   776     val exclss' =
   777       disc_eqnss
   778       |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
   779         #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
   780         #> maps (uncurry (map o pair)
   781           #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
   782               ((c, c', a orelse a'), (x, s_not (s_conjs y)))
   783             ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
   784             ||> Logic.list_implies
   785             ||> curry Logic.list_all (map dest_Free fun_args))))
   786   in
   787     map (list_comb o rpair corec_args) corecs
   788     |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
   789     |> map2 currys arg_Tss
   790     |> Syntax.check_terms lthy
   791     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
   792       bs mxs
   793     |> rpair exclss'
   794   end;
   795 
   796 fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
   797     (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
   798   if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
   799     let
   800       val n = 0 upto length ctr_specs
   801         |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
   802       val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
   803         |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
   804       val extra_disc_eqn = {
   805         fun_name = Binding.name_of fun_binding,
   806         fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
   807         fun_args = fun_args,
   808         ctr = #ctr (nth ctr_specs n),
   809         ctr_no = n,
   810         disc = #disc (nth ctr_specs n),
   811         prems = maps (s_not_conj o #prems) disc_eqns,
   812         auto_gen = true,
   813         maybe_ctr_rhs = NONE,
   814         maybe_code_rhs = NONE,
   815         user_eqn = undef_const};
   816     in
   817       chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
   818     end;
   819 
   820 fun find_corec_calls has_call basic_ctr_specs {ctr, sel, rhs_term, ...} =
   821   let
   822     val sel_no = find_first (equal ctr o #ctr) basic_ctr_specs
   823       |> find_index (equal sel) o #sels o the;
   824     fun find (Abs (_, _, b)) = find b
   825       | find (t as _ $ _) = strip_comb t |>> find ||> maps find |> (op @)
   826       | find f = if is_Free f andalso has_call f then [f] else [];
   827   in
   828     find rhs_term
   829     |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
   830   end;
   831 
   832 fun add_primcorec_ursive maybe_tac seq fixes specs of_specs lthy =
   833   let
   834     val (bs, mxs) = map_split (apfst fst) fixes;
   835     val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
   836 
   837     val fun_names = map Binding.name_of bs;
   838     val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
   839     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   840     val eqns_data =
   841       fold_map2 (dissect_coeqn lthy seq has_call fun_names basic_ctr_specss) (map snd specs)
   842         of_specs []
   843       |> flat o fst;
   844 
   845     val callssss =
   846       map_filter (try (fn Sel x => x)) eqns_data
   847       |> partition_eq ((op =) o pairself #fun_name)
   848       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   849       |> map (flat o snd)
   850       |> map2 (fold o find_corec_calls has_call) basic_ctr_specss
   851       |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
   852         (ctr, map (K []) sels))) basic_ctr_specss);
   853 
   854 (*
   855 val _ = tracing ("callssss = " ^ @{make_string} callssss);
   856 *)
   857 
   858     val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
   859           strong_coinduct_thms), lthy') =
   860       corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
   861     val actual_nn = length bs;
   862     val corec_specs = take actual_nn corec_specs'; (*###*)
   863     val ctr_specss = map #ctr_specs corec_specs;
   864 
   865     val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
   866       |> partition_eq ((op =) o pairself #fun_name)
   867       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   868       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
   869     val _ = disc_eqnss' |> map (fn x =>
   870       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
   871         primrec_error_eqns "excess discriminator equations in definition"
   872           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
   873 
   874     val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
   875       |> partition_eq ((op =) o pairself #fun_name)
   876       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   877       |> map (flat o snd);
   878 
   879     val arg_Tss = map (binder_types o snd o fst) fixes;
   880     val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
   881     val (defs, exclss') =
   882       build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
   883 
   884     fun excl_tac (c, c', a) =
   885       if a orelse c = c' orelse seq then SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
   886       else maybe_tac;
   887 
   888 (*
   889 val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
   890  space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
   891 *)
   892 
   893     val exclss'' = exclss' |> map (map (fn (idx, t) =>
   894       (idx, (Option.map (Goal.prove lthy [] [] t #> Thm.close_derivation) (excl_tac idx), t))));
   895     val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
   896     val (goal_idxss, goalss) = exclss''
   897       |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
   898       |> split_list o map split_list;
   899 
   900     fun prove thmss' def_thms' lthy =
   901       let
   902         val def_thms = map (snd o snd) def_thms';
   903 
   904         val exclss' = map (op ~~) (goal_idxss ~~ thmss');
   905         fun mk_exclsss excls n =
   906           (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
   907           |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
   908         val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
   909           |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
   910 
   911         fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
   912             ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
   913           if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else
   914             let
   915               val {disc_corec, ...} = nth ctr_specs ctr_no;
   916               val k = 1 + ctr_no;
   917               val m = length prems;
   918               val t =
   919                 list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
   920                 |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
   921                 |> HOLogic.mk_Trueprop
   922                 |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   923                 |> curry Logic.list_all (map dest_Free fun_args);
   924             in
   925               if prems = [@{term False}] then [] else
   926               mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
   927               |> K |> Goal.prove lthy [] [] t
   928               |> Thm.close_derivation
   929               |> pair (#disc (nth ctr_specs ctr_no))
   930               |> single
   931             end;
   932 
   933         fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
   934             : corec_spec) (disc_eqns : coeqn_data_disc list) exclsss
   935             ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : coeqn_data_sel) =
   936           let
   937             val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
   938             val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
   939             val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
   940                 (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
   941             val sel_corec = find_index (equal sel) (#sels ctr_spec)
   942               |> nth (#sel_corecs ctr_spec);
   943             val k = 1 + ctr_no;
   944             val m = length prems;
   945             val t =
   946               list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
   947               |> curry betapply sel
   948               |> rpair (abstract (List.rev fun_args) rhs_term)
   949               |> HOLogic.mk_Trueprop o HOLogic.mk_eq
   950               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   951               |> curry Logic.list_all (map dest_Free fun_args);
   952             val (distincts, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
   953           in
   954             mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
   955               nested_map_idents nested_map_comps sel_corec k m exclsss
   956             |> K |> Goal.prove lthy [] [] t
   957             |> Thm.close_derivation
   958             |> pair sel
   959           end;
   960 
   961         fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
   962             (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
   963           (* don't try to prove theorems when some sel_eqns are missing *)
   964           if not (exists (equal ctr o #ctr) disc_eqns)
   965               andalso not (exists (equal ctr o #ctr) sel_eqns)
   966             orelse
   967               filter (equal ctr o #ctr) sel_eqns
   968               |> fst o finds ((op =) o apsnd #sel) sels
   969               |> exists (null o snd)
   970           then [] else
   971             let
   972               val (fun_name, fun_T, fun_args, prems, maybe_rhs) =
   973                 (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
   974                 |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
   975                   #maybe_ctr_rhs x))
   976                 ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], NONE))
   977                 |> the o merge_options;
   978               val m = length prems;
   979               val t = (if is_some maybe_rhs then the maybe_rhs else
   980                   filter (equal ctr o #ctr) sel_eqns
   981                   |> fst o finds ((op =) o apsnd #sel) sels
   982                   |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
   983                   |> curry list_comb ctr)
   984                 |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
   985                   map Bound (length fun_args - 1 downto 0)))
   986                 |> HOLogic.mk_Trueprop
   987                 |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   988                 |> curry Logic.list_all (map dest_Free fun_args);
   989               val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
   990               val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
   991             in
   992               if prems = [@{term False}] then [] else
   993                 mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
   994                 |> K |> Goal.prove lthy [] [] t
   995                 |> Thm.close_derivation
   996                 |> pair ctr
   997                 |> single
   998             end;
   999 
  1000         fun prove_code disc_eqns sel_eqns ctr_alist ctr_specs =
  1001           let
  1002             val (fun_name, fun_T, fun_args, maybe_rhs) =
  1003               (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
  1004                find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
  1005               |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #maybe_code_rhs x))
  1006               ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, NONE))
  1007               |> the o merge_options;
  1008 
  1009             val bound_Ts = List.rev (map fastype_of fun_args);
  1010 
  1011             val lhs = list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
  1012             val maybe_rhs_info =
  1013               (case maybe_rhs of
  1014                 SOME rhs =>
  1015                 let
  1016                   val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
  1017                   val cond_ctrs =
  1018                     fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
  1019                   val ctr_thms = map (the o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
  1020                 in SOME (rhs, raw_rhs, ctr_thms) end
  1021               | NONE =>
  1022                 let
  1023                   fun prove_code_ctr {ctr, sels, ...} =
  1024                     if not (exists (equal ctr o fst) ctr_alist) then NONE else
  1025                       let
  1026                         val prems = find_first (equal ctr o #ctr) disc_eqns
  1027                           |> Option.map #prems |> the_default [];
  1028                         val t =
  1029                           filter (equal ctr o #ctr) sel_eqns
  1030                           |> fst o finds ((op =) o apsnd #sel) sels
  1031                           |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
  1032                             #-> abstract)
  1033                           |> curry list_comb ctr;
  1034                       in
  1035                         SOME (prems, t)
  1036                       end;
  1037                   val maybe_ctr_conds_argss = map prove_code_ctr ctr_specs;
  1038                 in
  1039                   if exists is_none maybe_ctr_conds_argss then NONE else
  1040                     let
  1041                       val rhs = fold_rev (fn SOME (prems, u) => fn t => mk_If (s_conjs prems) u t)
  1042                         maybe_ctr_conds_argss
  1043                         (Const (@{const_name Code.abort}, @{typ String.literal} -->
  1044                             (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
  1045                           HOLogic.mk_literal fun_name $
  1046                           absdummy @{typ unit} (incr_boundvars 1 lhs));
  1047                     in SOME (rhs, rhs, map snd ctr_alist) end
  1048                 end);
  1049           in
  1050             (case maybe_rhs_info of
  1051               NONE => []
  1052             | SOME (rhs, raw_rhs, ctr_thms) =>
  1053               let
  1054                 val ms = map (Logic.count_prems o prop_of) ctr_thms;
  1055                 val (raw_t, t) = (raw_rhs, rhs)
  1056                   |> pairself
  1057                     (curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
  1058                       map Bound (length fun_args - 1 downto 0)))
  1059                     #> HOLogic.mk_Trueprop
  1060                     #> curry Logic.list_all (map dest_Free fun_args));
  1061                 val (distincts, discIs, sel_splits, sel_split_asms) =
  1062                   case_thms_of_term lthy bound_Ts raw_rhs;
  1063 
  1064                 val raw_code_thm = mk_primcorec_raw_code_of_ctr_tac lthy distincts discIs sel_splits
  1065                     sel_split_asms ms ctr_thms
  1066                   |> K |> Goal.prove lthy [] [] raw_t
  1067                   |> Thm.close_derivation;
  1068               in
  1069                 mk_primcorec_code_of_raw_code_tac lthy distincts sel_splits raw_code_thm
  1070                 |> K |> Goal.prove lthy [] [] t
  1071                 |> Thm.close_derivation
  1072                 |> single
  1073               end)
  1074           end;
  1075 
  1076         val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
  1077         val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
  1078         val disc_thmss = map (map snd) disc_alists;
  1079         val sel_thmss = map (map snd) sel_alists;
  1080 
  1081         val ctr_alists = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
  1082           ctr_specss;
  1083         val ctr_thmss = map (map snd) ctr_alists;
  1084 
  1085         val code_thmss = map4 prove_code disc_eqnss sel_eqnss ctr_alists ctr_specss;
  1086 
  1087         val simp_thmss = map2 append disc_thmss sel_thmss
  1088 
  1089         val common_name = mk_common_name fun_names;
  1090 
  1091         val notes =
  1092           [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
  1093            (codeN, code_thmss, code_nitpicksimp_attrs),
  1094            (ctrN, ctr_thmss, []),
  1095            (discN, disc_thmss, simp_attrs),
  1096            (selN, sel_thmss, simp_attrs),
  1097            (simpsN, simp_thmss, []),
  1098            (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
  1099           |> maps (fn (thmN, thmss, attrs) =>
  1100             map2 (fn fun_name => fn thms =>
  1101                 ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
  1102               fun_names (take actual_nn thmss))
  1103           |> filter_out (null o fst o hd o snd);
  1104 
  1105         val common_notes =
  1106           [(coinductN, if n2m then [coinduct_thm] else [], []),
  1107            (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
  1108           |> filter_out (null o #2)
  1109           |> map (fn (thmN, thms, attrs) =>
  1110             ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
  1111       in
  1112         lthy |> Local_Theory.notes (notes @ common_notes) |> snd
  1113       end;
  1114 
  1115     fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
  1116   in
  1117     (goalss, after_qed, lthy')
  1118   end;
  1119 
  1120 fun add_primcorec_ursive_cmd maybe_tac seq (raw_fixes, raw_specs') lthy =
  1121   let
  1122     val (raw_specs, of_specs) = split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
  1123     val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
  1124   in
  1125     add_primcorec_ursive maybe_tac seq fixes specs of_specs lthy
  1126     handle ERROR str => primrec_error str
  1127   end
  1128   handle Primrec_Error (str, eqns) =>
  1129     if null eqns
  1130     then error ("primcorec error:\n  " ^ str)
  1131     else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
  1132       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
  1133 
  1134 val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
  1135   lthy
  1136   |> Proof.theorem NONE after_qed goalss
  1137   |> Proof.refine (Method.primitive_text I)
  1138   |> Seq.hd) ooo add_primcorec_ursive_cmd NONE;
  1139 
  1140 val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
  1141   lthy
  1142   |> after_qed (map (fn [] => []
  1143       | _ => primrec_error "need exclusiveness proofs - use primcorecursive instead of primcorec")
  1144     goalss)) ooo add_primcorec_ursive_cmd (SOME (fn {context = ctxt, ...} => auto_tac ctxt));
  1145 
  1146 end;