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