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