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