src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
author panny
Thu, 19 Sep 2013 00:32:33 +0200
changeset 54859 e176d6d3345f
parent 54857 03fac7082137
child 54862 9e64151359e8
permissions -rw-r--r--
generate more theorems (e.g. for types with only one constructor)
     1 (*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar.ML
     2     Author:     Lorenz Panny, TU Muenchen
     3     Copyright   2013
     4 
     5 Recursor and corecursor sugar.
     6 *)
     7 
     8 signature BNF_FP_REC_SUGAR =
     9 sig
    10   val add_primrec_cmd: (binding * string option * mixfix) list ->
    11     (Attrib.binding * string) list -> local_theory -> local_theory;
    12   val add_primcorec_cmd: bool ->
    13     (binding * string option * mixfix) list * (Attrib.binding * string) list -> Proof.context ->
    14     Proof.state
    15 end;
    16 
    17 structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
    18 struct
    19 
    20 open BNF_Util
    21 open BNF_FP_Util
    22 open BNF_FP_Rec_Sugar_Util
    23 open BNF_FP_Rec_Sugar_Tactics
    24 
    25 exception Primrec_Error of string * term list;
    26 
    27 fun primrec_error str = raise Primrec_Error (str, []);
    28 fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
    29 fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
    30 
    31 fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
    32 
    33 val free_name = try (fn Free (v, _) => v);
    34 val const_name = try (fn Const (v, _) => v);
    35 val undef_const = Const (@{const_name undefined}, dummyT);
    36 
    37 fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
    38   |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
    39 val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
    40 fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
    41   strip_qnt_body @{const_name all} t)
    42 fun mk_not @{const True} = @{const False}
    43   | mk_not @{const False} = @{const True}
    44   | mk_not (@{const Not} $ t) = t
    45   | mk_not (@{const Trueprop} $ t) = @{const Trueprop} $ mk_not t
    46   | mk_not t = HOLogic.mk_not t
    47 val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
    48 val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
    49 fun invert_prems [t] = map mk_not (HOLogic.disjuncts t)
    50   | invert_prems ts = [mk_disjs (map mk_not ts)];
    51 fun invert_prems_disj [t] = map mk_not (HOLogic.disjuncts t)
    52   | invert_prems_disj ts = [mk_disjs (map (mk_conjs o map mk_not o HOLogic.disjuncts) ts)];
    53 fun abstract vs =
    54   let fun a n (t $ u) = a n t $ a n u
    55         | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
    56         | a n t = let val idx = find_index (equal t) vs in
    57             if idx < 0 then t else Bound (n + idx) end
    58   in a 0 end;
    59 
    60 val simp_attrs = @{attributes [simp]};
    61 
    62 
    63 (* Primrec *)
    64 
    65 type eqn_data = {
    66   fun_name: string,
    67   rec_type: typ,
    68   ctr: term,
    69   ctr_args: term list,
    70   left_args: term list,
    71   right_args: term list,
    72   res_type: typ,
    73   rhs_term: term,
    74   user_eqn: term
    75 };
    76 
    77 fun dissect_eqn lthy fun_names eqn' =
    78   let
    79     val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
    80       handle TERM _ =>
    81         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    82     val (lhs, rhs) = HOLogic.dest_eq eqn
    83         handle TERM _ =>
    84           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    85     val (fun_name, args) = strip_comb lhs
    86       |>> (fn x => if is_Free x then fst (dest_Free x)
    87           else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
    88     val (left_args, rest) = take_prefix is_Free args;
    89     val (nonfrees, right_args) = take_suffix is_Free rest;
    90     val _ = length nonfrees = 1 orelse if length nonfrees = 0 then
    91       primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
    92       primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
    93     val _ = member (op =) fun_names fun_name orelse
    94       primrec_error_eqn "malformed function equation (does not start with function name)" eqn
    95 
    96     val (ctr, ctr_args) = strip_comb (the_single nonfrees);
    97     val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
    98       primrec_error_eqn "partially applied constructor in pattern" eqn;
    99     val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
   100       primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
   101         "\" in left-hand side") eqn end;
   102     val _ = forall is_Free ctr_args orelse
   103       primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
   104     val _ =
   105       let val b = fold_aterms (fn x as Free (v, _) =>
   106         if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
   107         not (member (op =) fun_names v) andalso
   108         not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
   109       in
   110         null b orelse
   111         primrec_error_eqn ("extra variable(s) in right-hand side: " ^
   112           commas (map (Syntax.string_of_term lthy) b)) eqn
   113       end;
   114   in
   115     {fun_name = fun_name,
   116      rec_type = body_type (type_of ctr),
   117      ctr = ctr,
   118      ctr_args = ctr_args,
   119      left_args = left_args,
   120      right_args = right_args,
   121      res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
   122      rhs_term = rhs,
   123      user_eqn = eqn'}
   124   end;
   125 
   126 fun rewrite_map_arg get_ctr_pos rec_type res_type =
   127   let
   128     val pT = HOLogic.mk_prodT (rec_type, res_type);
   129 
   130     val maybe_suc = Option.map (fn x => x + 1);
   131     fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
   132       | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
   133       | subst d t =
   134         let
   135           val (u, vs) = strip_comb t;
   136           val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1;
   137         in
   138           if ctr_pos >= 0 then
   139             if d = SOME ~1 andalso length vs = ctr_pos then
   140               list_comb (permute_args ctr_pos (snd_const pT), vs)
   141             else if length vs > ctr_pos andalso is_some d
   142                 andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
   143               list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
   144             else
   145               primrec_error_eqn ("recursive call not directly applied to constructor argument") t
   146           else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then
   147             list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs)
   148           else
   149             list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
   150         end
   151   in
   152     subst (SOME ~1)
   153   end;
   154 
   155 fun subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls t =
   156   let
   157     fun subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
   158       | subst bound_Ts (t as g' $ y) =
   159         let
   160           val maybe_direct_y' = AList.lookup (op =) direct_calls y;
   161           val maybe_indirect_y' = AList.lookup (op =) indirect_calls y;
   162           val (g, g_args) = strip_comb g';
   163           val ctr_pos = try (get_ctr_pos o the) (free_name g) |> the_default ~1;
   164           val _ = ctr_pos < 0 orelse length g_args >= ctr_pos orelse
   165             primrec_error_eqn "too few arguments in recursive call" t;
   166         in
   167           if not (member (op =) ctr_args y) then
   168             pairself (subst bound_Ts) (g', y) |> (op $)
   169           else if ctr_pos >= 0 then
   170             list_comb (the maybe_direct_y', g_args)
   171           else if is_some maybe_indirect_y' then
   172             (if has_call g' then t else y)
   173             |> massage_indirect_rec_call lthy has_call
   174               (rewrite_map_arg get_ctr_pos) bound_Ts y (the maybe_indirect_y')
   175             |> (if has_call g' then I else curry (op $) g')
   176           else
   177             t
   178         end
   179       | subst _ t = t
   180   in
   181     subst [] t
   182     |> tap (fn u => has_call u andalso (* FIXME detect this case earlier *)
   183       primrec_error_eqn "recursive call not directly applied to constructor argument" t)
   184   end;
   185 
   186 fun build_rec_arg lthy funs_data has_call ctr_spec maybe_eqn_data =
   187   if is_none maybe_eqn_data then undef_const else
   188     let
   189       val eqn_data = the maybe_eqn_data;
   190       val t = #rhs_term eqn_data;
   191       val ctr_args = #ctr_args eqn_data;
   192 
   193       val calls = #calls ctr_spec;
   194       val n_args = fold (curry (op +) o (fn Direct_Rec _ => 2 | _ => 1)) calls 0;
   195 
   196       val no_calls' = tag_list 0 calls
   197         |> map_filter (try (apsnd (fn No_Rec n => n | Direct_Rec (n, _) => n)));
   198       val direct_calls' = tag_list 0 calls
   199         |> map_filter (try (apsnd (fn Direct_Rec (_, n) => n)));
   200       val indirect_calls' = tag_list 0 calls
   201         |> map_filter (try (apsnd (fn Indirect_Rec n => n)));
   202 
   203       fun make_direct_type T = dummyT; (* FIXME? *)
   204 
   205       val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data;
   206 
   207       fun make_indirect_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T =>
   208         let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in
   209           if is_some maybe_res_type
   210           then HOLogic.mk_prodT (T, the maybe_res_type)
   211           else make_indirect_type T end))
   212         | make_indirect_type T = T;
   213 
   214       val args = replicate n_args ("", dummyT)
   215         |> Term.rename_wrt_term t
   216         |> map Free
   217         |> fold (fn (ctr_arg_idx, arg_idx) =>
   218             nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
   219           no_calls'
   220         |> fold (fn (ctr_arg_idx, arg_idx) =>
   221             nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_direct_type)))
   222           direct_calls'
   223         |> fold (fn (ctr_arg_idx, arg_idx) =>
   224             nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_indirect_type)))
   225           indirect_calls';
   226 
   227       val fun_name_ctr_pos_list =
   228         map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
   229       val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
   230       val direct_calls = map (apfst (nth ctr_args) o apsnd (nth args)) direct_calls';
   231       val indirect_calls = map (apfst (nth ctr_args) o apsnd (nth args)) indirect_calls';
   232 
   233       val abstractions = args @ #left_args eqn_data @ #right_args eqn_data;
   234     in
   235       t
   236       |> subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls
   237       |> fold_rev lambda abstractions
   238     end;
   239 
   240 fun build_defs lthy bs mxs funs_data rec_specs has_call =
   241   let
   242     val n_funs = length funs_data;
   243 
   244     val ctr_spec_eqn_data_list' =
   245       (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
   246       |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
   247           ##> (fn x => null x orelse
   248             primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
   249     val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
   250       primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
   251 
   252     val ctr_spec_eqn_data_list =
   253       ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
   254 
   255     val recs = take n_funs rec_specs |> map #recx;
   256     val rec_args = ctr_spec_eqn_data_list
   257       |> sort ((op <) o pairself (#offset o fst) |> make_ord)
   258       |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
   259     val ctr_poss = map (fn x =>
   260       if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
   261         primrec_error ("inconstant constructor pattern position for function " ^
   262           quote (#fun_name (hd x)))
   263       else
   264         hd x |> #left_args |> length) funs_data;
   265   in
   266     (recs, ctr_poss)
   267     |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
   268     |> Syntax.check_terms lthy
   269     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
   270   end;
   271 
   272 fun find_rec_calls has_call eqn_data =
   273   let
   274     fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg
   275       | find (t as _ $ _) ctr_arg =
   276         let
   277           val (f', args') = strip_comb t;
   278           val n = find_index (equal ctr_arg) args';
   279         in
   280           if n < 0 then
   281             find f' ctr_arg @ maps (fn x => find x ctr_arg) args'
   282           else
   283             let val (f, args) = chop n args' |>> curry list_comb f' in
   284               if has_call f then
   285                 f :: maps (fn x => find x ctr_arg) args
   286               else
   287                 find f ctr_arg @ maps (fn x => find x ctr_arg) args
   288             end
   289         end
   290       | find _ _ = [];
   291   in
   292     map (find (#rhs_term eqn_data)) (#ctr_args eqn_data)
   293     |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
   294   end;
   295 
   296 fun add_primrec fixes specs lthy =
   297   let
   298     val (bs, mxs) = map_split (apfst fst) fixes;
   299     val fun_names = map Binding.name_of bs;
   300     val eqns_data = map (snd #> dissect_eqn lthy fun_names) specs;
   301     val funs_data = eqns_data
   302       |> partition_eq ((op =) o pairself #fun_name)
   303       |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
   304       |> map (fn (x, y) => the_single y handle List.Empty =>
   305           primrec_error ("missing equations for function " ^ quote x));
   306 
   307     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   308     val arg_Ts = map (#rec_type o hd) funs_data;
   309     val res_Ts = map (#res_type o hd) funs_data;
   310     val callssss = funs_data
   311       |> map (partition_eq ((op =) o pairself #ctr))
   312       |> map (maps (map_filter (find_rec_calls has_call)));
   313 
   314     fun get_indices t = map (fst #>> Binding.name_of #> Free) fixes
   315       |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   316       |> map_filter I;
   317     val ((nontriv, rec_specs, _, induct_thm, induct_thms), lthy') =
   318       rec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
   319 
   320     val actual_nn = length funs_data;
   321 
   322     val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
   323       map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
   324         primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
   325           " is not a constructor in left-hand side") user_eqn) eqns_data end;
   326 
   327     val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
   328 
   329     fun prove def_thms' {ctr_specs, nested_map_idents, nested_map_comps, ...} induct_thm fun_data
   330         lthy =
   331       let
   332         val fun_name = #fun_name (hd fun_data);
   333         val def_thms = map (snd o snd) def_thms';
   334         val simp_thms = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
   335           |> fst
   336           |> map_filter (try (fn (x, [y]) =>
   337             (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
   338           |> map (fn (user_eqn, num_extra_args, rec_thm) =>
   339             mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
   340             |> K |> Goal.prove lthy [] [] user_eqn)
   341 
   342         val notes =
   343           [(inductN, if actual_nn > 1 then [induct_thm] else [], []),
   344            (simpsN, simp_thms, simp_attrs)]
   345           |> filter_out (null o #2)
   346           |> map (fn (thmN, thms, attrs) =>
   347             ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]));
   348       in
   349         lthy |> Local_Theory.notes notes
   350       end;
   351 
   352     val common_name = mk_common_name fun_names;
   353 
   354     val common_notes =
   355       [(inductN, if nontriv andalso actual_nn > 1 then [induct_thm] else [], [])]
   356       |> filter_out (null o #2)
   357       |> map (fn (thmN, thms, attrs) =>
   358         ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
   359   in
   360     lthy'
   361     |> fold_map Local_Theory.define defs
   362     |-> snd oo (fn def_thms' => fold_map3 (prove def_thms') (take actual_nn rec_specs)
   363       (take actual_nn induct_thms) funs_data)
   364     |> snd o Local_Theory.notes common_notes
   365   end;
   366 
   367 fun add_primrec_cmd raw_fixes raw_specs lthy =
   368   let
   369     val _ = let val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes) in null d orelse
   370       primrec_error ("duplicate function name(s): " ^ commas d) end;
   371     val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy);
   372   in
   373     add_primrec fixes specs lthy
   374       handle ERROR str => primrec_error str
   375   end
   376   handle Primrec_Error (str, eqns) =>
   377     if null eqns
   378     then error ("primrec_new error:\n  " ^ str)
   379     else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
   380       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns))
   381 
   382 
   383 
   384 (* Primcorec *)
   385 
   386 type co_eqn_data_disc = {
   387   fun_name: string,
   388   fun_T: typ,
   389   fun_args: term list,
   390   ctr: term,
   391   ctr_no: int, (*###*)
   392   disc: term,
   393   prems: term list,
   394   user_eqn: term
   395 };
   396 type co_eqn_data_sel = {
   397   fun_name: string,
   398   fun_T: typ,
   399   fun_args: term list,
   400   ctr: term,
   401   sel: term,
   402   rhs_term: term,
   403   user_eqn: term
   404 };
   405 datatype co_eqn_data =
   406   Disc of co_eqn_data_disc |
   407   Sel of co_eqn_data_sel;
   408 
   409 fun co_dissect_eqn_disc sequential fun_names corec_specs prems' concl matchedsss =
   410   let
   411     fun find_subterm p = let (* FIXME \<exists>? *)
   412       fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
   413         | f t = if p t then SOME t else NONE
   414       in f end;
   415 
   416     val applied_fun = concl
   417       |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
   418       |> the
   419       handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
   420     val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
   421     val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   422 
   423     val discs = map #disc ctr_specs;
   424     val ctrs = map #ctr ctr_specs;
   425     val not_disc = head_of concl = @{term Not};
   426     val _ = not_disc andalso length ctrs <> 2 andalso
   427       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
   428     val disc = find_subterm (member (op =) discs o head_of) concl;
   429     val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   430         |> (fn SOME t => let val n = find_index (equal t) ctrs in
   431           if n >= 0 then SOME n else NONE end | _ => NONE);
   432     val _ = is_some disc orelse is_some eq_ctr0 orelse
   433       primrec_error_eqn "no discriminator in equation" concl;
   434     val ctr_no' =
   435       if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
   436     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
   437     val ctr = #ctr (nth ctr_specs ctr_no);
   438 
   439     val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
   440     val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
   441     val prems = map (abstract (List.rev fun_args)) prems';
   442     val real_prems =
   443       (if catch_all orelse sequential then maps invert_prems_disj matchedss else []) @
   444       (if catch_all then [] else prems);
   445 
   446     val matchedsss' = AList.delete (op =) fun_name matchedsss
   447       |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [real_prems]);
   448 
   449     val user_eqn =
   450       (real_prems, betapply (#disc (nth ctr_specs ctr_no), applied_fun))
   451       |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop
   452       |> Logic.list_implies;
   453   in
   454     (Disc {
   455       fun_name = fun_name,
   456       fun_T = fun_T,
   457       fun_args = fun_args,
   458       ctr = ctr,
   459       ctr_no = ctr_no,
   460       disc = #disc (nth ctr_specs ctr_no),
   461       prems = real_prems,
   462       user_eqn = user_eqn
   463     }, matchedsss')
   464   end;
   465 
   466 fun co_dissect_eqn_sel fun_names corec_specs eqn' eqn =
   467   let
   468     val (lhs, rhs) = HOLogic.dest_eq eqn
   469       handle TERM _ =>
   470         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   471     val sel = head_of lhs;
   472     val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
   473       handle TERM _ =>
   474         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   475     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name)
   476       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   477     val (ctr_spec, sel) = #ctr_specs corec_spec
   478       |> the o get_index (try (the o find_first (equal sel) o #sels))
   479       |>> nth (#ctr_specs corec_spec);
   480     val user_eqn = drop_All eqn';
   481   in
   482     Sel {
   483       fun_name = fun_name,
   484       fun_T = fun_T,
   485       fun_args = fun_args,
   486       ctr = #ctr ctr_spec,
   487       sel = sel,
   488       rhs_term = rhs,
   489       user_eqn = user_eqn
   490     }
   491   end;
   492 
   493 fun co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss =
   494   let 
   495     val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
   496     val fun_name = head_of lhs |> fst o dest_Free;
   497     val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   498     val (ctr, ctr_args) = strip_comb rhs;
   499     val {disc, sels, ...} = the (find_first (equal ctr o #ctr) ctr_specs)
   500       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   501 
   502     val disc_imp_rhs = betapply (disc, lhs);
   503     val (maybe_eqn_data_disc, matchedsss') = if length ctr_specs = 1
   504       then (NONE, matchedsss)
   505       else apfst SOME (co_dissect_eqn_disc
   506           sequential fun_names corec_specs imp_prems disc_imp_rhs matchedsss);
   507 
   508     val sel_imp_rhss = (sels ~~ ctr_args)
   509       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   510 
   511 val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
   512  (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
   513  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
   514 
   515     val eqns_data_sel =
   516       map (co_dissect_eqn_sel fun_names corec_specs eqn') sel_imp_rhss;
   517   in
   518     (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
   519   end;
   520 
   521 fun co_dissect_eqn sequential fun_names corec_specs eqn' matchedsss =
   522   let
   523     val eqn = drop_All eqn'
   524       handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   525     val (imp_prems, imp_rhs) = Logic.strip_horn eqn
   526       |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
   527 
   528     val head = imp_rhs
   529       |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
   530       |> head_of;
   531 
   532     val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
   533 
   534     val discs = maps #ctr_specs corec_specs |> map #disc;
   535     val sels = maps #ctr_specs corec_specs |> maps #sels;
   536     val ctrs = maps #ctr_specs corec_specs |> map #ctr;
   537   in
   538     if member (op =) discs head orelse
   539       is_some maybe_rhs andalso
   540         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   541       co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedsss
   542       |>> single
   543     else if member (op =) sels head then
   544       ([co_dissect_eqn_sel fun_names corec_specs eqn' imp_rhs], matchedsss)
   545     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
   546       co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss
   547     else
   548       primrec_error_eqn "malformed function equation" eqn
   549   end;
   550 
   551 fun build_corec_arg_disc ctr_specs {fun_args, ctr_no, prems, ...} =
   552   if is_none (#pred (nth ctr_specs ctr_no)) then I else
   553     mk_conjs prems
   554     |> curry subst_bounds (List.rev fun_args)
   555     |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
   556     |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
   557 
   558 fun build_corec_arg_no_call sel_eqns sel =
   559   find_first (equal sel o #sel) sel_eqns
   560   |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
   561   |> the_default undef_const
   562   |> K;
   563 
   564 fun build_corec_arg_direct_call lthy has_call sel_eqns sel =
   565   let
   566     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   567     fun rewrite is_end U _ t =
   568       if U = @{typ bool} then @{term True} |> has_call t ? K @{term False} (* stop? *)
   569       else if is_end = has_call t then undef_const
   570       else if is_end then t (* end *)
   571       else HOLogic.mk_tuple (snd (strip_comb t)); (* continue *)
   572     fun massage rhs_term is_end t = massage_direct_corec_call
   573       lthy has_call (rewrite is_end) [] (range_type (fastype_of t)) rhs_term;
   574   in
   575     if is_none maybe_sel_eqn then K I else
   576       abs_tuple (#fun_args (the maybe_sel_eqn)) oo massage (#rhs_term (the maybe_sel_eqn))
   577   end;
   578 
   579 fun build_corec_arg_indirect_call lthy has_call sel_eqns sel =
   580   let
   581     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   582     fun rewrite _ _ =
   583       let
   584         fun subst (Abs (v, T, b)) = Abs (v, T, subst b)
   585           | subst (t as _ $ _) =
   586             let val (u, vs) = strip_comb t in
   587               if is_Free u andalso has_call u then
   588                 Const (@{const_name Inr}, dummyT) $
   589                   (try (foldr1 (fn (x, y) => Const (@{const_name Pair}, dummyT) $ x $ y)) vs
   590                     |> the_default (hd vs))
   591               else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
   592                 list_comb (u |> map_types (K dummyT), map subst vs)
   593               else
   594                 list_comb (subst u, map subst vs)
   595             end
   596           | subst t = t;
   597       in
   598         subst
   599       end;
   600     fun massage rhs_term t = massage_indirect_corec_call
   601       lthy has_call rewrite [] (range_type (fastype_of t)) rhs_term;
   602   in
   603     if is_none maybe_sel_eqn then I else
   604       abs_tuple (#fun_args (the maybe_sel_eqn)) o massage (#rhs_term (the maybe_sel_eqn))
   605   end;
   606 
   607 fun build_corec_args_sel lthy has_call all_sel_eqns ctr_spec =
   608   let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
   609     if null sel_eqns then I else
   610       let
   611         val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
   612 
   613         val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
   614         val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list;
   615         val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list;
   616       in
   617         I
   618         #> fold (fn (sel, n) => nth_map n
   619           (build_corec_arg_no_call sel_eqns sel)) no_calls'
   620         #> fold (fn (sel, (q, g, h)) =>
   621           let val f = build_corec_arg_direct_call lthy has_call sel_eqns sel in
   622             nth_map h (f false) o nth_map g (f true) o nth_map q (f true) end) direct_calls'
   623         #> fold (fn (sel, n) => nth_map n
   624           (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
   625       end
   626   end;
   627 
   628 fun co_build_defs lthy bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss =
   629   let
   630     val corec_specs' = take (length bs) corec_specs;
   631     val corecs = map #corec corec_specs';
   632     val ctr_specss = map #ctr_specs corec_specs';
   633     val corec_args = hd corecs
   634       |> fst o split_last o binder_types o fastype_of
   635       |> map (Const o pair @{const_name undefined})
   636       |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
   637       |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
   638     fun currys Ts t = if length Ts <= 1 then t else
   639       t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
   640         (length Ts - 1 downto 0 |> map Bound)
   641       |> fold_rev (Term.abs o pair Name.uu) Ts;
   642 
   643 val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
   644  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
   645 
   646     val exclss' =
   647       disc_eqnss
   648       |> map (map (fn {fun_args, ctr_no, prems, ...} => (fun_args, ctr_no, prems))
   649         #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
   650         #> maps (uncurry (map o pair)
   651           #> map (fn ((fun_args, c, x), (_, c', y)) => ((c, c'), (x, mk_not (mk_conjs y)))
   652             ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
   653             ||> Logic.list_implies
   654             ||> curry Logic.list_all (map dest_Free fun_args))))
   655   in
   656     map (list_comb o rpair corec_args) corecs
   657     |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
   658     |> map2 currys arg_Tss
   659     |> Syntax.check_terms lthy
   660     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
   661     |> rpair exclss'
   662   end;
   663 
   664 fun mk_real_disc_eqns fun_binding arg_Ts {ctr_specs, ...} sel_eqns disc_eqns =
   665   if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
   666     let
   667       val n = 0 upto length ctr_specs
   668         |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
   669       val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
   670         |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
   671       val extra_disc_eqn = {
   672         fun_name = Binding.name_of fun_binding,
   673         fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
   674         fun_args = fun_args,
   675         ctr = #ctr (nth ctr_specs n),
   676         ctr_no = n,
   677         disc = #disc (nth ctr_specs n),
   678         prems = maps (invert_prems o #prems) disc_eqns,
   679         user_eqn = undef_const};
   680     in
   681       chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
   682     end;
   683 
   684 fun add_primcorec sequential fixes specs lthy =
   685   let
   686     val (bs, mxs) = map_split (apfst fst) fixes;
   687     val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
   688 
   689     (* copied from primrec_new *)
   690     fun get_indices t = map (fst #>> Binding.name_of #> Free) fixes
   691       |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   692       |> map_filter I;
   693 
   694     val callssss = []; (* FIXME *)
   695 
   696     val ((nontriv, corec_specs', _, coinduct_thm, strong_co_induct_thm, coinduct_thmss,
   697           strong_coinduct_thmss), lthy') =
   698       corec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
   699 
   700     val fun_names = map Binding.name_of bs;
   701     val corec_specs = take (length fun_names) corec_specs'; (*###*)
   702 
   703     val (eqns_data, _) =
   704       fold_map (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) []
   705       |>> flat;
   706 
   707     val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
   708       |> partition_eq ((op =) o pairself #fun_name)
   709       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   710       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
   711     val _ = disc_eqnss' |> map (fn x =>
   712       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
   713         primrec_error_eqns "excess discriminator equations in definition"
   714           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
   715 
   716     val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
   717       |> partition_eq ((op =) o pairself #fun_name)
   718       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   719       |> map (flat o snd);
   720 
   721     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   722     val arg_Tss = map (binder_types o snd o fst) fixes;
   723     val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
   724     val (defs, exclss') =
   725       co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
   726 
   727     (* try to prove (automatically generated) tautologies by ourselves *)
   728     val exclss'' = exclss'
   729       |> map (map (apsnd
   730         (`(try (fn t => Goal.prove lthy [] [] t (mk_primcorec_assumption_tac lthy |> K))))));
   731     val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
   732     val (obligation_idxss, obligationss) = exclss''
   733       |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
   734       |> split_list o map split_list;
   735 
   736     fun prove thmss' def_thms' lthy =
   737       let
   738         val def_thms = map (snd o snd) def_thms';
   739 
   740         val exclss' = map (op ~~) (obligation_idxss ~~ thmss');
   741         fun mk_exclsss excls n =
   742           (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
   743           |-> fold (fn ((c, c'), thm) => nth_map c (nth_map c' (K [thm])));
   744         val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
   745           |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
   746 
   747         fun prove_disc {ctr_specs, ...} exclsss
   748             {fun_name, fun_T, fun_args, ctr_no, prems, user_eqn, ...} =
   749           if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else
   750             let
   751               val {disc_corec, ...} = nth ctr_specs ctr_no;
   752               val k = 1 + ctr_no;
   753               val m = length prems;
   754               val t =
   755                 list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
   756                 |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
   757                 |> HOLogic.mk_Trueprop
   758                 |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   759                 |> curry Logic.list_all (map dest_Free fun_args);
   760             in
   761               mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
   762               |> K |> Goal.prove lthy [] [] t
   763               |> pair (#disc (nth ctr_specs ctr_no))
   764               |> single
   765             end;
   766 
   767         fun prove_sel {ctr_specs, nested_maps, nested_map_idents, nested_map_comps, ...}
   768             disc_eqns exclsss {fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} =
   769           let
   770             val (SOME ctr_spec) = find_first (equal ctr o #ctr) ctr_specs;
   771             val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
   772             val prems = the_default (maps (invert_prems o #prems) disc_eqns)
   773                 (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
   774             val sel_corec = find_index (equal sel) (#sels ctr_spec)
   775               |> nth (#sel_corecs ctr_spec);
   776             val k = 1 + ctr_no;
   777             val m = length prems;
   778             val t =
   779               list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
   780               |> curry betapply sel
   781               |> rpair (abstract (List.rev fun_args) rhs_term)
   782               |> HOLogic.mk_Trueprop o HOLogic.mk_eq
   783               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   784               |> curry Logic.list_all (map dest_Free fun_args);
   785           in
   786             mk_primcorec_ctr_or_sel_tac lthy def_thms sel_corec k m exclsss
   787               nested_maps nested_map_idents nested_map_comps
   788             |> K |> Goal.prove lthy [] [] t
   789             |> pair sel
   790           end;
   791 
   792         fun prove_ctr (_, disc_thms) (_, sel_thms') disc_eqns sel_eqns
   793             {ctr, disc, sels, collapse, ...} =
   794 let val _ = tracing ("disc = " ^ @{make_string} disc); in
   795           if not (exists (equal ctr o #ctr) disc_eqns)
   796               andalso not (exists (equal ctr o #ctr) sel_eqns)
   797 andalso (warning ("no eqns for ctr " ^ Syntax.string_of_term lthy ctr); true)
   798             orelse (* don't try to prove theorems when some sel_eqns are missing *)
   799               filter (equal ctr o #ctr) sel_eqns
   800               |> fst o finds ((op =) o apsnd #sel) sels
   801               |> exists (null o snd)
   802 andalso (warning ("sel_eqn(s) missing for ctr " ^ Syntax.string_of_term lthy ctr); true)
   803           then [] else
   804             let
   805 val _ = tracing ("ctr = " ^ Syntax.string_of_term lthy ctr);
   806 val _ = tracing (the_default "NO disc_eqn" (Option.map (curry (op ^) "disc = " o Syntax.string_of_term lthy o #disc) (find_first (equal ctr o #ctr) disc_eqns)));
   807               val (fun_name, fun_T, fun_args, prems) =
   808                 (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
   809                 |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x))
   810                 ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, []))
   811                 |> the o merge_options;
   812               val m = length prems;
   813               val t = sel_eqns
   814                 |> fst o finds ((op =) o apsnd #sel) sels
   815                 |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
   816                 |> curry list_comb ctr
   817                 |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
   818                   map Bound (length fun_args - 1 downto 0)))
   819                 |> HOLogic.mk_Trueprop
   820                 |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   821                 |> curry Logic.list_all (map dest_Free fun_args);
   822               val maybe_disc_thm = AList.lookup (op =) disc_thms disc;
   823               val sel_thms = map snd (filter (member (op =) sels o fst) sel_thms');
   824 val _ = tracing ("t = " ^ Syntax.string_of_term lthy t);
   825 val _ = tracing ("m = " ^ @{make_string} m);
   826 val _ = tracing ("collapse = " ^ @{make_string} collapse);
   827 val _ = tracing ("maybe_disc_thm = " ^ @{make_string} maybe_disc_thm);
   828 val _ = tracing ("sel_thms = " ^ @{make_string} sel_thms);
   829             in
   830               mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
   831               |> K |> Goal.prove lthy [] [] t
   832               |> single
   833 (*handle ERROR x => (warning x; []))*)
   834 end
   835           end;
   836 
   837         val (disc_notes, disc_thmss) =
   838           fun_names ~~ map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss
   839           |> `(map (fn (fun_name, thms) =>
   840             ((Binding.qualify true fun_name (@{binding disc}), simp_attrs), [(map snd thms, [])])));
   841         val (sel_notes, sel_thmss) =
   842           fun_names ~~ map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss
   843           |> `(map (fn (fun_name, thms) =>
   844             ((Binding.qualify true fun_name (@{binding sel}), simp_attrs), [(map snd thms, [])])));
   845         val ctr_notes =
   846           fun_names ~~ map5 (maps oooo prove_ctr) disc_thmss sel_thmss
   847             disc_eqnss sel_eqnss (map #ctr_specs corec_specs)
   848           |> map (fn (fun_name, thms) =>
   849             ((Binding.qualify true fun_name (@{binding ctr}), simp_attrs), [(thms, [])]));
   850       in
   851         lthy |> snd o Local_Theory.notes
   852           (filter (not o null o fst o the_single o snd) (disc_notes @ sel_notes @ ctr_notes))
   853       end;
   854   in
   855     lthy'
   856     |> Proof.theorem NONE (curry (op #->) (fold_map Local_Theory.define defs) o prove) obligationss
   857     |> Proof.refine (Method.primitive_text I)
   858     |> Seq.hd
   859   end
   860 
   861 fun add_primcorec_cmd seq (raw_fixes, raw_specs) lthy =
   862   let
   863     val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy);
   864   in
   865     add_primcorec seq fixes specs lthy
   866     handle ERROR str => primrec_error str
   867   end
   868   handle Primrec_Error (str, eqns) =>
   869     if null eqns
   870     then error ("primcorec error:\n  " ^ str)
   871     else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
   872       space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns))
   873 
   874 end;