src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
author blanchet
Sat, 26 Oct 2013 12:54:21 +0200
changeset 55656 5151b84d0668
parent 55654 0a06b51ffa56
child 55658 1c26d55b8d73
permissions -rw-r--r--
tuned error message
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
traytel@55150
    10
  val add_primrec: (binding * typ option * mixfix) list ->
traytel@55150
    11
    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
blanchet@54440
    12
  val add_primrec_cmd: (binding * string option * mixfix) list ->
traytel@55150
    13
    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
traytel@55150
    14
  val add_primrec_global: (binding * typ option * mixfix) list ->
traytel@55150
    15
    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
traytel@55150
    16
  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
traytel@55150
    17
    (binding * typ option * mixfix) list ->
traytel@55150
    18
    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
traytel@55150
    19
  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
traytel@55150
    20
    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
blanchet@54890
    21
  val add_primcorecursive_cmd: bool ->
panny@54968
    22
    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
panny@54968
    23
    Proof.context -> Proof.state
panny@54959
    24
  val add_primcorec_cmd: bool ->
panny@54968
    25
    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
panny@54968
    26
    local_theory -> local_theory
blanchet@54440
    27
end;
blanchet@54440
    28
blanchet@54440
    29
structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
blanchet@54440
    30
struct
blanchet@54440
    31
blanchet@54440
    32
open BNF_Util
blanchet@54440
    33
open BNF_FP_Util
blanchet@54440
    34
open BNF_FP_Rec_Sugar_Util
blanchet@54440
    35
open BNF_FP_Rec_Sugar_Tactics
blanchet@54440
    36
panny@55517
    37
val codeN = "code";
panny@55517
    38
val ctrN = "ctr";
panny@55517
    39
val discN = "disc";
panny@55517
    40
val selN = "sel";
blanchet@54928
    41
blanchet@55597
    42
val nitpicksimp_attrs = @{attributes [nitpick_simp]};
blanchet@54931
    43
val simp_attrs = @{attributes [simp]};
blanchet@55597
    44
val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
blanchet@55597
    45
val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
blanchet@54931
    46
blanchet@54440
    47
exception Primrec_Error of string * term list;
blanchet@54440
    48
blanchet@54440
    49
fun primrec_error str = raise Primrec_Error (str, []);
blanchet@54440
    50
fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
blanchet@54440
    51
fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
blanchet@54440
    52
panny@54495
    53
fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
panny@54495
    54
panny@54494
    55
val free_name = try (fn Free (v, _) => v);
panny@54494
    56
val const_name = try (fn Const (v, _) => v);
panny@54495
    57
val undef_const = Const (@{const_name undefined}, dummyT);
panny@54494
    58
panny@54495
    59
fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
panny@54857
    60
  |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
panny@54538
    61
val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
panny@54791
    62
fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
panny@54791
    63
  strip_qnt_body @{const_name all} t)
panny@54857
    64
fun abstract vs =
panny@54857
    65
  let fun a n (t $ u) = a n t $ a n u
panny@54857
    66
        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
panny@54857
    67
        | a n t = let val idx = find_index (equal t) vs in
panny@54857
    68
            if idx < 0 then t else Bound (n + idx) end
panny@54857
    69
  in a 0 end;
panny@54872
    70
fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u;
panny@54872
    71
fun mk_tuple1 Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 Ts));
blanchet@54440
    72
blanchet@54931
    73
fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
blanchet@54931
    74
  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
blanchet@54931
    75
  |> map_filter I;
blanchet@54440
    76
blanchet@54447
    77
blanchet@54447
    78
(* Primrec *)
blanchet@54447
    79
blanchet@54440
    80
type eqn_data = {
blanchet@54440
    81
  fun_name: string,
blanchet@54440
    82
  rec_type: typ,
blanchet@54440
    83
  ctr: term,
blanchet@54440
    84
  ctr_args: term list,
blanchet@54440
    85
  left_args: term list,
blanchet@54440
    86
  right_args: term list,
blanchet@54440
    87
  res_type: typ,
blanchet@54440
    88
  rhs_term: term,
blanchet@54440
    89
  user_eqn: term
blanchet@54440
    90
};
blanchet@54440
    91
blanchet@54440
    92
fun dissect_eqn lthy fun_names eqn' =
blanchet@54440
    93
  let
panny@54791
    94
    val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
panny@54791
    95
      handle TERM _ =>
panny@54791
    96
        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
blanchet@54440
    97
    val (lhs, rhs) = HOLogic.dest_eq eqn
blanchet@54440
    98
        handle TERM _ =>
blanchet@54440
    99
          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
blanchet@54440
   100
    val (fun_name, args) = strip_comb lhs
blanchet@54440
   101
      |>> (fn x => if is_Free x then fst (dest_Free x)
blanchet@54440
   102
          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
blanchet@54440
   103
    val (left_args, rest) = take_prefix is_Free args;
blanchet@54440
   104
    val (nonfrees, right_args) = take_suffix is_Free rest;
blanchet@54967
   105
    val num_nonfrees = length nonfrees;
blanchet@54967
   106
    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
blanchet@54440
   107
      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
blanchet@54440
   108
      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
blanchet@54440
   109
    val _ = member (op =) fun_names fun_name orelse
blanchet@54440
   110
      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
blanchet@54440
   111
blanchet@54440
   112
    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
blanchet@54440
   113
    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
blanchet@54440
   114
      primrec_error_eqn "partially applied constructor in pattern" eqn;
blanchet@54440
   115
    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
blanchet@54440
   116
      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
blanchet@54440
   117
        "\" in left-hand side") eqn end;
blanchet@54440
   118
    val _ = forall is_Free ctr_args orelse
blanchet@54440
   119
      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
blanchet@54440
   120
    val _ =
blanchet@54440
   121
      let val b = fold_aterms (fn x as Free (v, _) =>
blanchet@54440
   122
        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
blanchet@54440
   123
        not (member (op =) fun_names v) andalso
blanchet@54440
   124
        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
blanchet@54440
   125
      in
blanchet@54440
   126
        null b orelse
blanchet@54440
   127
        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
blanchet@54440
   128
          commas (map (Syntax.string_of_term lthy) b)) eqn
blanchet@54440
   129
      end;
blanchet@54440
   130
  in
blanchet@54440
   131
    {fun_name = fun_name,
blanchet@54440
   132
     rec_type = body_type (type_of ctr),
blanchet@54440
   133
     ctr = ctr,
blanchet@54440
   134
     ctr_args = ctr_args,
blanchet@54440
   135
     left_args = left_args,
blanchet@54440
   136
     right_args = right_args,
blanchet@54440
   137
     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
blanchet@54440
   138
     rhs_term = rhs,
blanchet@54440
   139
     user_eqn = eqn'}
blanchet@54440
   140
  end;
blanchet@54440
   141
panny@54538
   142
fun rewrite_map_arg get_ctr_pos rec_type res_type =
blanchet@54440
   143
  let
blanchet@54440
   144
    val pT = HOLogic.mk_prodT (rec_type, res_type);
blanchet@54440
   145
panny@54494
   146
    val maybe_suc = Option.map (fn x => x + 1);
panny@54494
   147
    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
panny@54494
   148
      | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
panny@54494
   149
      | subst d t =
panny@54495
   150
        let
panny@54495
   151
          val (u, vs) = strip_comb t;
panny@54538
   152
          val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1;
panny@54495
   153
        in
panny@54538
   154
          if ctr_pos >= 0 then
panny@54494
   155
            if d = SOME ~1 andalso length vs = ctr_pos then
panny@54494
   156
              list_comb (permute_args ctr_pos (snd_const pT), vs)
panny@54494
   157
            else if length vs > ctr_pos andalso is_some d
panny@54494
   158
                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
panny@54494
   159
              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
blanchet@54440
   160
            else
panny@54494
   161
              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
panny@54494
   162
          else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then
panny@54494
   163
            list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs)
blanchet@54440
   164
          else
panny@54494
   165
            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
blanchet@54440
   166
        end
blanchet@54440
   167
  in
panny@54494
   168
    subst (SOME ~1)
blanchet@54440
   169
  end;
blanchet@54440
   170
blanchet@55622
   171
fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
blanchet@54440
   172
  let
blanchet@55621
   173
    fun try_nested_rec bound_Ts y t =
blanchet@55621
   174
      AList.lookup (op =) nested_calls y
blanchet@55621
   175
      |> Option.map (fn y' =>
blanchet@55621
   176
        massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
blanchet@55621
   177
blanchet@55618
   178
    fun subst bound_Ts (t as g' $ y) =
blanchet@55622
   179
        let
blanchet@55622
   180
          fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
blanchet@55622
   181
          val y_head = head_of y;
blanchet@55622
   182
        in
blanchet@55611
   183
          if not (member (op =) ctr_args y_head) then
blanchet@55622
   184
            subst_rec ()
blanchet@55611
   185
          else
blanchet@55621
   186
            (case try_nested_rec bound_Ts y_head t of
blanchet@55621
   187
              SOME t' => t'
blanchet@55620
   188
            | NONE =>
blanchet@55620
   189
              let val (g, g_args) = strip_comb g' in
blanchet@55620
   190
                (case try (get_ctr_pos o the) (free_name g) of
blanchet@55620
   191
                  SOME ctr_pos =>
blanchet@55620
   192
                  (length g_args >= ctr_pos orelse
blanchet@55620
   193
                   primrec_error_eqn "too few arguments in recursive call" t;
blanchet@55620
   194
                   (case AList.lookup (op =) mutual_calls y of
blanchet@55620
   195
                     SOME y' => list_comb (y', g_args)
blanchet@55622
   196
                   | NONE => subst_rec ()))
blanchet@55622
   197
                | NONE => subst_rec ())
blanchet@55620
   198
              end)
blanchet@55611
   199
        end
blanchet@55618
   200
      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
panny@54487
   201
      | subst _ t = t
blanchet@55622
   202
blanchet@55622
   203
    fun subst' t =
blanchet@55622
   204
      if has_call t then
blanchet@55622
   205
        (* FIXME detect this case earlier? *)
blanchet@55622
   206
        primrec_error_eqn "recursive call not directly applied to constructor argument" t
blanchet@55622
   207
      else
blanchet@55622
   208
        try_nested_rec [] (head_of t) t |> the_default t
panny@54487
   209
  in
blanchet@55622
   210
    subst' o subst []
panny@54487
   211
  end;
blanchet@54440
   212
blanchet@55140
   213
fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
blanchet@55140
   214
    (maybe_eqn_data : eqn_data option) =
panny@54495
   215
  if is_none maybe_eqn_data then undef_const else
blanchet@54440
   216
    let
blanchet@54440
   217
      val eqn_data = the maybe_eqn_data;
blanchet@54440
   218
      val t = #rhs_term eqn_data;
blanchet@54440
   219
      val ctr_args = #ctr_args eqn_data;
blanchet@54440
   220
blanchet@54440
   221
      val calls = #calls ctr_spec;
blanchet@55652
   222
      val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
blanchet@54440
   223
blanchet@54440
   224
      val no_calls' = tag_list 0 calls
blanchet@55652
   225
        |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
blanchet@55554
   226
      val mutual_calls' = tag_list 0 calls
blanchet@55652
   227
        |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
blanchet@55554
   228
      val nested_calls' = tag_list 0 calls
blanchet@55652
   229
        |> map_filter (try (apsnd (fn Nested_Rec p => p)));
blanchet@54440
   230
blanchet@54440
   231
      val args = replicate n_args ("", dummyT)
blanchet@54440
   232
        |> Term.rename_wrt_term t
blanchet@54440
   233
        |> map Free
blanchet@55653
   234
        |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
blanchet@54440
   235
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
blanchet@54440
   236
          no_calls'
blanchet@55652
   237
        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
blanchet@55652
   238
            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
blanchet@55554
   239
          mutual_calls'
blanchet@55652
   240
        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
blanchet@55652
   241
            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
blanchet@55554
   242
          nested_calls';
blanchet@54440
   243
panny@54538
   244
      val fun_name_ctr_pos_list =
panny@54538
   245
        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
panny@54538
   246
      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
blanchet@55652
   247
      val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
blanchet@55652
   248
      val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
blanchet@54440
   249
panny@54538
   250
      val abstractions = args @ #left_args eqn_data @ #right_args eqn_data;
blanchet@54440
   251
    in
panny@54487
   252
      t
blanchet@55554
   253
      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
panny@54538
   254
      |> fold_rev lambda abstractions
panny@54487
   255
    end;
blanchet@54440
   256
blanchet@55140
   257
fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
blanchet@54440
   258
  let
blanchet@54440
   259
    val n_funs = length funs_data;
blanchet@54440
   260
blanchet@54440
   261
    val ctr_spec_eqn_data_list' =
blanchet@54440
   262
      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
blanchet@54440
   263
      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
blanchet@54440
   264
          ##> (fn x => null x orelse
blanchet@54440
   265
            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
blanchet@54440
   266
    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
blanchet@54440
   267
      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
blanchet@54440
   268
blanchet@54440
   269
    val ctr_spec_eqn_data_list =
blanchet@54440
   270
      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
blanchet@54440
   271
blanchet@54440
   272
    val recs = take n_funs rec_specs |> map #recx;
blanchet@54440
   273
    val rec_args = ctr_spec_eqn_data_list
blanchet@54440
   274
      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
panny@54495
   275
      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
blanchet@54440
   276
    val ctr_poss = map (fn x =>
blanchet@54440
   277
      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
blanchet@54440
   278
        primrec_error ("inconstant constructor pattern position for function " ^
blanchet@54440
   279
          quote (#fun_name (hd x)))
blanchet@54440
   280
      else
blanchet@54440
   281
        hd x |> #left_args |> length) funs_data;
blanchet@54440
   282
  in
blanchet@54440
   283
    (recs, ctr_poss)
blanchet@54440
   284
    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
blanchet@54440
   285
    |> Syntax.check_terms lthy
blanchet@55607
   286
    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
blanchet@55607
   287
      bs mxs
blanchet@54440
   288
  end;
blanchet@54440
   289
blanchet@55654
   290
fun massage_comp ctxt has_call bound_Ts t =
blanchet@55654
   291
  massage_nested_corec_call ctxt has_call (K (K (K I))) bound_Ts (fastype_of1 (bound_Ts, t)) t;
blanchet@55654
   292
blanchet@55654
   293
fun find_rec_calls ctxt has_call (eqn_data : eqn_data) =
blanchet@54440
   294
  let
blanchet@55654
   295
    fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
blanchet@55654
   296
      | find bound_Ts (t as _ $ _) ctr_arg =
blanchet@54440
   297
        let
blanchet@55654
   298
          val typof = curry fastype_of1 bound_Ts;
blanchet@54440
   299
          val (f', args') = strip_comb t;
blanchet@55654
   300
          val n = find_index (equal ctr_arg o head_of) args';
blanchet@54440
   301
        in
blanchet@54440
   302
          if n < 0 then
blanchet@55654
   303
            find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
blanchet@54440
   304
          else
blanchet@55654
   305
            let
blanchet@55654
   306
              val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
blanchet@55654
   307
              val (arg_head, arg_args) = Term.strip_comb arg;
blanchet@55654
   308
            in
panny@54495
   309
              if has_call f then
blanchet@55654
   310
                mk_partial_compN (length arg_args) (typof f) (typof arg_head) f ::
blanchet@55654
   311
                maps (fn x => find bound_Ts x ctr_arg) args
blanchet@54440
   312
              else
blanchet@55654
   313
                find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
blanchet@54440
   314
            end
blanchet@54440
   315
        end
blanchet@55654
   316
      | find _ _ _ = [];
blanchet@54440
   317
  in
blanchet@55654
   318
    map (find [] (#rhs_term eqn_data)) (#ctr_args eqn_data)
blanchet@54440
   319
    |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
blanchet@54440
   320
  end;
blanchet@54440
   321
traytel@55150
   322
fun prepare_primrec fixes specs lthy =
blanchet@54440
   323
  let
traytel@54489
   324
    val (bs, mxs) = map_split (apfst fst) fixes;
blanchet@54440
   325
    val fun_names = map Binding.name_of bs;
traytel@55150
   326
    val eqns_data = map (dissect_eqn lthy fun_names) specs;
blanchet@54440
   327
    val funs_data = eqns_data
blanchet@54440
   328
      |> partition_eq ((op =) o pairself #fun_name)
blanchet@54440
   329
      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
blanchet@54440
   330
      |> map (fn (x, y) => the_single y handle List.Empty =>
blanchet@54440
   331
          primrec_error ("missing equations for function " ^ quote x));
blanchet@54440
   332
panny@54495
   333
    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
blanchet@54440
   334
    val arg_Ts = map (#rec_type o hd) funs_data;
blanchet@54440
   335
    val res_Ts = map (#res_type o hd) funs_data;
blanchet@54440
   336
    val callssss = funs_data
blanchet@54440
   337
      |> map (partition_eq ((op =) o pairself #ctr))
blanchet@55654
   338
      |> map (maps (map_filter (find_rec_calls lthy has_call)));
blanchet@54440
   339
blanchet@54967
   340
    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
blanchet@54931
   341
      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
blanchet@54440
   342
blanchet@54440
   343
    val actual_nn = length funs_data;
blanchet@54440
   344
blanchet@54440
   345
    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
blanchet@54440
   346
      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
blanchet@54440
   347
        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
blanchet@54440
   348
          " is not a constructor in left-hand side") user_eqn) eqns_data end;
blanchet@54440
   349
panny@54495
   350
    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
blanchet@54440
   351
traytel@55150
   352
    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
traytel@55150
   353
        (fun_data : eqn_data list) =
blanchet@54440
   354
      let
blanchet@54440
   355
        val def_thms = map (snd o snd) def_thms';
traytel@55150
   356
        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
blanchet@54440
   357
          |> fst
blanchet@54440
   358
          |> map_filter (try (fn (x, [y]) =>
blanchet@54440
   359
            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
blanchet@54440
   360
          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
blanchet@54466
   361
            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
blanchet@55628
   362
            |> K |> Goal.prove lthy [] [] user_eqn
blanchet@55628
   363
            |> Thm.close_derivation);
traytel@55150
   364
        val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
traytel@55150
   365
      in
traytel@55150
   366
        (poss, simp_thmss)
traytel@55150
   367
      end;
blanchet@54440
   368
traytel@55150
   369
    val notes =
traytel@55150
   370
      (if n2m then map2 (fn name => fn thm =>
traytel@55150
   371
        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
traytel@55150
   372
      |> map (fn (prefix, thmN, thms, attrs) =>
traytel@55150
   373
        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
blanchet@54440
   374
blanchet@54440
   375
    val common_name = mk_common_name fun_names;
blanchet@54440
   376
blanchet@54440
   377
    val common_notes =
traytel@55150
   378
      (if n2m then [(inductN, [induct_thm], [])] else [])
blanchet@54440
   379
      |> map (fn (thmN, thms, attrs) =>
blanchet@54440
   380
        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
blanchet@54440
   381
  in
traytel@55150
   382
    (((fun_names, defs),
traytel@55150
   383
      fn lthy => fn defs =>
traytel@55150
   384
        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
traytel@55150
   385
      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
blanchet@54440
   386
  end;
blanchet@54440
   387
traytel@55150
   388
(* primrec definition *)
traytel@55150
   389
traytel@55150
   390
fun add_primrec_simple fixes ts lthy =
blanchet@54440
   391
  let
traytel@55150
   392
    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
traytel@55150
   393
      handle ERROR str => primrec_error str;
blanchet@54440
   394
  in
traytel@55150
   395
    lthy
traytel@55150
   396
    |> fold_map Local_Theory.define defs
traytel@55150
   397
    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
blanchet@54440
   398
  end
blanchet@54440
   399
  handle Primrec_Error (str, eqns) =>
blanchet@54440
   400
    if null eqns
blanchet@54440
   401
    then error ("primrec_new error:\n  " ^ str)
blanchet@54440
   402
    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
panny@54959
   403
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
blanchet@54440
   404
traytel@55150
   405
local
traytel@55150
   406
traytel@55165
   407
fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
traytel@55150
   408
  let
traytel@55150
   409
    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
traytel@55150
   410
    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
traytel@55150
   411
traytel@55150
   412
    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
traytel@55150
   413
traytel@55150
   414
    val mk_notes =
traytel@55150
   415
      flat ooo map3 (fn poss => fn prefix => fn thms =>
traytel@55150
   416
        let
traytel@55150
   417
          val (bs, attrss) = map_split (fst o nth specs) poss;
traytel@55150
   418
          val notes =
traytel@55150
   419
            map3 (fn b => fn attrs => fn thm =>
blanchet@55597
   420
              ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
traytel@55150
   421
            bs attrss thms;
traytel@55150
   422
        in
traytel@55150
   423
          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
traytel@55150
   424
        end);
traytel@55150
   425
  in
traytel@55150
   426
    lthy
traytel@55150
   427
    |> add_primrec_simple fixes (map snd specs)
traytel@55150
   428
    |-> (fn (names, (ts, (posss, simpss))) =>
traytel@55150
   429
      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
traytel@55150
   430
      #> Local_Theory.notes (mk_notes posss names simpss)
traytel@55150
   431
      #>> pair ts o map snd)
traytel@55150
   432
  end;
traytel@55150
   433
traytel@55150
   434
in
traytel@55150
   435
traytel@55150
   436
val add_primrec = gen_primrec Specification.check_spec;
traytel@55150
   437
val add_primrec_cmd = gen_primrec Specification.read_spec;
traytel@55150
   438
traytel@55150
   439
end;
traytel@55150
   440
traytel@55150
   441
fun add_primrec_global fixes specs thy =
traytel@55150
   442
  let
traytel@55150
   443
    val lthy = Named_Target.theory_init thy;
traytel@55150
   444
    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
traytel@55150
   445
    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
traytel@55150
   446
  in ((ts, simps'), Local_Theory.exit_global lthy') end;
traytel@55150
   447
traytel@55150
   448
fun add_primrec_overloaded ops fixes specs thy =
traytel@55150
   449
  let
traytel@55150
   450
    val lthy = Overloading.overloading ops thy;
traytel@55150
   451
    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
traytel@55150
   452
    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
traytel@55150
   453
  in ((ts, simps'), Local_Theory.exit_global lthy') end;
traytel@55150
   454
blanchet@54440
   455
blanchet@54440
   456
blanchet@54447
   457
(* Primcorec *)
blanchet@54440
   458
blanchet@55605
   459
type coeqn_data_disc = {
blanchet@54440
   460
  fun_name: string,
panny@54857
   461
  fun_T: typ,
panny@54538
   462
  fun_args: term list,
panny@54857
   463
  ctr: term,
panny@54478
   464
  ctr_no: int, (*###*)
panny@54857
   465
  disc: term,
panny@54791
   466
  prems: term list,
panny@54959
   467
  auto_gen: bool,
panny@55549
   468
  maybe_ctr_rhs: term option,
panny@55549
   469
  maybe_code_rhs: term option,
blanchet@54440
   470
  user_eqn: term
blanchet@54440
   471
};
blanchet@55138
   472
blanchet@55605
   473
type coeqn_data_sel = {
blanchet@54440
   474
  fun_name: string,
panny@54857
   475
  fun_T: typ,
panny@54538
   476
  fun_args: term list,
panny@54478
   477
  ctr: term,
panny@54478
   478
  sel: term,
blanchet@54440
   479
  rhs_term: term,
blanchet@54440
   480
  user_eqn: term
blanchet@54440
   481
};
blanchet@55138
   482
blanchet@55605
   483
datatype coeqn_data =
blanchet@55605
   484
  Disc of coeqn_data_disc |
blanchet@55605
   485
  Sel of coeqn_data_sel;
blanchet@54440
   486
panny@55612
   487
fun dissect_coeqn_disc seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
panny@55612
   488
    maybe_ctr_rhs maybe_code_rhs prems' concl matchedsss =
blanchet@54440
   489
  let
blanchet@54440
   490
    fun find_subterm p = let (* FIXME \<exists>? *)
panny@54538
   491
      fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
blanchet@54440
   492
        | f t = if p t then SOME t else NONE
blanchet@54440
   493
      in f end;
blanchet@54440
   494
panny@54791
   495
    val applied_fun = concl
panny@54791
   496
      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
panny@54791
   497
      |> the
blanchet@55656
   498
      handle Option.Option => primrec_error_eqn "malformed discriminator formula" concl;
panny@54857
   499
    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
panny@55612
   500
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
blanchet@54440
   501
panny@55612
   502
    val discs = map #disc basic_ctr_specs;
panny@55612
   503
    val ctrs = map #ctr basic_ctr_specs;
panny@54791
   504
    val not_disc = head_of concl = @{term Not};
panny@54538
   505
    val _ = not_disc andalso length ctrs <> 2 andalso
panny@54791
   506
      primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
panny@55612
   507
    val disc' = find_subterm (member (op =) discs o head_of) concl;
panny@54791
   508
    val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
blanchet@54440
   509
        |> (fn SOME t => let val n = find_index (equal t) ctrs in
blanchet@54440
   510
          if n >= 0 then SOME n else NONE end | _ => NONE);
panny@55612
   511
    val _ = is_some disc' orelse is_some eq_ctr0 orelse
panny@54791
   512
      primrec_error_eqn "no discriminator in equation" concl;
blanchet@54440
   513
    val ctr_no' =
panny@55612
   514
      if is_none disc' then the eq_ctr0 else find_index (equal (head_of (the disc'))) discs;
blanchet@54440
   515
    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
panny@55612
   516
    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
blanchet@54440
   517
panny@54791
   518
    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
panny@54857
   519
    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
panny@54857
   520
    val prems = map (abstract (List.rev fun_args)) prems';
panny@54857
   521
    val real_prems =
blanchet@55519
   522
      (if catch_all orelse seq then maps s_not_conj matchedss else []) @
panny@54791
   523
      (if catch_all then [] else prems);
blanchet@54440
   524
panny@54857
   525
    val matchedsss' = AList.delete (op =) fun_name matchedsss
panny@55517
   526
      |> cons (fun_name, if seq then matchedss @ [prems] else matchedss @ [real_prems]);
panny@54791
   527
panny@54791
   528
    val user_eqn =
panny@55549
   529
      (real_prems, concl)
panny@55549
   530
      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
panny@55549
   531
      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
blanchet@54440
   532
  in
panny@54478
   533
    (Disc {
blanchet@54440
   534
      fun_name = fun_name,
panny@54857
   535
      fun_T = fun_T,
panny@54538
   536
      fun_args = fun_args,
panny@54857
   537
      ctr = ctr,
blanchet@54440
   538
      ctr_no = ctr_no,
panny@55612
   539
      disc = disc,
panny@54791
   540
      prems = real_prems,
panny@54959
   541
      auto_gen = catch_all,
panny@55549
   542
      maybe_ctr_rhs = maybe_ctr_rhs,
panny@55549
   543
      maybe_code_rhs = maybe_code_rhs,
panny@54791
   544
      user_eqn = user_eqn
panny@54857
   545
    }, matchedsss')
blanchet@54440
   546
  end;
blanchet@54440
   547
panny@55612
   548
fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn' of_spec
panny@55612
   549
    eqn =
blanchet@54440
   550
  let
blanchet@54440
   551
    val (lhs, rhs) = HOLogic.dest_eq eqn
blanchet@54440
   552
      handle TERM _ =>
blanchet@54440
   553
        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
blanchet@54440
   554
    val sel = head_of lhs;
panny@54857
   555
    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
blanchet@54440
   556
      handle TERM _ =>
blanchet@54440
   557
        primrec_error_eqn "malformed selector argument in left-hand side" eqn;
panny@55612
   558
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
blanchet@54440
   559
      handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
panny@55612
   560
    val {ctr, ...} =
panny@54968
   561
      if is_some of_spec
panny@55612
   562
      then the (find_first (equal (the of_spec) o #ctr) basic_ctr_specs)
panny@55612
   563
      else filter (exists (equal sel) o #sels) basic_ctr_specs |> the_single
panny@54968
   564
        handle List.Empty => primrec_error_eqn "ambiguous selector - use \"of\"" eqn;
panny@54791
   565
    val user_eqn = drop_All eqn';
blanchet@54440
   566
  in
panny@54478
   567
    Sel {
blanchet@54440
   568
      fun_name = fun_name,
panny@54857
   569
      fun_T = fun_T,
panny@54538
   570
      fun_args = fun_args,
panny@55612
   571
      ctr = ctr,
panny@54478
   572
      sel = sel,
blanchet@54440
   573
      rhs_term = rhs,
panny@54791
   574
      user_eqn = user_eqn
blanchet@54440
   575
    }
blanchet@54440
   576
  end;
blanchet@54440
   577
panny@55612
   578
fun dissect_coeqn_ctr seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
panny@55612
   579
    maybe_code_rhs prems concl matchedsss =
blanchet@55047
   580
  let
panny@55517
   581
    val (lhs, rhs) = HOLogic.dest_eq concl;
panny@55549
   582
    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
panny@55612
   583
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
blanchet@55526
   584
    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
panny@55612
   585
    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) basic_ctr_specs)
blanchet@54440
   586
      handle Option.Option => primrec_error_eqn "not a constructor" ctr;
panny@54478
   587
panny@55517
   588
    val disc_concl = betapply (disc, lhs);
panny@55612
   589
    val (maybe_eqn_data_disc, matchedsss') = if length basic_ctr_specs = 1
panny@54857
   590
      then (NONE, matchedsss)
panny@55612
   591
      else apfst SOME (dissect_coeqn_disc seq fun_names basic_ctr_specss
panny@55549
   592
          (SOME (abstract (List.rev fun_args) rhs)) maybe_code_rhs prems disc_concl matchedsss);
blanchet@54440
   593
blanchet@55611
   594
    val sel_concls = sels ~~ ctr_args
blanchet@54440
   595
      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
blanchet@54440
   596
blanchet@54993
   597
(*
panny@55517
   598
val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
panny@55517
   599
 (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
panny@55549
   600
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
panny@55549
   601
 "\nfor premise(s)\n    \<cdot> " ^
panny@55549
   602
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
blanchet@54993
   603
*)
blanchet@54440
   604
panny@55612
   605
    val eqns_data_sel =
panny@55612
   606
      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn' (SOME ctr)) sel_concls;
blanchet@54440
   607
  in
panny@54857
   608
    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
blanchet@54440
   609
  end;
blanchet@54440
   610
panny@55612
   611
fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss =
panny@55517
   612
  let
panny@55517
   613
    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
panny@55549
   614
    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
panny@55612
   615
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
panny@55517
   616
panny@55517
   617
    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
panny@55612
   618
        if member ((op =) o apsnd #ctr) basic_ctr_specs ctr
panny@55517
   619
        then cons (ctr, cs)
panny@55517
   620
        else primrec_error_eqn "not a constructor" ctr) [] rhs' []
panny@55517
   621
      |> AList.group (op =);
panny@55517
   622
blanchet@55520
   623
    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
panny@55517
   624
    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
panny@55517
   625
        binder_types (fastype_of ctr)
panny@55517
   626
        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
panny@55517
   627
          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
panny@55517
   628
        |> curry list_comb ctr
panny@55517
   629
        |> curry HOLogic.mk_eq lhs);
panny@55517
   630
  in
panny@55612
   631
    fold_map2 (dissect_coeqn_ctr false fun_names basic_ctr_specss eqn'
panny@55549
   632
        (SOME (abstract (List.rev fun_args) rhs)))
panny@55549
   633
      ctr_premss ctr_concls matchedsss
panny@55517
   634
  end;
panny@55517
   635
panny@55612
   636
fun dissect_coeqn lthy seq has_call fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
panny@55612
   637
    eqn' of_spec matchedsss =
blanchet@54440
   638
  let
panny@54791
   639
    val eqn = drop_All eqn'
panny@54791
   640
      handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
panny@55517
   641
    val (prems, concl) = Logic.strip_horn eqn
panny@54478
   642
      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
blanchet@54440
   643
panny@55517
   644
    val head = concl
blanchet@54440
   645
      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
blanchet@54440
   646
      |> head_of;
blanchet@54440
   647
panny@55517
   648
    val maybe_rhs = concl |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
blanchet@54440
   649
panny@55612
   650
    val discs = maps (map #disc) basic_ctr_specss;
panny@55612
   651
    val sels = maps (maps #sels) basic_ctr_specss;
panny@55612
   652
    val ctrs = maps (map #ctr) basic_ctr_specss;
blanchet@54440
   653
  in
blanchet@54440
   654
    if member (op =) discs head orelse
blanchet@54440
   655
      is_some maybe_rhs andalso
blanchet@54440
   656
        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
panny@55612
   657
      dissect_coeqn_disc seq fun_names basic_ctr_specss NONE NONE prems concl matchedsss
blanchet@54440
   658
      |>> single
blanchet@54440
   659
    else if member (op =) sels head then
panny@55612
   660
      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn' of_spec concl], matchedsss)
panny@55517
   661
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
blanchet@55526
   662
      member (op =) ctrs (head_of (unfold_let (the maybe_rhs))) then
panny@55612
   663
      dissect_coeqn_ctr seq fun_names basic_ctr_specss eqn' NONE prems concl matchedsss
panny@55517
   664
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
panny@55517
   665
      null prems then
panny@55612
   666
      dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss
panny@55517
   667
      |>> flat
blanchet@54440
   668
    else
blanchet@54440
   669
      primrec_error_eqn "malformed function equation" eqn
blanchet@54440
   670
  end;
blanchet@54440
   671
blanchet@55139
   672
fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
blanchet@55605
   673
    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
panny@54791
   674
  if is_none (#pred (nth ctr_specs ctr_no)) then I else
blanchet@55520
   675
    s_conjs prems
panny@54791
   676
    |> curry subst_bounds (List.rev fun_args)
panny@54791
   677
    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
panny@54791
   678
    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
blanchet@54440
   679
blanchet@55605
   680
fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
panny@54857
   681
  find_first (equal sel o #sel) sel_eqns
panny@54857
   682
  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
panny@54857
   683
  |> the_default undef_const
panny@54548
   684
  |> K;
panny@54497
   685
blanchet@55605
   686
fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
panny@54497
   687
  let
panny@54548
   688
    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
panny@54497
   689
  in
panny@55013
   690
    if is_none maybe_sel_eqn then (I, I, I) else
panny@55013
   691
    let
panny@55013
   692
      val {fun_args, rhs_term, ... } = the maybe_sel_eqn;
panny@55549
   693
      val bound_Ts = List.rev (map fastype_of fun_args);
blanchet@55027
   694
      fun rewrite_q _ t = if has_call t then @{term False} else @{term True};
blanchet@55027
   695
      fun rewrite_g _ t = if has_call t then undef_const else t;
panny@55036
   696
      fun rewrite_h bound_Ts t =
panny@55036
   697
        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
blanchet@55554
   698
      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
panny@55549
   699
        |> abs_tuple fun_args;
panny@55013
   700
    in
panny@55013
   701
      (massage rewrite_q,
panny@55013
   702
       massage rewrite_g,
panny@55013
   703
       massage rewrite_h)
panny@55013
   704
    end
panny@54497
   705
  end;
panny@54497
   706
blanchet@55605
   707
fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
panny@54548
   708
  let
panny@54548
   709
    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
panny@55036
   710
  in
panny@55036
   711
    if is_none maybe_sel_eqn then I else
panny@55036
   712
    let
panny@55036
   713
      val {fun_args, rhs_term, ...} = the maybe_sel_eqn;
panny@55549
   714
      val bound_Ts = List.rev (map fastype_of fun_args);
panny@55036
   715
      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
panny@55036
   716
        | rewrite bound_Ts U T (t as _ $ _) =
panny@55036
   717
          let val (u, vs) = strip_comb t in
panny@55036
   718
            if is_Free u andalso has_call u then
panny@55036
   719
              Inr_const U T $ mk_tuple1 bound_Ts vs
panny@55036
   720
            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
panny@55036
   721
              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
panny@55036
   722
            else
panny@55036
   723
              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
panny@55036
   724
          end
panny@55036
   725
        | rewrite _ U T t =
panny@55036
   726
          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
panny@55036
   727
      fun massage t =
panny@55549
   728
        rhs_term
blanchet@55554
   729
        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
panny@54872
   730
        |> abs_tuple fun_args;
panny@55036
   731
    in
panny@55036
   732
      massage
panny@55036
   733
    end
panny@54548
   734
  end;
panny@54497
   735
blanchet@55605
   736
fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
blanchet@55139
   737
    (ctr_spec : corec_ctr_spec) =
panny@54478
   738
  let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
panny@54478
   739
    if null sel_eqns then I else
panny@54478
   740
      let
panny@54478
   741
        val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
panny@54478
   742
panny@54478
   743
        val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
blanchet@55554
   744
        val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
blanchet@55554
   745
        val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
panny@54478
   746
      in
panny@54497
   747
        I
panny@54872
   748
        #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
panny@54497
   749
        #> fold (fn (sel, (q, g, h)) =>
blanchet@55554
   750
          let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
blanchet@55554
   751
            nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
panny@54497
   752
        #> fold (fn (sel, n) => nth_map n
blanchet@55554
   753
          (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
panny@54478
   754
      end
blanchet@54440
   755
  end;
blanchet@54440
   756
blanchet@55605
   757
fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
blanchet@55605
   758
    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
panny@54791
   759
  let
panny@55612
   760
    val corecs = map #corec corec_specs;
panny@55612
   761
    val ctr_specss = map #ctr_specs corec_specs;
panny@54497
   762
    val corec_args = hd corecs
panny@54497
   763
      |> fst o split_last o binder_types o fastype_of
panny@54497
   764
      |> map (Const o pair @{const_name undefined})
panny@54857
   765
      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
panny@54497
   766
      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
panny@54872
   767
    fun currys [] t = t
panny@54872
   768
      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
panny@54872
   769
          |> fold_rev (Term.abs o pair Name.uu) Ts;
panny@54538
   770
blanchet@54993
   771
(*
panny@54497
   772
val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
panny@54548
   773
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
blanchet@54993
   774
*)
blanchet@54440
   775
panny@54791
   776
    val exclss' =
panny@54857
   777
      disc_eqnss
panny@54959
   778
      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
panny@54791
   779
        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
panny@54791
   780
        #> maps (uncurry (map o pair)
panny@54959
   781
          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
blanchet@55520
   782
              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
panny@54791
   783
            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
panny@54791
   784
            ||> Logic.list_implies
panny@54791
   785
            ||> curry Logic.list_all (map dest_Free fun_args))))
blanchet@54440
   786
  in
blanchet@54440
   787
    map (list_comb o rpair corec_args) corecs
blanchet@54440
   788
    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
blanchet@54440
   789
    |> map2 currys arg_Tss
blanchet@54440
   790
    |> Syntax.check_terms lthy
blanchet@55607
   791
    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
blanchet@55607
   792
      bs mxs
panny@54791
   793
    |> rpair exclss'
blanchet@54440
   794
  end;
blanchet@54440
   795
blanchet@55139
   796
fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
blanchet@55605
   797
    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
panny@54857
   798
  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
panny@54857
   799
    let
panny@54857
   800
      val n = 0 upto length ctr_specs
panny@54857
   801
        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
panny@54859
   802
      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
panny@54859
   803
        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
panny@54857
   804
      val extra_disc_eqn = {
panny@54857
   805
        fun_name = Binding.name_of fun_binding,
panny@54857
   806
        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
panny@54859
   807
        fun_args = fun_args,
panny@54857
   808
        ctr = #ctr (nth ctr_specs n),
panny@54857
   809
        ctr_no = n,
panny@54857
   810
        disc = #disc (nth ctr_specs n),
blanchet@55519
   811
        prems = maps (s_not_conj o #prems) disc_eqns,
panny@54959
   812
        auto_gen = true,
panny@55549
   813
        maybe_ctr_rhs = NONE,
panny@55549
   814
        maybe_code_rhs = NONE,
panny@54857
   815
        user_eqn = undef_const};
panny@54857
   816
    in
panny@54857
   817
      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
panny@54857
   818
    end;
panny@54857
   819
panny@55612
   820
fun find_corec_calls has_call basic_ctr_specs {ctr, sel, rhs_term, ...} =
panny@55612
   821
  let
panny@55612
   822
    val sel_no = find_first (equal ctr o #ctr) basic_ctr_specs
panny@55612
   823
      |> find_index (equal sel) o #sels o the;
panny@55612
   824
    fun find (Abs (_, _, b)) = find b
panny@55612
   825
      | find (t as _ $ _) = strip_comb t |>> find ||> maps find |> (op @)
panny@55612
   826
      | find f = if is_Free f andalso has_call f then [f] else [];
panny@55612
   827
  in
panny@55612
   828
    find rhs_term
panny@55612
   829
    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
panny@55612
   830
  end;
panny@55612
   831
blanchet@55629
   832
fun add_primcorec_ursive maybe_tac seq fixes specs of_specs lthy =
blanchet@54440
   833
  let
traytel@54489
   834
    val (bs, mxs) = map_split (apfst fst) fixes;
blanchet@54440
   835
    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
blanchet@54440
   836
panny@55612
   837
    val fun_names = map Binding.name_of bs;
panny@55612
   838
    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
panny@55612
   839
    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
panny@55612
   840
    val eqns_data =
panny@55612
   841
      fold_map2 (dissect_coeqn lthy seq has_call fun_names basic_ctr_specss) (map snd specs)
panny@55612
   842
        of_specs []
panny@55612
   843
      |> flat o fst;
panny@55612
   844
panny@55612
   845
    val callssss =
panny@55612
   846
      map_filter (try (fn Sel x => x)) eqns_data
panny@55612
   847
      |> partition_eq ((op =) o pairself #fun_name)
panny@55612
   848
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
panny@55613
   849
      |> map (flat o snd)
panny@55613
   850
      |> map2 (fold o find_corec_calls has_call) basic_ctr_specss
panny@55612
   851
      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
panny@55612
   852
        (ctr, map (K []) sels))) basic_ctr_specss);
panny@55612
   853
panny@55612
   854
(*
panny@55612
   855
val _ = tracing ("callssss = " ^ @{make_string} callssss);
panny@55612
   856
*)
blanchet@54440
   857
blanchet@54967
   858
    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
blanchet@54934
   859
          strong_coinduct_thms), lthy') =
blanchet@54931
   860
      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
blanchet@54967
   861
    val actual_nn = length bs;
blanchet@54967
   862
    val corec_specs = take actual_nn corec_specs'; (*###*)
blanchet@55630
   863
    val ctr_specss = map #ctr_specs corec_specs;
blanchet@54440
   864
panny@54857
   865
    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
panny@54791
   866
      |> partition_eq ((op =) o pairself #fun_name)
panny@54857
   867
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
panny@54791
   868
      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
panny@54857
   869
    val _ = disc_eqnss' |> map (fn x =>
panny@54857
   870
      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
blanchet@55656
   871
        primrec_error_eqns "excess discriminator formula in definition"
panny@54857
   872
          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
panny@54791
   873
panny@54791
   874
    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
panny@54791
   875
      |> partition_eq ((op =) o pairself #fun_name)
panny@54857
   876
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
panny@54791
   877
      |> map (flat o snd);
panny@54791
   878
panny@54497
   879
    val arg_Tss = map (binder_types o snd o fst) fixes;
panny@54859
   880
    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
panny@54791
   881
    val (defs, exclss') =
blanchet@55605
   882
      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
panny@54791
   883
blanchet@55060
   884
    fun excl_tac (c, c', a) =
blanchet@55629
   885
      if a orelse c = c' orelse seq then SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
blanchet@55629
   886
      else maybe_tac;
panny@54959
   887
blanchet@54993
   888
(*
panny@54959
   889
val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
panny@54959
   890
 space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
blanchet@54993
   891
*)
panny@54959
   892
panny@54959
   893
    val exclss'' = exclss' |> map (map (fn (idx, t) =>
blanchet@55629
   894
      (idx, (Option.map (Goal.prove lthy [] [] t #> Thm.close_derivation) (excl_tac idx), t))));
panny@54791
   895
    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
blanchet@55632
   896
    val (goal_idxss, goalss) = exclss''
panny@54791
   897
      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
panny@54791
   898
      |> split_list o map split_list;
panny@54791
   899
panny@54791
   900
    fun prove thmss' def_thms' lthy =
panny@54791
   901
      let
panny@54791
   902
        val def_thms = map (snd o snd) def_thms';
panny@54791
   903
blanchet@55632
   904
        val exclss' = map (op ~~) (goal_idxss ~~ thmss');
panny@54791
   905
        fun mk_exclsss excls n =
panny@54791
   906
          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
panny@54959
   907
          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
panny@54791
   908
        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
panny@54791
   909
          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
panny@54791
   910
blanchet@55139
   911
        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
blanchet@55605
   912
            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
panny@54859
   913
          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else
panny@54857
   914
            let
panny@54859
   915
              val {disc_corec, ...} = nth ctr_specs ctr_no;
panny@54857
   916
              val k = 1 + ctr_no;
panny@54857
   917
              val m = length prems;
panny@54857
   918
              val t =
panny@54857
   919
                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
panny@54857
   920
                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
panny@54857
   921
                |> HOLogic.mk_Trueprop
panny@54857
   922
                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
panny@54857
   923
                |> curry Logic.list_all (map dest_Free fun_args);
panny@54857
   924
            in
panny@55549
   925
              if prems = [@{term False}] then [] else
panny@54857
   926
              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
panny@54857
   927
              |> K |> Goal.prove lthy [] [] t
blanchet@55628
   928
              |> Thm.close_derivation
panny@54857
   929
              |> pair (#disc (nth ctr_specs ctr_no))
panny@54857
   930
              |> single
panny@54857
   931
            end;
panny@54857
   932
blanchet@55139
   933
        fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
blanchet@55605
   934
            : corec_spec) (disc_eqns : coeqn_data_disc list) exclsss
blanchet@55605
   935
            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : coeqn_data_sel) =
panny@54791
   936
          let
blanchet@55046
   937
            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
panny@54857
   938
            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
blanchet@55519
   939
            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
panny@54857
   940
                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
panny@54857
   941
            val sel_corec = find_index (equal sel) (#sels ctr_spec)
panny@54857
   942
              |> nth (#sel_corecs ctr_spec);
panny@54791
   943
            val k = 1 + ctr_no;
panny@54791
   944
            val m = length prems;
panny@54791
   945
            val t =
panny@54857
   946
              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
panny@54857
   947
              |> curry betapply sel
panny@54857
   948
              |> rpair (abstract (List.rev fun_args) rhs_term)
panny@54857
   949
              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
panny@54791
   950
              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
panny@54857
   951
              |> curry Logic.list_all (map dest_Free fun_args);
blanchet@55062
   952
            val (distincts, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
panny@54791
   953
          in
blanchet@55055
   954
            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
blanchet@55047
   955
              nested_map_idents nested_map_comps sel_corec k m exclsss
panny@54791
   956
            |> K |> Goal.prove lthy [] [] t
blanchet@55628
   957
            |> Thm.close_derivation
panny@54857
   958
            |> pair sel
panny@54791
   959
          end;
panny@54791
   960
blanchet@55605
   961
        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
blanchet@55605
   962
            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
panny@55549
   963
          (* don't try to prove theorems when some sel_eqns are missing *)
panny@54857
   964
          if not (exists (equal ctr o #ctr) disc_eqns)
panny@54859
   965
              andalso not (exists (equal ctr o #ctr) sel_eqns)
panny@55549
   966
            orelse
panny@54857
   967
              filter (equal ctr o #ctr) sel_eqns
panny@54857
   968
              |> fst o finds ((op =) o apsnd #sel) sels
panny@54857
   969
              |> exists (null o snd)
panny@54857
   970
          then [] else
panny@54857
   971
            let
panny@55549
   972
              val (fun_name, fun_T, fun_args, prems, maybe_rhs) =
panny@54859
   973
                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
panny@55549
   974
                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
panny@55549
   975
                  #maybe_ctr_rhs x))
panny@55549
   976
                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], NONE))
panny@54859
   977
                |> the o merge_options;
panny@54857
   978
              val m = length prems;
panny@55549
   979
              val t = (if is_some maybe_rhs then the maybe_rhs else
panny@55549
   980
                  filter (equal ctr o #ctr) sel_eqns
panny@55549
   981
                  |> fst o finds ((op =) o apsnd #sel) sels
panny@55549
   982
                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
panny@55549
   983
                  |> curry list_comb ctr)
panny@54857
   984
                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
panny@54857
   985
                  map Bound (length fun_args - 1 downto 0)))
panny@54857
   986
                |> HOLogic.mk_Trueprop
panny@54857
   987
                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
panny@54857
   988
                |> curry Logic.list_all (map dest_Free fun_args);
blanchet@54928
   989
              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
blanchet@54928
   990
              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
panny@54857
   991
            in
panny@55549
   992
              if prems = [@{term False}] then [] else
panny@55549
   993
                mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
panny@55549
   994
                |> K |> Goal.prove lthy [] [] t
blanchet@55628
   995
                |> Thm.close_derivation
panny@55549
   996
                |> pair ctr
panny@55549
   997
                |> single
panny@55013
   998
            end;
panny@54857
   999
blanchet@55630
  1000
        fun prove_code disc_eqns sel_eqns ctr_alist ctr_specs =
panny@55550
  1001
          let
panny@55549
  1002
            val (fun_name, fun_T, fun_args, maybe_rhs) =
panny@55549
  1003
              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
panny@55549
  1004
               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
panny@55549
  1005
              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #maybe_code_rhs x))
panny@55549
  1006
              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, NONE))
panny@55549
  1007
              |> the o merge_options;
panny@55549
  1008
blanchet@55625
  1009
            val bound_Ts = List.rev (map fastype_of fun_args);
blanchet@55625
  1010
panny@55572
  1011
            val lhs = list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
blanchet@55625
  1012
            val maybe_rhs_info =
blanchet@55625
  1013
              (case maybe_rhs of
blanchet@55625
  1014
                SOME rhs =>
blanchet@55625
  1015
                let
blanchet@55625
  1016
                  val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
blanchet@55625
  1017
                  val cond_ctrs =
blanchet@55625
  1018
                    fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
blanchet@55625
  1019
                  val ctr_thms = map (the o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
blanchet@55625
  1020
                in SOME (rhs, raw_rhs, ctr_thms) end
blanchet@55625
  1021
              | NONE =>
blanchet@55625
  1022
                let
blanchet@55625
  1023
                  fun prove_code_ctr {ctr, sels, ...} =
blanchet@55625
  1024
                    if not (exists (equal ctr o fst) ctr_alist) then NONE else
blanchet@55625
  1025
                      let
blanchet@55625
  1026
                        val prems = find_first (equal ctr o #ctr) disc_eqns
blanchet@55625
  1027
                          |> Option.map #prems |> the_default [];
blanchet@55625
  1028
                        val t =
blanchet@55625
  1029
                          filter (equal ctr o #ctr) sel_eqns
blanchet@55625
  1030
                          |> fst o finds ((op =) o apsnd #sel) sels
blanchet@55625
  1031
                          |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
blanchet@55625
  1032
                            #-> abstract)
blanchet@55625
  1033
                          |> curry list_comb ctr;
blanchet@55625
  1034
                      in
blanchet@55625
  1035
                        SOME (prems, t)
blanchet@55625
  1036
                      end;
blanchet@55625
  1037
                  val maybe_ctr_conds_argss = map prove_code_ctr ctr_specs;
blanchet@55625
  1038
                in
blanchet@55625
  1039
                  if exists is_none maybe_ctr_conds_argss then NONE else
blanchet@55625
  1040
                    let
blanchet@55625
  1041
                      val rhs = fold_rev (fn SOME (prems, u) => fn t => mk_If (s_conjs prems) u t)
blanchet@55625
  1042
                        maybe_ctr_conds_argss
blanchet@55625
  1043
                        (Const (@{const_name Code.abort}, @{typ String.literal} -->
blanchet@55625
  1044
                            (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
blanchet@55630
  1045
                          HOLogic.mk_literal fun_name $
blanchet@55625
  1046
                          absdummy @{typ unit} (incr_boundvars 1 lhs));
blanchet@55625
  1047
                    in SOME (rhs, rhs, map snd ctr_alist) end
blanchet@55625
  1048
                end);
blanchet@55625
  1049
          in
blanchet@55625
  1050
            (case maybe_rhs_info of
blanchet@55625
  1051
              NONE => []
blanchet@55625
  1052
            | SOME (rhs, raw_rhs, ctr_thms) =>
panny@55549
  1053
              let
panny@55549
  1054
                val ms = map (Logic.count_prems o prop_of) ctr_thms;
blanchet@55624
  1055
                val (raw_t, t) = (raw_rhs, rhs)
blanchet@55624
  1056
                  |> pairself
blanchet@55624
  1057
                    (curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
blanchet@55624
  1058
                      map Bound (length fun_args - 1 downto 0)))
blanchet@55624
  1059
                    #> HOLogic.mk_Trueprop
blanchet@55624
  1060
                    #> curry Logic.list_all (map dest_Free fun_args));
panny@55550
  1061
                val (distincts, discIs, sel_splits, sel_split_asms) =
blanchet@55624
  1062
                  case_thms_of_term lthy bound_Ts raw_rhs;
panny@55550
  1063
blanchet@55624
  1064
                val raw_code_thm = mk_primcorec_raw_code_of_ctr_tac lthy distincts discIs sel_splits
panny@55549
  1065
                    sel_split_asms ms ctr_thms
blanchet@55628
  1066
                  |> K |> Goal.prove lthy [] [] raw_t
blanchet@55628
  1067
                  |> Thm.close_derivation;
blanchet@55624
  1068
              in
blanchet@55626
  1069
                mk_primcorec_code_of_raw_code_tac lthy distincts sel_splits raw_code_thm
blanchet@55624
  1070
                |> K |> Goal.prove lthy [] [] t
blanchet@55628
  1071
                |> Thm.close_derivation
blanchet@55624
  1072
                |> single
blanchet@55625
  1073
              end)
blanchet@55625
  1074
          end;
panny@55549
  1075
blanchet@54928
  1076
        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
blanchet@54928
  1077
        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
blanchet@54928
  1078
        val disc_thmss = map (map snd) disc_alists;
blanchet@54928
  1079
        val sel_thmss = map (map snd) sel_alists;
panny@55549
  1080
panny@55549
  1081
        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
blanchet@55630
  1082
          ctr_specss;
panny@55549
  1083
        val ctr_thmss = map (map snd) ctr_alists;
panny@55549
  1084
blanchet@55630
  1085
        val code_thmss = map4 prove_code disc_eqnss sel_eqnss ctr_alists ctr_specss;
blanchet@54928
  1086
blanchet@55167
  1087
        val simp_thmss = map2 append disc_thmss sel_thmss
blanchet@54932
  1088
blanchet@54934
  1089
        val common_name = mk_common_name fun_names;
blanchet@54934
  1090
blanchet@54928
  1091
        val notes =
blanchet@54967
  1092
          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
blanchet@55597
  1093
           (codeN, code_thmss, code_nitpicksimp_attrs),
blanchet@54934
  1094
           (ctrN, ctr_thmss, []),
blanchet@54928
  1095
           (discN, disc_thmss, simp_attrs),
blanchet@54932
  1096
           (selN, sel_thmss, simp_attrs),
blanchet@54934
  1097
           (simpsN, simp_thmss, []),
blanchet@54967
  1098
           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
blanchet@54928
  1099
          |> maps (fn (thmN, thmss, attrs) =>
blanchet@54928
  1100
            map2 (fn fun_name => fn thms =>
blanchet@54928
  1101
                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
blanchet@54967
  1102
              fun_names (take actual_nn thmss))
blanchet@54928
  1103
          |> filter_out (null o fst o hd o snd);
blanchet@54934
  1104
blanchet@54934
  1105
        val common_notes =
blanchet@54967
  1106
          [(coinductN, if n2m then [coinduct_thm] else [], []),
blanchet@54967
  1107
           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
blanchet@54934
  1108
          |> filter_out (null o #2)
blanchet@54934
  1109
          |> map (fn (thmN, thms, attrs) =>
blanchet@54934
  1110
            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
panny@54791
  1111
      in
blanchet@55167
  1112
        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
panny@54791
  1113
      end;
panny@54959
  1114
panny@54959
  1115
    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
blanchet@54440
  1116
  in
blanchet@55629
  1117
    (goalss, after_qed, lthy')
panny@54959
  1118
  end;
blanchet@54440
  1119
blanchet@55629
  1120
fun add_primcorec_ursive_cmd maybe_tac seq (raw_fixes, raw_specs') lthy =
blanchet@54440
  1121
  let
panny@54968
  1122
    val (raw_specs, of_specs) = split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
panny@54968
  1123
    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
blanchet@54440
  1124
  in
blanchet@55629
  1125
    add_primcorec_ursive maybe_tac seq fixes specs of_specs lthy
blanchet@54440
  1126
    handle ERROR str => primrec_error str
blanchet@54440
  1127
  end
blanchet@54440
  1128
  handle Primrec_Error (str, eqns) =>
blanchet@54440
  1129
    if null eqns
blanchet@54440
  1130
    then error ("primcorec error:\n  " ^ str)
blanchet@54440
  1131
    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
panny@54959
  1132
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
panny@54959
  1133
blanchet@55629
  1134
val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
blanchet@55629
  1135
  lthy
blanchet@55629
  1136
  |> Proof.theorem NONE after_qed goalss
blanchet@55629
  1137
  |> Proof.refine (Method.primitive_text I)
blanchet@55629
  1138
  |> Seq.hd) ooo add_primcorec_ursive_cmd NONE;
blanchet@55629
  1139
blanchet@55629
  1140
val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
blanchet@55629
  1141
  lthy
blanchet@55629
  1142
  |> after_qed (map (fn [] => []
blanchet@55629
  1143
      | _ => primrec_error "need exclusiveness proofs - use primcorecursive instead of primcorec")
blanchet@55629
  1144
    goalss)) ooo add_primcorec_ursive_cmd (SOME (fn {context = ctxt, ...} => auto_tac ctxt));
blanchet@54440
  1145
blanchet@54440
  1146
end;