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