src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
author blanchet
Fri, 30 Aug 2013 11:37:22 +0200
changeset 54441 cfef783090eb
parent 54440 ae49b835ca01
child 54447 8af01463b2d3
permissions -rw-r--r--
polished newly included files after moving
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@54440
    12
end;
blanchet@54440
    13
blanchet@54440
    14
(* TODO:
blanchet@54440
    15
     - error handling for indirect calls
blanchet@54440
    16
*)
blanchet@54440
    17
blanchet@54440
    18
structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
blanchet@54440
    19
struct
blanchet@54440
    20
blanchet@54440
    21
open BNF_Util
blanchet@54440
    22
open BNF_FP_Util
blanchet@54440
    23
open BNF_FP_Rec_Sugar_Util
blanchet@54440
    24
open BNF_FP_Rec_Sugar_Tactics
blanchet@54440
    25
blanchet@54440
    26
exception Primrec_Error of string * term list;
blanchet@54440
    27
blanchet@54440
    28
fun primrec_error str = raise Primrec_Error (str, []);
blanchet@54440
    29
fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
blanchet@54440
    30
fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
blanchet@54440
    31
blanchet@54440
    32
fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
blanchet@54440
    33
blanchet@54440
    34
val simp_attrs = @{attributes [simp]};
blanchet@54440
    35
blanchet@54440
    36
blanchet@54440
    37
type eqn_data = {
blanchet@54440
    38
  fun_name: string,
blanchet@54440
    39
  rec_type: typ,
blanchet@54440
    40
  ctr: term,
blanchet@54440
    41
  ctr_args: term list,
blanchet@54440
    42
  left_args: term list,
blanchet@54440
    43
  right_args: term list,
blanchet@54440
    44
  res_type: typ,
blanchet@54440
    45
  rhs_term: term,
blanchet@54440
    46
  user_eqn: term
blanchet@54440
    47
};
blanchet@54440
    48
blanchet@54440
    49
fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
blanchet@54441
    50
  |> fold (K (fn u => Abs (Name.uu, dummyT, u))) (0 upto n);
blanchet@54440
    51
blanchet@54440
    52
fun dissect_eqn lthy fun_names eqn' =
blanchet@54440
    53
  let
blanchet@54440
    54
    val eqn = subst_bounds (strip_qnt_vars @{const_name all} eqn' |> map Free |> rev,
blanchet@54440
    55
        strip_qnt_body @{const_name all} eqn') |> HOLogic.dest_Trueprop
blanchet@54440
    56
        handle TERM _ =>
blanchet@54440
    57
          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
blanchet@54440
    58
    val (lhs, rhs) = HOLogic.dest_eq eqn
blanchet@54440
    59
        handle TERM _ =>
blanchet@54440
    60
          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
blanchet@54440
    61
    val (fun_name, args) = strip_comb lhs
blanchet@54440
    62
      |>> (fn x => if is_Free x then fst (dest_Free x)
blanchet@54440
    63
          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
blanchet@54440
    64
    val (left_args, rest) = take_prefix is_Free args;
blanchet@54440
    65
    val (nonfrees, right_args) = take_suffix is_Free rest;
blanchet@54440
    66
    val _ = length nonfrees = 1 orelse if length nonfrees = 0 then
blanchet@54440
    67
      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
blanchet@54440
    68
      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
blanchet@54440
    69
    val _ = member (op =) fun_names fun_name orelse
blanchet@54440
    70
      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
blanchet@54440
    71
blanchet@54440
    72
    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
blanchet@54440
    73
    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
blanchet@54440
    74
      primrec_error_eqn "partially applied constructor in pattern" eqn;
blanchet@54440
    75
    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
blanchet@54440
    76
      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
blanchet@54440
    77
        "\" in left-hand side") eqn end;
blanchet@54440
    78
    val _ = forall is_Free ctr_args orelse
blanchet@54440
    79
      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
blanchet@54440
    80
    val _ =
blanchet@54440
    81
      let val b = fold_aterms (fn x as Free (v, _) =>
blanchet@54440
    82
        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
blanchet@54440
    83
        not (member (op =) fun_names v) andalso
blanchet@54440
    84
        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
blanchet@54440
    85
      in
blanchet@54440
    86
        null b orelse
blanchet@54440
    87
        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
blanchet@54440
    88
          commas (map (Syntax.string_of_term lthy) b)) eqn
blanchet@54440
    89
      end;
blanchet@54440
    90
  in
blanchet@54440
    91
    {fun_name = fun_name,
blanchet@54440
    92
     rec_type = body_type (type_of ctr),
blanchet@54440
    93
     ctr = ctr,
blanchet@54440
    94
     ctr_args = ctr_args,
blanchet@54440
    95
     left_args = left_args,
blanchet@54440
    96
     right_args = right_args,
blanchet@54440
    97
     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
blanchet@54440
    98
     rhs_term = rhs,
blanchet@54440
    99
     user_eqn = eqn'}
blanchet@54440
   100
  end;
blanchet@54440
   101
blanchet@54440
   102
(* substitutes (f ls x rs) by (y ls rs) for all f: get_idx f \<ge> 0, (x,y) \<in> substs *)
blanchet@54440
   103
fun subst_direct_calls get_idx get_ctr_pos substs = 
blanchet@54440
   104
  let
blanchet@54440
   105
    fun subst (Abs (v, T, b)) = Abs (v, T, subst b)
blanchet@54440
   106
      | subst t =
blanchet@54440
   107
        let
blanchet@54440
   108
          val (f, args) = strip_comb t;
blanchet@54440
   109
          val idx = get_idx f;
blanchet@54440
   110
          val ctr_pos  = if idx >= 0 then get_ctr_pos idx else ~1;
blanchet@54440
   111
        in
blanchet@54440
   112
          if idx < 0 then
blanchet@54440
   113
            list_comb (f, map subst args)
blanchet@54440
   114
          else if ctr_pos >= length args then
blanchet@54440
   115
            primrec_error_eqn "too few arguments in recursive call" t
blanchet@54440
   116
          else
blanchet@54440
   117
            let
blanchet@54440
   118
              val (key, repl) = the (find_first (equal (nth args ctr_pos) o fst) substs)
blanchet@54440
   119
                handle Option.Option => primrec_error_eqn
blanchet@54440
   120
                  "recursive call not directly applied to constructor argument" t;
blanchet@54440
   121
            in
blanchet@54440
   122
              remove (op =) key args |> map subst |> curry list_comb repl
blanchet@54440
   123
            end
blanchet@54440
   124
        end
blanchet@54440
   125
  in subst end;
blanchet@54440
   126
blanchet@54440
   127
(* FIXME get rid of funs_data or get_indices *)
blanchet@54440
   128
fun rewrite_map_arg funs_data get_indices y rec_type res_type =
blanchet@54440
   129
  let
blanchet@54440
   130
    val pT = HOLogic.mk_prodT (rec_type, res_type);
blanchet@54440
   131
    val fstx = fst_const pT;
blanchet@54440
   132
    val sndx = snd_const pT;
blanchet@54440
   133
blanchet@54440
   134
    val SOME ({fun_name, left_args, ...} :: _) =
blanchet@54440
   135
      find_first (equal rec_type o #rec_type o hd) funs_data;
blanchet@54440
   136
    val ctr_pos = length left_args;
blanchet@54440
   137
blanchet@54440
   138
    fun subst _ d (t as Bound d') = t |> d = d' ? curry (op $) fstx
blanchet@54440
   139
      | subst l d (Abs (v, T, b)) = Abs (v, if d < 0 then pT else T, subst l (d + 1) b)
blanchet@54440
   140
      | subst l d t =
blanchet@54440
   141
        let val (u, vs) = strip_comb t in
blanchet@54440
   142
          if try (fst o dest_Free) u = SOME fun_name then
blanchet@54440
   143
            if l andalso length vs = ctr_pos then
blanchet@54440
   144
              list_comb (sndx |> permute_args ctr_pos, vs)
blanchet@54440
   145
            else if length vs <= ctr_pos then
blanchet@54440
   146
              primrec_error_eqn "too few arguments in recursive call" t
blanchet@54440
   147
            else if nth vs ctr_pos |> member (op =) [y, Bound d] then
blanchet@54440
   148
              list_comb (sndx $ nth vs ctr_pos, nth_drop ctr_pos vs |> map (subst false d))
blanchet@54440
   149
            else
blanchet@54440
   150
              primrec_error_eqn "recursive call not directly applied to constructor argument" t
blanchet@54440
   151
          else if try (fst o dest_Const) u = SOME @{const_name comp} then
blanchet@54440
   152
            (hd vs |> get_indices |> null orelse
blanchet@54440
   153
              primrec_error_eqn "recursive call not directly applied to constructor argument" t;
blanchet@54440
   154
            list_comb
blanchet@54440
   155
              (u |> map_types (strip_type #>> (fn Ts => Ts
blanchet@54440
   156
                   |> nth_map (length Ts - 1) (K pT)
blanchet@54440
   157
                   |> nth_map (length Ts - 2) (strip_type #>> nth_map 0 (K pT) #> (op --->)))
blanchet@54440
   158
                 #> (op --->)),
blanchet@54440
   159
              nth_map 1 (subst l d) vs))
blanchet@54440
   160
          else
blanchet@54440
   161
            list_comb (u, map (subst false d) vs)
blanchet@54440
   162
        end
blanchet@54440
   163
  in
blanchet@54440
   164
    subst true ~1
blanchet@54440
   165
  end;
blanchet@54440
   166
blanchet@54440
   167
(* FIXME get rid of funs_data or get_indices *)
blanchet@54440
   168
fun subst_indirect_call lthy funs_data get_indices (y, y') =
blanchet@54440
   169
  let
blanchet@54440
   170
    fun massage massage_map_arg bound_Ts =
blanchet@54440
   171
      massage_indirect_rec_call lthy (not o null o get_indices) massage_map_arg bound_Ts y y';
blanchet@54440
   172
    fun subst bound_Ts (t as _ $ _) =
blanchet@54440
   173
        let
blanchet@54440
   174
          val (f', args') = strip_comb t;
blanchet@54440
   175
          val fun_arg_idx = find_index (exists_subterm (not o null o get_indices)) args';
blanchet@54440
   176
          val arg_idx = find_index (exists_subterm (equal y)) args';
blanchet@54440
   177
          val (f, args) = chop (arg_idx + 1) args' |>> curry list_comb f';
blanchet@54440
   178
          val _ = fun_arg_idx < 0 orelse arg_idx >= 0 orelse
blanchet@54440
   179
            primrec_error_eqn "recursive call not applied to constructor argument" t;
blanchet@54440
   180
        in
blanchet@54440
   181
          if fun_arg_idx <> arg_idx andalso fun_arg_idx >= 0 then
blanchet@54440
   182
            if nth args' arg_idx = y then
blanchet@54440
   183
              list_comb (massage (rewrite_map_arg funs_data get_indices y) bound_Ts f, args)
blanchet@54440
   184
            else
blanchet@54440
   185
              primrec_error_eqn "recursive call not directly applied to constructor argument" f
blanchet@54440
   186
          else
blanchet@54440
   187
            list_comb (f', map (subst bound_Ts) args')
blanchet@54440
   188
        end
blanchet@54440
   189
      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
blanchet@54440
   190
      | subst bound_Ts t = t |> t = y ? massage (K I |> K) bound_Ts;
blanchet@54440
   191
  in subst [] end;
blanchet@54440
   192
blanchet@54440
   193
fun build_rec_arg lthy get_indices funs_data ctr_spec maybe_eqn_data =
blanchet@54440
   194
  if is_some maybe_eqn_data then
blanchet@54440
   195
    let
blanchet@54440
   196
      val eqn_data = the maybe_eqn_data;
blanchet@54440
   197
      val t = #rhs_term eqn_data;
blanchet@54440
   198
      val ctr_args = #ctr_args eqn_data;
blanchet@54440
   199
blanchet@54440
   200
      val calls = #calls ctr_spec;
blanchet@54440
   201
      val n_args = fold (curry (op +) o (fn Direct_Rec _ => 2 | _ => 1)) calls 0;
blanchet@54440
   202
blanchet@54440
   203
      val no_calls' = tag_list 0 calls
blanchet@54440
   204
        |> map_filter (try (apsnd (fn No_Rec n => n | Direct_Rec (n, _) => n)));
blanchet@54440
   205
      val direct_calls' = tag_list 0 calls
blanchet@54440
   206
        |> map_filter (try (apsnd (fn Direct_Rec (_, n) => n)));
blanchet@54440
   207
      val indirect_calls' = tag_list 0 calls
blanchet@54440
   208
        |> map_filter (try (apsnd (fn Indirect_Rec n => n)));
blanchet@54440
   209
blanchet@54440
   210
      fun make_direct_type T = dummyT; (* FIXME? *)
blanchet@54440
   211
blanchet@54440
   212
      val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data;
blanchet@54440
   213
blanchet@54440
   214
      fun make_indirect_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T =>
blanchet@54440
   215
        let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in
blanchet@54440
   216
          if is_some maybe_res_type
blanchet@54440
   217
          then HOLogic.mk_prodT (T, the maybe_res_type)
blanchet@54440
   218
          else make_indirect_type T end))
blanchet@54440
   219
        | make_indirect_type T = T;
blanchet@54440
   220
blanchet@54440
   221
      val args = replicate n_args ("", dummyT)
blanchet@54440
   222
        |> Term.rename_wrt_term t
blanchet@54440
   223
        |> map Free
blanchet@54440
   224
        |> fold (fn (ctr_arg_idx, arg_idx) =>
blanchet@54440
   225
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
blanchet@54440
   226
          no_calls'
blanchet@54440
   227
        |> fold (fn (ctr_arg_idx, arg_idx) =>
blanchet@54440
   228
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_direct_type)))
blanchet@54440
   229
          direct_calls'
blanchet@54440
   230
        |> fold (fn (ctr_arg_idx, arg_idx) =>
blanchet@54440
   231
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_indirect_type)))
blanchet@54440
   232
          indirect_calls';
blanchet@54440
   233
blanchet@54440
   234
      val direct_calls = map (apfst (nth ctr_args) o apsnd (nth args)) direct_calls';
blanchet@54440
   235
      val indirect_calls = map (apfst (nth ctr_args) o apsnd (nth args)) indirect_calls';
blanchet@54440
   236
blanchet@54440
   237
      val get_idx = (fn Free (v, _) => find_index (equal v o #fun_name o hd) funs_data | _ => ~1);
blanchet@54440
   238
blanchet@54440
   239
      val t' = t
blanchet@54440
   240
        |> fold (subst_indirect_call lthy funs_data get_indices) indirect_calls
blanchet@54440
   241
        |> subst_direct_calls get_idx (length o #left_args o hd o nth funs_data) direct_calls;
blanchet@54440
   242
blanchet@54440
   243
      val abstractions = map dest_Free (args @ #left_args eqn_data @ #right_args eqn_data);
blanchet@54440
   244
    in
blanchet@54440
   245
      t' |> fold_rev absfree abstractions
blanchet@54440
   246
    end
blanchet@54440
   247
  else Const (@{const_name undefined}, dummyT)
blanchet@54440
   248
blanchet@54440
   249
fun build_defs lthy bs funs_data rec_specs get_indices =
blanchet@54440
   250
  let
blanchet@54440
   251
    val n_funs = length funs_data;
blanchet@54440
   252
blanchet@54440
   253
    val ctr_spec_eqn_data_list' =
blanchet@54440
   254
      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
blanchet@54440
   255
      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
blanchet@54440
   256
          ##> (fn x => null x orelse
blanchet@54440
   257
            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
blanchet@54440
   258
    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
blanchet@54440
   259
      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
blanchet@54440
   260
blanchet@54440
   261
    val ctr_spec_eqn_data_list =
blanchet@54440
   262
      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
blanchet@54440
   263
blanchet@54440
   264
    val recs = take n_funs rec_specs |> map #recx;
blanchet@54440
   265
    val rec_args = ctr_spec_eqn_data_list
blanchet@54440
   266
      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
blanchet@54440
   267
      |> map (uncurry (build_rec_arg lthy get_indices funs_data) o apsnd (try the_single));
blanchet@54440
   268
    val ctr_poss = map (fn x =>
blanchet@54440
   269
      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
blanchet@54440
   270
        primrec_error ("inconstant constructor pattern position for function " ^
blanchet@54440
   271
          quote (#fun_name (hd x)))
blanchet@54440
   272
      else
blanchet@54440
   273
        hd x |> #left_args |> length) funs_data;
blanchet@54440
   274
  in
blanchet@54440
   275
    (recs, ctr_poss)
blanchet@54440
   276
    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
blanchet@54440
   277
    |> Syntax.check_terms lthy
blanchet@54440
   278
    |> map2 (fn b => fn t => ((b, NoSyn), ((Binding.map_name Thm.def_name b, []), t))) bs
blanchet@54440
   279
  end;
blanchet@54440
   280
blanchet@54440
   281
fun find_rec_calls get_indices eqn_data =
blanchet@54440
   282
  let
blanchet@54440
   283
    fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg
blanchet@54440
   284
      | find (t as _ $ _) ctr_arg =
blanchet@54440
   285
        let
blanchet@54440
   286
          val (f', args') = strip_comb t;
blanchet@54440
   287
          val n = find_index (equal ctr_arg) args';
blanchet@54440
   288
        in
blanchet@54440
   289
          if n < 0 then
blanchet@54440
   290
            find f' ctr_arg @ maps (fn x => find x ctr_arg) args'
blanchet@54440
   291
          else
blanchet@54440
   292
            let val (f, args) = chop n args' |>> curry list_comb f' in
blanchet@54440
   293
              if exists_subterm (not o null o get_indices) f then
blanchet@54440
   294
                f :: maps (fn x => find x ctr_arg) args
blanchet@54440
   295
              else
blanchet@54440
   296
                find f ctr_arg @ maps (fn x => find x ctr_arg) args
blanchet@54440
   297
            end
blanchet@54440
   298
        end
blanchet@54440
   299
      | find _ _ = [];
blanchet@54440
   300
  in
blanchet@54440
   301
    map (find (#rhs_term eqn_data)) (#ctr_args eqn_data)
blanchet@54440
   302
    |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
blanchet@54440
   303
  end;
blanchet@54440
   304
blanchet@54440
   305
fun add_primrec fixes specs lthy =
blanchet@54440
   306
  let
blanchet@54440
   307
    val bs = map (fst o fst) fixes;
blanchet@54440
   308
    val fun_names = map Binding.name_of bs;
blanchet@54440
   309
    val eqns_data = map (snd #> dissect_eqn lthy fun_names) specs;
blanchet@54440
   310
    val funs_data = eqns_data
blanchet@54440
   311
      |> partition_eq ((op =) o pairself #fun_name)
blanchet@54440
   312
      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
blanchet@54440
   313
      |> map (fn (x, y) => the_single y handle List.Empty =>
blanchet@54440
   314
          primrec_error ("missing equations for function " ^ quote x));
blanchet@54440
   315
blanchet@54440
   316
    fun get_indices t = map (fst #>> Binding.name_of #> Free) fixes
blanchet@54440
   317
      |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
blanchet@54440
   318
      |> map_filter I;
blanchet@54440
   319
blanchet@54440
   320
    val arg_Ts = map (#rec_type o hd) funs_data;
blanchet@54440
   321
    val res_Ts = map (#res_type o hd) funs_data;
blanchet@54440
   322
    val callssss = funs_data
blanchet@54440
   323
      |> map (partition_eq ((op =) o pairself #ctr))
blanchet@54440
   324
      |> map (maps (map_filter (find_rec_calls get_indices)));
blanchet@54440
   325
blanchet@54440
   326
    val ((nontriv, rec_specs, _, induct_thm, induct_thms), lthy') =
blanchet@54440
   327
      rec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
blanchet@54440
   328
blanchet@54440
   329
    val actual_nn = length funs_data;
blanchet@54440
   330
blanchet@54440
   331
    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
blanchet@54440
   332
      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
blanchet@54440
   333
        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
blanchet@54440
   334
          " is not a constructor in left-hand side") user_eqn) eqns_data end;
blanchet@54440
   335
blanchet@54440
   336
    val defs = build_defs lthy' bs funs_data rec_specs get_indices;
blanchet@54440
   337
blanchet@54440
   338
    fun prove def_thms' {ctr_specs, nested_map_id's, nested_map_comps, ...} induct_thm fun_data
blanchet@54440
   339
        lthy =
blanchet@54440
   340
      let
blanchet@54440
   341
        val fun_name = #fun_name (hd fun_data);
blanchet@54440
   342
        val def_thms = map (snd o snd) def_thms';
blanchet@54440
   343
        val simp_thms = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
blanchet@54440
   344
          |> fst
blanchet@54440
   345
          |> map_filter (try (fn (x, [y]) =>
blanchet@54440
   346
            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
blanchet@54440
   347
          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
blanchet@54440
   348
            mk_primrec_tac lthy num_extra_args nested_map_id's nested_map_comps def_thms rec_thm
blanchet@54440
   349
            |> K |> Goal.prove lthy [] [] user_eqn)
blanchet@54440
   350
blanchet@54440
   351
        val notes =
blanchet@54440
   352
          [(inductN, if actual_nn > 1 then [induct_thm] else [], []),
blanchet@54440
   353
           (simpsN, simp_thms, simp_attrs)]
blanchet@54440
   354
          |> filter_out (null o #2)
blanchet@54440
   355
          |> map (fn (thmN, thms, attrs) =>
blanchet@54440
   356
            ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]));
blanchet@54440
   357
      in
blanchet@54440
   358
        lthy |> Local_Theory.notes notes
blanchet@54440
   359
      end;
blanchet@54440
   360
blanchet@54440
   361
    val common_name = mk_common_name fun_names;
blanchet@54440
   362
blanchet@54440
   363
    val common_notes =
blanchet@54440
   364
      [(inductN, if nontriv andalso actual_nn > 1 then [induct_thm] else [], [])]
blanchet@54440
   365
      |> filter_out (null o #2)
blanchet@54440
   366
      |> map (fn (thmN, thms, attrs) =>
blanchet@54440
   367
        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
blanchet@54440
   368
  in
blanchet@54440
   369
    lthy'
blanchet@54440
   370
    |> fold_map Local_Theory.define defs
blanchet@54440
   371
    |-> (fn def_thms' => fold_map3 (prove def_thms') (take actual_nn rec_specs)
blanchet@54440
   372
      (take actual_nn induct_thms) funs_data)
blanchet@54440
   373
    |> snd
blanchet@54440
   374
    |> Local_Theory.notes common_notes |> snd
blanchet@54440
   375
  end;
blanchet@54440
   376
blanchet@54440
   377
fun add_primrec_cmd raw_fixes raw_specs lthy =
blanchet@54440
   378
  let
blanchet@54440
   379
    val _ = let val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes) in null d orelse
blanchet@54440
   380
      primrec_error ("duplicate function name(s): " ^ commas d) end;
blanchet@54440
   381
    val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy);
blanchet@54440
   382
  in
blanchet@54440
   383
    add_primrec fixes specs lthy
blanchet@54440
   384
      handle ERROR str => primrec_error str
blanchet@54440
   385
  end
blanchet@54440
   386
  handle Primrec_Error (str, eqns) =>
blanchet@54440
   387
    if null eqns
blanchet@54440
   388
    then error ("primrec_new error:\n  " ^ str)
blanchet@54440
   389
    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
blanchet@54440
   390
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns))
blanchet@54440
   391
blanchet@54440
   392
blanchet@54440
   393
val _ = Outer_Syntax.local_theory
blanchet@54440
   394
  @{command_spec "primrec_new"}
blanchet@54440
   395
  "define primitive recursive functions"
blanchet@54440
   396
  (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_primrec_cmd);
blanchet@54440
   397
blanchet@54440
   398
blanchet@54440
   399
blanchet@54440
   400
blanchet@54440
   401
blanchet@54440
   402
blanchet@54440
   403
blanchet@54440
   404
blanchet@54440
   405
blanchet@54440
   406
blanchet@54440
   407
blanchet@54440
   408
blanchet@54440
   409
blanchet@54440
   410
blanchet@54440
   411
blanchet@54440
   412
blanchet@54440
   413
type co_eqn_data_dtr_disc = {
blanchet@54440
   414
  fun_name: string,
blanchet@54440
   415
  ctr_no: int,
blanchet@54440
   416
  cond: term,
blanchet@54440
   417
  user_eqn: term
blanchet@54440
   418
};
blanchet@54440
   419
type co_eqn_data_dtr_sel = {
blanchet@54440
   420
  fun_name: string,
blanchet@54440
   421
  ctr_no: int,
blanchet@54440
   422
  sel_no: int,
blanchet@54440
   423
  fun_args: term list,
blanchet@54440
   424
  rhs_term: term,
blanchet@54440
   425
  user_eqn: term
blanchet@54440
   426
};
blanchet@54440
   427
datatype co_eqn_data =
blanchet@54440
   428
  Dtr_Disc of co_eqn_data_dtr_disc |
blanchet@54440
   429
  Dtr_Sel of co_eqn_data_dtr_sel
blanchet@54440
   430
blanchet@54440
   431
fun co_dissect_eqn_disc sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps =
blanchet@54440
   432
  let
blanchet@54440
   433
    fun find_subterm p = let (* FIXME \<exists>? *)
blanchet@54440
   434
      fun f (t as u $ v) =
blanchet@54440
   435
        fold_rev (curry merge_options) [if p t then SOME t else NONE, f u, f v] NONE
blanchet@54440
   436
        | f t = if p t then SOME t else NONE
blanchet@54440
   437
      in f end;
blanchet@54440
   438
blanchet@54440
   439
    val fun_name = imp_rhs
blanchet@54440
   440
      |> perhaps (try HOLogic.dest_not)
blanchet@54440
   441
      |> `(try (fst o dest_Free o head_of o snd o dest_comb))
blanchet@54440
   442
      ||> (try (fst o dest_Free o head_of o fst o HOLogic.dest_eq))
blanchet@54440
   443
      |> the o merge_options;
blanchet@54440
   444
    val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name)
blanchet@54440
   445
      handle Option.Option => primrec_error_eqn "malformed discriminator equation" imp_rhs;
blanchet@54440
   446
blanchet@54440
   447
    val discs = #ctr_specs corec_spec |> map #disc;
blanchet@54440
   448
    val ctrs = #ctr_specs corec_spec |> map #ctr;
blanchet@54440
   449
    val n_ctrs = length ctrs;
blanchet@54440
   450
    val not_disc = head_of imp_rhs = @{term Not};
blanchet@54440
   451
    val _ = not_disc andalso n_ctrs <> 2 andalso
blanchet@54440
   452
      primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" imp_rhs;
blanchet@54440
   453
    val disc = find_subterm (member (op =) discs o head_of) imp_rhs;
blanchet@54440
   454
    val eq_ctr0 = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
blanchet@54440
   455
        |> (fn SOME t => let val n = find_index (equal t) ctrs in
blanchet@54440
   456
          if n >= 0 then SOME n else NONE end | _ => NONE);
blanchet@54440
   457
    val _ = is_some disc orelse is_some eq_ctr0 orelse
blanchet@54440
   458
      primrec_error_eqn "no discriminator in equation" imp_rhs;
blanchet@54440
   459
    val ctr_no' =
blanchet@54440
   460
      if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
blanchet@54440
   461
    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
blanchet@54440
   462
    val fun_args = if is_none disc
blanchet@54440
   463
      then imp_rhs |> perhaps (try HOLogic.dest_not) |> HOLogic.dest_eq |> fst |> strip_comb |> snd
blanchet@54440
   464
      else the disc |> the_single o snd o strip_comb
blanchet@54440
   465
        |> (fn t => if try (fst o dest_Free o head_of) t = SOME fun_name
blanchet@54440
   466
          then snd (strip_comb t) else []);
blanchet@54440
   467
blanchet@54440
   468
    val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
blanchet@54440
   469
    val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
blanchet@54440
   470
    val catch_all = try (fst o dest_Free o the_single) imp_lhs' = SOME Name.uu_;
blanchet@54440
   471
    val matched_conds = filter (equal fun_name o fst) matched_conds_ps |> map snd;
blanchet@54440
   472
    val imp_lhs = mk_conjs imp_lhs';
blanchet@54440
   473
    val cond =
blanchet@54440
   474
      if catch_all then
blanchet@54440
   475
        if null matched_conds then fold_rev absfree (map dest_Free fun_args) @{const True} else
blanchet@54440
   476
          (strip_abs_vars (hd matched_conds),
blanchet@54440
   477
            mk_disjs (map strip_abs_body matched_conds) |> HOLogic.mk_not)
blanchet@54440
   478
          |-> fold_rev (fn (v, T) => fn u => Abs (v, T, u))
blanchet@54440
   479
      else if sequential then
blanchet@54440
   480
        HOLogic.mk_conj (HOLogic.mk_not (mk_disjs (map strip_abs_body matched_conds)), imp_lhs)
blanchet@54440
   481
        |> fold_rev absfree (map dest_Free fun_args)
blanchet@54440
   482
      else
blanchet@54440
   483
        imp_lhs |> fold_rev absfree (map dest_Free fun_args);
blanchet@54440
   484
    val matched_cond =
blanchet@54440
   485
      if sequential then fold_rev absfree (map dest_Free fun_args) imp_lhs else cond;
blanchet@54440
   486
blanchet@54440
   487
    val matched_conds_ps' = if catch_all
blanchet@54440
   488
      then (fun_name, cond) :: filter (not_equal fun_name o fst) matched_conds_ps
blanchet@54440
   489
      else (fun_name, matched_cond) :: matched_conds_ps;
blanchet@54440
   490
  in
blanchet@54440
   491
    (Dtr_Disc {
blanchet@54440
   492
      fun_name = fun_name,
blanchet@54440
   493
      ctr_no = ctr_no,
blanchet@54440
   494
      cond = cond,
blanchet@54440
   495
      user_eqn = eqn'
blanchet@54440
   496
    }, matched_conds_ps')
blanchet@54440
   497
  end;
blanchet@54440
   498
blanchet@54440
   499
fun co_dissect_eqn_sel fun_name_corec_spec_list eqn' eqn =
blanchet@54440
   500
  let
blanchet@54440
   501
    val (lhs, rhs) = HOLogic.dest_eq eqn
blanchet@54440
   502
      handle TERM _ =>
blanchet@54440
   503
        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
blanchet@54440
   504
    val sel = head_of lhs;
blanchet@54440
   505
    val (fun_name, fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst (fst o dest_Free)
blanchet@54440
   506
      handle TERM _ =>
blanchet@54440
   507
        primrec_error_eqn "malformed selector argument in left-hand side" eqn;
blanchet@54440
   508
    val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name)
blanchet@54440
   509
      handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
blanchet@54440
   510
    val ((ctr_spec, ctr_no), sel) = #ctr_specs corec_spec
blanchet@54440
   511
      |> the o get_index (try (the o find_first (equal sel) o #sels))
blanchet@54440
   512
      |>> `(nth (#ctr_specs corec_spec));
blanchet@54440
   513
    val sel_no = find_index (equal sel) (#sels ctr_spec);
blanchet@54440
   514
  in
blanchet@54440
   515
    Dtr_Sel {
blanchet@54440
   516
      fun_name = fun_name,
blanchet@54440
   517
      ctr_no = ctr_no,
blanchet@54440
   518
      sel_no = sel_no,
blanchet@54440
   519
      fun_args = fun_args,
blanchet@54440
   520
      rhs_term = rhs,
blanchet@54440
   521
      user_eqn = eqn'
blanchet@54440
   522
    }
blanchet@54440
   523
  end;
blanchet@54440
   524
blanchet@54440
   525
fun co_dissect_eqn_ctr sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps =
blanchet@54440
   526
  let 
blanchet@54440
   527
    val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
blanchet@54440
   528
    val fun_name = head_of lhs |> fst o dest_Free;
blanchet@54440
   529
    val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name);
blanchet@54440
   530
    val (ctr, ctr_args) = strip_comb rhs;
blanchet@54440
   531
    val ctr_spec = the (find_first (equal ctr o #ctr) (#ctr_specs corec_spec))
blanchet@54440
   532
      handle Option.Option => primrec_error_eqn "not a constructor" ctr;
blanchet@54440
   533
    val disc_imp_rhs = betapply (#disc ctr_spec, lhs);
blanchet@54440
   534
    val (eqn_data_disc, matched_conds_ps') = co_dissect_eqn_disc
blanchet@54440
   535
        sequential fun_name_corec_spec_list eqn' imp_lhs' disc_imp_rhs matched_conds_ps;
blanchet@54440
   536
blanchet@54440
   537
    val sel_imp_rhss = (#sels ctr_spec ~~ ctr_args)
blanchet@54440
   538
      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
blanchet@54440
   539
blanchet@54440
   540
val _ = warning ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
blanchet@54440
   541
 Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> " ^
blanchet@54440
   542
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
blanchet@54440
   543
blanchet@54440
   544
    val eqns_data_sel =
blanchet@54440
   545
      map (co_dissect_eqn_sel fun_name_corec_spec_list @{const True}(*###*)) sel_imp_rhss;
blanchet@54440
   546
  in
blanchet@54440
   547
    (eqn_data_disc :: eqns_data_sel, matched_conds_ps')
blanchet@54440
   548
  end;
blanchet@54440
   549
blanchet@54440
   550
fun co_dissect_eqn sequential fun_name_corec_spec_list eqn' matched_conds_ps =
blanchet@54440
   551
  let
blanchet@54440
   552
    val eqn = subst_bounds (strip_qnt_vars @{const_name all} eqn' |> map Free |> rev,
blanchet@54440
   553
        strip_qnt_body @{const_name all} eqn')
blanchet@54440
   554
        handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
blanchet@54440
   555
    val (imp_lhs', imp_rhs) =
blanchet@54440
   556
      (map HOLogic.dest_Trueprop (Logic.strip_imp_prems eqn),
blanchet@54440
   557
       HOLogic.dest_Trueprop (Logic.strip_imp_concl eqn));
blanchet@54440
   558
blanchet@54440
   559
    val head = imp_rhs
blanchet@54440
   560
      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
blanchet@54440
   561
      |> head_of;
blanchet@54440
   562
blanchet@54440
   563
    val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
blanchet@54440
   564
blanchet@54440
   565
    val fun_names = map fst fun_name_corec_spec_list;
blanchet@54440
   566
    val discs = maps (#ctr_specs o snd) fun_name_corec_spec_list |> map #disc;
blanchet@54440
   567
    val sels = maps (#ctr_specs o snd) fun_name_corec_spec_list |> maps #sels;
blanchet@54440
   568
    val ctrs = maps (#ctr_specs o snd) fun_name_corec_spec_list |> map #ctr;
blanchet@54440
   569
  in
blanchet@54440
   570
    if member (op =) discs head orelse
blanchet@54440
   571
      is_some maybe_rhs andalso
blanchet@54440
   572
        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
blanchet@54440
   573
      co_dissect_eqn_disc sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps
blanchet@54440
   574
      |>> single
blanchet@54440
   575
    else if member (op =) sels head then
blanchet@54440
   576
      ([co_dissect_eqn_sel fun_name_corec_spec_list eqn' imp_rhs], matched_conds_ps)
blanchet@54440
   577
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
blanchet@54440
   578
      co_dissect_eqn_ctr sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds_ps
blanchet@54440
   579
    else
blanchet@54440
   580
      primrec_error_eqn "malformed function equation" eqn
blanchet@54440
   581
  end;
blanchet@54440
   582
blanchet@54440
   583
fun build_corec_args_discs ctr_specs disc_eqns =
blanchet@54440
   584
  let
blanchet@54440
   585
    val conds = map #cond disc_eqns;
blanchet@54440
   586
    val args =
blanchet@54440
   587
      if length ctr_specs = 1 then []
blanchet@54440
   588
      else if length disc_eqns = length ctr_specs then
blanchet@54440
   589
        fst (split_last conds)
blanchet@54440
   590
      else if length disc_eqns = length ctr_specs - 1 then
blanchet@54440
   591
        let val n = 0 upto length ctr_specs - 1
blanchet@54440
   592
            |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns)) (*###*) in
blanchet@54440
   593
          if n = length ctr_specs - 1 then
blanchet@54440
   594
            conds
blanchet@54440
   595
          else
blanchet@54440
   596
            split_last conds
blanchet@54440
   597
            ||> (fn t => fold_rev absfree (strip_abs_vars t) (strip_abs_body t |> HOLogic.mk_not))
blanchet@54440
   598
            |>> chop n
blanchet@54440
   599
            |> (fn ((l, r), x) => l @ (x :: r))
blanchet@54440
   600
        end
blanchet@54440
   601
      else
blanchet@54440
   602
        0 upto length ctr_specs - 1
blanchet@54440
   603
        |> map (fn idx => find_first (equal idx o #ctr_no) disc_eqns
blanchet@54440
   604
          |> Option.map #cond
blanchet@54440
   605
          |> the_default (Const (@{const_name undefined}, dummyT)))
blanchet@54440
   606
        |> fst o split_last;
blanchet@54440
   607
    fun finish t =
blanchet@54440
   608
      let val n = length (fastype_of t |> binder_types) in
blanchet@54440
   609
        if t = Const (@{const_name undefined}, dummyT) then t
blanchet@54440
   610
        else if n = 0 then Abs (Name.uu_, @{typ unit}, t)
blanchet@54440
   611
        else if n = 1 then t
blanchet@54440
   612
        else Const (@{const_name prod_case}, dummyT) $ t
blanchet@54440
   613
      end;
blanchet@54440
   614
  in
blanchet@54440
   615
    map finish args
blanchet@54440
   616
  end;
blanchet@54440
   617
blanchet@54440
   618
fun build_corec_args_sel sel_eqns ctr_spec =
blanchet@54440
   619
  let
blanchet@54440
   620
    (* FIXME *)
blanchet@54440
   621
    val n_args = fold (curry (op +)) (map (fn Direct_Corec _ => 3 | _ => 1) (#calls ctr_spec)) 0;
blanchet@54440
   622
  in
blanchet@54440
   623
    replicate n_args (Const (@{const_name undefined}, dummyT))
blanchet@54440
   624
  end;
blanchet@54440
   625
blanchet@54440
   626
fun co_build_defs lthy sequential bs arg_Tss fun_name_corec_spec_list eqns_data =
blanchet@54440
   627
  let
blanchet@54440
   628
    val fun_names = map Binding.name_of bs;
blanchet@54440
   629
blanchet@54440
   630
(*    fun group _ [] = [] (* FIXME \<exists>? *)
blanchet@54440
   631
      | group eq (x :: xs) =
blanchet@54440
   632
        let val (xs', ys) = List.partition (eq x) xs in (x :: xs') :: group eq ys end;*)
blanchet@54440
   633
    val disc_eqnss = map_filter (try (fn Dtr_Disc x => x)) eqns_data
blanchet@54440
   634
      |> partition_eq ((op =) o pairself #fun_name)
blanchet@54440
   635
      |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
blanchet@54440
   636
      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
blanchet@54440
   637
blanchet@54440
   638
    val _ = disc_eqnss |> map (fn x =>
blanchet@54440
   639
      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
blanchet@54440
   640
        primrec_error_eqns "excess discriminator equations in definition"
blanchet@54440
   641
          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
blanchet@54440
   642
blanchet@54440
   643
val _ = warning ("disc_eqnss:\n    \<cdot> " ^ space_implode "\n    \<cdot> " (map @{make_string} disc_eqnss));
blanchet@54440
   644
blanchet@54440
   645
    val sel_eqnss = map_filter (try (fn Dtr_Sel x => x)) eqns_data
blanchet@54440
   646
      |> partition_eq ((op =) o pairself #fun_name)
blanchet@54440
   647
      |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
blanchet@54440
   648
      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
blanchet@54440
   649
blanchet@54440
   650
val _ = warning ("sel_eqnss:\n    \<cdot> " ^ space_implode "\n    \<cdot> " (map @{make_string} sel_eqnss));
blanchet@54440
   651
blanchet@54440
   652
    fun splice (xs :: xss) (ys :: yss) = xs @ ys @ splice xss yss (* FIXME \<exists>? *)
blanchet@54440
   653
      | splice xss yss = flat xss @ flat yss;
blanchet@54440
   654
    val corecs = map (#corec o snd) fun_name_corec_spec_list;
blanchet@54440
   655
    val corec_args = (map snd fun_name_corec_spec_list ~~ disc_eqnss ~~ sel_eqnss)
blanchet@54440
   656
      |> maps (fn (({ctr_specs, ...}, disc_eqns), sel_eqns) =>
blanchet@54440
   657
        splice (build_corec_args_discs ctr_specs disc_eqns |> map single)
blanchet@54440
   658
          (map (build_corec_args_sel sel_eqns) ctr_specs));
blanchet@54440
   659
blanchet@54440
   660
val _ = warning ("corecursor arguments:\n    \<cdot> " ^
blanchet@54440
   661
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) corec_args));
blanchet@54440
   662
blanchet@54440
   663
    fun uneq_pairs_rev xs = xs (* FIXME \<exists>? *)
blanchet@54440
   664
      |> these o try (split_last #> (fn (ys, y) => uneq_pairs_rev ys @ map (pair y) ys));
blanchet@54440
   665
    val proof_obligations = if sequential then [] else
blanchet@54440
   666
      maps (uneq_pairs_rev o map #cond) disc_eqnss
blanchet@54440
   667
      |> map (fn (x, y) => ((strip_abs_body x, strip_abs_body y), strip_abs_vars x))
blanchet@54440
   668
      |> map (apfst (apsnd HOLogic.mk_not #> pairself HOLogic.mk_Trueprop
blanchet@54440
   669
        #> apfst (curry (op $) @{const ==>}) #> (op $)))
blanchet@54440
   670
      |> map (fn (t, abs_vars) => fold_rev (fn (v, T) => fn u =>
blanchet@54440
   671
          Const (@{const_name all}, (T --> @{typ prop}) --> @{typ prop}) $
blanchet@54440
   672
            Abs (v, T, u)) abs_vars t);
blanchet@54440
   673
blanchet@54440
   674
    fun currys Ts t = if length Ts <= 1 then t else
blanchet@54440
   675
      t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
blanchet@54440
   676
        (length Ts - 1 downto 0 |> map Bound)
blanchet@54441
   677
      |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts;
blanchet@54440
   678
  in
blanchet@54440
   679
    map (list_comb o rpair corec_args) corecs
blanchet@54440
   680
    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
blanchet@54440
   681
    |> map2 currys arg_Tss
blanchet@54440
   682
    |> Syntax.check_terms lthy
blanchet@54440
   683
    |> map2 (fn b => fn t => ((b, NoSyn), ((Binding.map_name Thm.def_name b, []), t))) bs
blanchet@54440
   684
    |> rpair proof_obligations
blanchet@54440
   685
  end;
blanchet@54440
   686
blanchet@54440
   687
fun add_primcorec sequential fixes specs lthy =
blanchet@54440
   688
  let
blanchet@54440
   689
    val bs = map (fst o fst) fixes;
blanchet@54440
   690
    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
blanchet@54440
   691
blanchet@54440
   692
    (* copied from primrec_new *)
blanchet@54440
   693
    fun get_indices t = map (fst #>> Binding.name_of #> Free) fixes
blanchet@54440
   694
      |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
blanchet@54440
   695
      |> map_filter I;
blanchet@54440
   696
blanchet@54440
   697
    val callssss = []; (* FIXME *)
blanchet@54440
   698
blanchet@54440
   699
    val ((nontriv, corec_specs, _, coinduct_thm, strong_co_induct_thm, coinduct_thmss,
blanchet@54440
   700
          strong_coinduct_thmss), lthy') =
blanchet@54440
   701
      corec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
blanchet@54440
   702
blanchet@54440
   703
    val fun_names = map Binding.name_of bs;
blanchet@54440
   704
blanchet@54440
   705
    val fun_name_corec_spec_list = (fun_names ~~ res_Ts, corec_specs)
blanchet@54440
   706
      |> uncurry (finds (fn ((v, T), {corec, ...}) => T = body_type (fastype_of corec))) |> fst
blanchet@54440
   707
      |> map (apfst fst #> apsnd the_single); (*###*)
blanchet@54440
   708
blanchet@54440
   709
    val (eqns_data, _) =
blanchet@54440
   710
      fold_map (co_dissect_eqn sequential fun_name_corec_spec_list) (map snd specs) []
blanchet@54440
   711
      |>> flat;
blanchet@54440
   712
blanchet@54440
   713
    val (defs, proof_obligations) =
blanchet@54440
   714
      co_build_defs lthy' sequential bs (map (binder_types o snd o fst) fixes)
blanchet@54440
   715
        fun_name_corec_spec_list eqns_data;
blanchet@54440
   716
  in
blanchet@54440
   717
    lthy'
blanchet@54440
   718
    |> fold_map Local_Theory.define defs |> snd
blanchet@54440
   719
    |> Proof.theorem NONE (K I) [map (rpair []) proof_obligations]
blanchet@54440
   720
    |> Proof.refine (Method.primitive_text I)
blanchet@54440
   721
    |> Seq.hd
blanchet@54440
   722
  end
blanchet@54440
   723
blanchet@54440
   724
fun add_primcorec_cmd seq (raw_fixes, raw_specs) lthy =
blanchet@54440
   725
  let
blanchet@54440
   726
    val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy);
blanchet@54440
   727
  in
blanchet@54440
   728
    add_primcorec seq fixes specs lthy
blanchet@54440
   729
    handle ERROR str => primrec_error str
blanchet@54440
   730
  end
blanchet@54440
   731
  handle Primrec_Error (str, eqns) =>
blanchet@54440
   732
    if null eqns
blanchet@54440
   733
    then error ("primcorec error:\n  " ^ str)
blanchet@54440
   734
    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
blanchet@54440
   735
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns))
blanchet@54440
   736
blanchet@54440
   737
val _ = Outer_Syntax.local_theory_to_proof
blanchet@54440
   738
  @{command_spec "primcorec"}
blanchet@54440
   739
  "define primitive corecursive functions"
blanchet@54440
   740
  (Parse.opt_keyword "sequential" -- (Parse.fixes -- Parse_Spec.where_alt_specs) >>
blanchet@54440
   741
    uncurry add_primcorec_cmd);
blanchet@54440
   742
blanchet@54440
   743
end;