src/HOL/BNF/Tools/bnf_lfp_rec_sugar.ML
author blanchet
Mon, 04 Nov 2013 16:53:43 +0100
changeset 55698 8fdb4dc08ed1
child 55708 4843082be7ef
permissions -rw-r--r--
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
     1 (*  Title:      HOL/BNF/Tools/bnf_lfp_rec_sugar.ML
     2     Author:     Lorenz Panny, TU Muenchen
     3     Author:     Jasmin Blanchette, TU Muenchen
     4     Copyright   2013
     5 
     6 Recursor sugar.
     7 *)
     8 
     9 signature BNF_LFP_REC_SUGAR =
    10 sig
    11   val add_primrec: (binding * typ option * mixfix) list ->
    12     (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
    13   val add_primrec_cmd: (binding * string option * mixfix) list ->
    14     (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
    15   val add_primrec_global: (binding * typ option * mixfix) list ->
    16     (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
    17   val add_primrec_overloaded: (string * (string * typ) * bool) list ->
    18     (binding * typ option * mixfix) list ->
    19     (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
    20   val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
    21     local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
    22 end;
    23 
    24 structure BNF_LFP_Rec_Sugar : BNF_LFP_REC_SUGAR =
    25 struct
    26 
    27 open Ctr_Sugar
    28 open BNF_Util
    29 open BNF_Tactics
    30 open BNF_Def
    31 open BNF_FP_Util
    32 open BNF_FP_Def_Sugar
    33 open BNF_FP_N2M_Sugar
    34 open BNF_FP_Rec_Sugar_Util
    35 
    36 val nitpicksimp_attrs = @{attributes [nitpick_simp]};
    37 val simp_attrs = @{attributes [simp]};
    38 val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
    39 
    40 exception Primrec_Error of string * term list;
    41 
    42 fun primrec_error str = raise Primrec_Error (str, []);
    43 fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
    44 fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
    45 
    46 datatype rec_call =
    47   No_Rec of int * typ |
    48   Mutual_Rec of (int * typ) * (int * typ) |
    49   Nested_Rec of int * typ;
    50 
    51 type rec_ctr_spec =
    52   {ctr: term,
    53    offset: int,
    54    calls: rec_call list,
    55    rec_thm: thm};
    56 
    57 type rec_spec =
    58   {recx: term,
    59    nested_map_idents: thm list,
    60    nested_map_comps: thm list,
    61    ctr_specs: rec_ctr_spec list};
    62 
    63 exception AINT_NO_MAP of term;
    64 
    65 fun ill_formed_rec_call ctxt t =
    66   error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
    67 fun invalid_map ctxt t =
    68   error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
    69 fun unexpected_rec_call ctxt t =
    70   error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
    71 
    72 fun massage_nested_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
    73   let
    74     fun check_no_call t = if has_call t then unexpected_rec_call ctxt t else ();
    75 
    76     val typof = curry fastype_of1 bound_Ts;
    77     val build_map_fst = build_map ctxt (fst_const o fst);
    78 
    79     val yT = typof y;
    80     val yU = typof y';
    81 
    82     fun y_of_y' () = build_map_fst (yU, yT) $ y';
    83     val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
    84 
    85     fun massage_mutual_fun U T t =
    86       (case t of
    87         Const (@{const_name comp}, _) $ t1 $ t2 =>
    88         mk_comp bound_Ts (tap check_no_call t1, massage_mutual_fun U T t2)
    89       | _ =>
    90         if has_call t then
    91           (case try HOLogic.dest_prodT U of
    92             SOME (U1, U2) => if U1 = T then raw_massage_fun T U2 t else invalid_map ctxt t
    93           | NONE => invalid_map ctxt t)
    94         else
    95           mk_comp bound_Ts (t, build_map_fst (U, T)));
    96 
    97     fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
    98         (case try (dest_map ctxt s) t of
    99           SOME (map0, fs) =>
   100           let
   101             val Type (_, ran_Ts) = range_type (typof t);
   102             val map' = mk_map (length fs) Us ran_Ts map0;
   103             val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
   104           in
   105             Term.list_comb (map', fs')
   106           end
   107         | NONE => raise AINT_NO_MAP t)
   108       | massage_map _ _ t = raise AINT_NO_MAP t
   109     and massage_map_or_map_arg U T t =
   110       if T = U then
   111         tap check_no_call t
   112       else
   113         massage_map U T t
   114         handle AINT_NO_MAP _ => massage_mutual_fun U T t;
   115 
   116     fun massage_call (t as t1 $ t2) =
   117         if has_call t then
   118           if t2 = y then
   119             massage_map yU yT (elim_y t1) $ y'
   120             handle AINT_NO_MAP t' => invalid_map ctxt t'
   121           else
   122             let val (g, xs) = Term.strip_comb t2 in
   123               if g = y then
   124                 if exists has_call xs then unexpected_rec_call ctxt t2
   125                 else Term.list_comb (massage_call (mk_compN (length xs) bound_Ts (t1, y)), xs)
   126               else
   127                 ill_formed_rec_call ctxt t
   128             end
   129         else
   130           elim_y t
   131       | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
   132   in
   133     massage_call
   134   end;
   135 
   136 fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
   137   let
   138     val thy = Proof_Context.theory_of lthy;
   139 
   140     val ((missing_arg_Ts, perm0_kks,
   141           fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
   142             co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy') =
   143       nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy;
   144 
   145     val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
   146 
   147     val indices = map #index fp_sugars;
   148     val perm_indices = map #index perm_fp_sugars;
   149 
   150     val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
   151     val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
   152     val perm_lfpTs = map (body_type o fastype_of o hd) perm_ctrss;
   153 
   154     val nn0 = length arg_Ts;
   155     val nn = length perm_lfpTs;
   156     val kks = 0 upto nn - 1;
   157     val perm_ns = map length perm_ctr_Tsss;
   158     val perm_mss = map (map length) perm_ctr_Tsss;
   159 
   160     val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
   161       perm_fp_sugars;
   162     val perm_fun_arg_Tssss =
   163       mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
   164 
   165     fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
   166     fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
   167 
   168     val induct_thms = unpermute0 (conj_dests nn induct_thm);
   169 
   170     val lfpTs = unpermute perm_lfpTs;
   171     val Cs = unpermute perm_Cs;
   172 
   173     val As_rho = tvar_subst thy (take nn0 lfpTs) arg_Ts;
   174     val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
   175 
   176     val substA = Term.subst_TVars As_rho;
   177     val substAT = Term.typ_subst_TVars As_rho;
   178     val substCT = Term.typ_subst_TVars Cs_rho;
   179     val substACT = substAT o substCT;
   180 
   181     val perm_Cs' = map substCT perm_Cs;
   182 
   183     fun offset_of_ctr 0 _ = 0
   184       | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
   185         length ctrs + offset_of_ctr (n - 1) ctr_sugars;
   186 
   187     fun call_of [i] [T] = (if exists_subtype_in Cs T then Nested_Rec else No_Rec) (i, substACT T)
   188       | call_of [i, i'] [T, T'] = Mutual_Rec ((i, substACT T), (i', substACT T'));
   189 
   190     fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
   191       let
   192         val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
   193         val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
   194         val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
   195       in
   196         {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
   197          rec_thm = rec_thm}
   198       end;
   199 
   200     fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
   201       let
   202         val ctrs = #ctrs (nth ctr_sugars index);
   203         val rec_thmss = co_rec_of (nth iter_thmsss index);
   204         val k = offset_of_ctr index ctr_sugars;
   205         val n = length ctrs;
   206       in
   207         map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thmss
   208       end;
   209 
   210     fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
   211       : fp_sugar) =
   212       {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
   213        nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
   214        nested_map_comps = map map_comp_of_bnf nested_bnfs,
   215        ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
   216   in
   217     ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms),
   218      lthy')
   219   end;
   220 
   221 val undef_const = Const (@{const_name undefined}, dummyT);
   222 
   223 fun permute_args n t =
   224   list_comb (t, map Bound (0 :: (n downto 1))) |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
   225 
   226 type eqn_data = {
   227   fun_name: string,
   228   rec_type: typ,
   229   ctr: term,
   230   ctr_args: term list,
   231   left_args: term list,
   232   right_args: term list,
   233   res_type: typ,
   234   rhs_term: term,
   235   user_eqn: term
   236 };
   237 
   238 fun dissect_eqn lthy fun_names eqn' =
   239   let
   240     val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
   241       handle TERM _ =>
   242         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
   243     val (lhs, rhs) = HOLogic.dest_eq eqn
   244         handle TERM _ =>
   245           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
   246     val (fun_name, args) = strip_comb lhs
   247       |>> (fn x => if is_Free x then fst (dest_Free x)
   248           else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
   249     val (left_args, rest) = take_prefix is_Free args;
   250     val (nonfrees, right_args) = take_suffix is_Free rest;
   251     val num_nonfrees = length nonfrees;
   252     val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
   253       primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
   254       primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
   255     val _ = member (op =) fun_names fun_name orelse
   256       primrec_error_eqn "malformed function equation (does not start with function name)" eqn
   257 
   258     val (ctr, ctr_args) = strip_comb (the_single nonfrees);
   259     val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
   260       primrec_error_eqn "partially applied constructor in pattern" eqn;
   261     val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
   262       primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
   263         "\" in left-hand side") eqn end;
   264     val _ = forall is_Free ctr_args orelse
   265       primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
   266     val _ =
   267       let val b = fold_aterms (fn x as Free (v, _) =>
   268         if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
   269         not (member (op =) fun_names v) andalso
   270         not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
   271       in
   272         null b orelse
   273         primrec_error_eqn ("extra variable(s) in right-hand side: " ^
   274           commas (map (Syntax.string_of_term lthy) b)) eqn
   275       end;
   276   in
   277     {fun_name = fun_name,
   278      rec_type = body_type (type_of ctr),
   279      ctr = ctr,
   280      ctr_args = ctr_args,
   281      left_args = left_args,
   282      right_args = right_args,
   283      res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
   284      rhs_term = rhs,
   285      user_eqn = eqn'}
   286   end;
   287 
   288 fun rewrite_map_arg get_ctr_pos rec_type res_type =
   289   let
   290     val pT = HOLogic.mk_prodT (rec_type, res_type);
   291 
   292     val maybe_suc = Option.map (fn x => x + 1);
   293     fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
   294       | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
   295       | subst d t =
   296         let
   297           val (u, vs) = strip_comb t;
   298           val ctr_pos = try (get_ctr_pos o fst o dest_Free) u |> the_default ~1;
   299         in
   300           if ctr_pos >= 0 then
   301             if d = SOME ~1 andalso length vs = ctr_pos then
   302               list_comb (permute_args ctr_pos (snd_const pT), vs)
   303             else if length vs > ctr_pos andalso is_some d
   304                 andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
   305               list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
   306             else
   307               primrec_error_eqn ("recursive call not directly applied to constructor argument") t
   308           else
   309             list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
   310         end
   311   in
   312     subst (SOME ~1)
   313   end;
   314 
   315 fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
   316   let
   317     fun try_nested_rec bound_Ts y t =
   318       AList.lookup (op =) nested_calls y
   319       |> Option.map (fn y' =>
   320         massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
   321 
   322     fun subst bound_Ts (t as g' $ y) =
   323         let
   324           fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
   325           val y_head = head_of y;
   326         in
   327           if not (member (op =) ctr_args y_head) then
   328             subst_rec ()
   329           else
   330             (case try_nested_rec bound_Ts y_head t of
   331               SOME t' => t'
   332             | NONE =>
   333               let val (g, g_args) = strip_comb g' in
   334                 (case try (get_ctr_pos o fst o dest_Free) g of
   335                   SOME ctr_pos =>
   336                   (length g_args >= ctr_pos orelse
   337                    primrec_error_eqn "too few arguments in recursive call" t;
   338                    (case AList.lookup (op =) mutual_calls y of
   339                      SOME y' => list_comb (y', g_args)
   340                    | NONE => subst_rec ()))
   341                 | NONE => subst_rec ())
   342               end)
   343         end
   344       | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
   345       | subst _ t = t
   346 
   347     fun subst' t =
   348       if has_call t then
   349         (* FIXME detect this case earlier? *)
   350         primrec_error_eqn "recursive call not directly applied to constructor argument" t
   351       else
   352         try_nested_rec [] (head_of t) t |> the_default t
   353   in
   354     subst' o subst []
   355   end;
   356 
   357 fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
   358     (maybe_eqn_data : eqn_data option) =
   359   (case maybe_eqn_data of
   360     NONE => undef_const
   361   | SOME {ctr_args, left_args, right_args, rhs_term = t, ...} =>
   362     let
   363       val calls = #calls ctr_spec;
   364       val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
   365 
   366       val no_calls' = tag_list 0 calls
   367         |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
   368       val mutual_calls' = tag_list 0 calls
   369         |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
   370       val nested_calls' = tag_list 0 calls
   371         |> map_filter (try (apsnd (fn Nested_Rec p => p)));
   372 
   373       val args = replicate n_args ("", dummyT)
   374         |> Term.rename_wrt_term t
   375         |> map Free
   376         |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
   377             nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
   378           no_calls'
   379         |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
   380             nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
   381           mutual_calls'
   382         |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
   383             nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
   384           nested_calls';
   385 
   386       val fun_name_ctr_pos_list =
   387         map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
   388       val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
   389       val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
   390       val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
   391     in
   392       t
   393       |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
   394       |> fold_rev lambda (args @ left_args @ right_args)
   395     end);
   396 
   397 fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
   398   let
   399     val n_funs = length funs_data;
   400 
   401     val ctr_spec_eqn_data_list' =
   402       (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
   403       |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
   404           ##> (fn x => null x orelse
   405             primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
   406     val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
   407       primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
   408 
   409     val ctr_spec_eqn_data_list =
   410       ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
   411 
   412     val recs = take n_funs rec_specs |> map #recx;
   413     val rec_args = ctr_spec_eqn_data_list
   414       |> sort ((op <) o pairself (#offset o fst) |> make_ord)
   415       |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
   416     val ctr_poss = map (fn x =>
   417       if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
   418         primrec_error ("inconstant constructor pattern position for function " ^
   419           quote (#fun_name (hd x)))
   420       else
   421         hd x |> #left_args |> length) funs_data;
   422   in
   423     (recs, ctr_poss)
   424     |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
   425     |> Syntax.check_terms lthy
   426     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
   427       bs mxs
   428   end;
   429 
   430 fun find_rec_calls has_call ({ctr, ctr_args, rhs_term, ...} : eqn_data) =
   431   let
   432     fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
   433       | find bound_Ts (t as _ $ _) ctr_arg =
   434         let
   435           val typof = curry fastype_of1 bound_Ts;
   436           val (f', args') = strip_comb t;
   437           val n = find_index (equal ctr_arg o head_of) args';
   438         in
   439           if n < 0 then
   440             find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
   441           else
   442             let
   443               val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
   444               val (arg_head, arg_args) = Term.strip_comb arg;
   445             in
   446               if has_call f then
   447                 mk_partial_compN (length arg_args) (typof arg_head) f ::
   448                 maps (fn x => find bound_Ts x ctr_arg) args
   449               else
   450                 find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
   451             end
   452         end
   453       | find _ _ _ = [];
   454   in
   455     map (find [] rhs_term) ctr_args
   456     |> (fn [] => NONE | callss => SOME (ctr, callss))
   457   end;
   458 
   459 fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
   460   unfold_thms_tac ctxt fun_defs THEN
   461   HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
   462   unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
   463   HEADGOAL (rtac refl);
   464 
   465 fun prepare_primrec fixes specs lthy =
   466   let
   467     val (bs, mxs) = map_split (apfst fst) fixes;
   468     val fun_names = map Binding.name_of bs;
   469     val eqns_data = map (dissect_eqn lthy fun_names) specs;
   470     val funs_data = eqns_data
   471       |> partition_eq ((op =) o pairself #fun_name)
   472       |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
   473       |> map (fn (x, y) => the_single y handle List.Empty =>
   474           primrec_error ("missing equations for function " ^ quote x));
   475 
   476     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   477     val arg_Ts = map (#rec_type o hd) funs_data;
   478     val res_Ts = map (#res_type o hd) funs_data;
   479     val callssss = funs_data
   480       |> map (partition_eq ((op =) o pairself #ctr))
   481       |> map (maps (map_filter (find_rec_calls has_call)));
   482 
   483     val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
   484       rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
   485 
   486     val actual_nn = length funs_data;
   487 
   488     val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
   489       map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
   490         primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
   491           " is not a constructor in left-hand side") user_eqn) eqns_data end;
   492 
   493     val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
   494 
   495     fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
   496         (fun_data : eqn_data list) =
   497       let
   498         val def_thms = map (snd o snd) def_thms';
   499         val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
   500           |> fst
   501           |> map_filter (try (fn (x, [y]) =>
   502             (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
   503           |> map (fn (user_eqn, num_extra_args, rec_thm) =>
   504             mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
   505             |> K |> Goal.prove lthy [] [] user_eqn
   506             |> Thm.close_derivation);
   507         val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
   508       in
   509         (poss, simp_thmss)
   510       end;
   511 
   512     val notes =
   513       (if n2m then map2 (fn name => fn thm =>
   514         (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
   515       |> map (fn (prefix, thmN, thms, attrs) =>
   516         ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
   517 
   518     val common_name = mk_common_name fun_names;
   519 
   520     val common_notes =
   521       (if n2m then [(inductN, [induct_thm], [])] else [])
   522       |> map (fn (thmN, thms, attrs) =>
   523         ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
   524   in
   525     (((fun_names, defs),
   526       fn lthy => fn defs =>
   527         split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
   528       lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
   529   end;
   530 
   531 (* primrec definition *)
   532 
   533 fun add_primrec_simple fixes ts lthy =
   534   let
   535     val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
   536       handle ERROR str => primrec_error str;
   537   in
   538     lthy
   539     |> fold_map Local_Theory.define defs
   540     |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
   541   end
   542   handle Primrec_Error (str, eqns) =>
   543     if null eqns
   544     then error ("primrec_new error:\n  " ^ str)
   545     else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
   546       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
   547 
   548 local
   549 
   550 fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
   551   let
   552     val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
   553     val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
   554 
   555     val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
   556 
   557     val mk_notes =
   558       flat ooo map3 (fn poss => fn prefix => fn thms =>
   559         let
   560           val (bs, attrss) = map_split (fst o nth specs) poss;
   561           val notes =
   562             map3 (fn b => fn attrs => fn thm =>
   563               ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
   564             bs attrss thms;
   565         in
   566           ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
   567         end);
   568   in
   569     lthy
   570     |> add_primrec_simple fixes (map snd specs)
   571     |-> (fn (names, (ts, (posss, simpss))) =>
   572       Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
   573       #> Local_Theory.notes (mk_notes posss names simpss)
   574       #>> pair ts o map snd)
   575   end;
   576 
   577 in
   578 
   579 val add_primrec = gen_primrec Specification.check_spec;
   580 val add_primrec_cmd = gen_primrec Specification.read_spec;
   581 
   582 end;
   583 
   584 fun add_primrec_global fixes specs thy =
   585   let
   586     val lthy = Named_Target.theory_init thy;
   587     val ((ts, simps), lthy') = add_primrec fixes specs lthy;
   588     val simps' = burrow (Proof_Context.export lthy' lthy) simps;
   589   in ((ts, simps'), Local_Theory.exit_global lthy') end;
   590 
   591 fun add_primrec_overloaded ops fixes specs thy =
   592   let
   593     val lthy = Overloading.overloading ops thy;
   594     val ((ts, simps), lthy') = add_primrec fixes specs lthy;
   595     val simps' = burrow (Proof_Context.export lthy' lthy) simps;
   596   in ((ts, simps'), Local_Theory.exit_global lthy') end;
   597 
   598 end;