src/HOL/BNF/Tools/bnf_gfp_rec_sugar.ML
author panny
Mon, 25 Nov 2013 20:25:22 +0100
changeset 55964 c822230fd22b
parent 55731 3ffb74b52ed6
child 55986 985f8b49c050
permissions -rw-r--r--
prevent exception when equations for a function are missing;
pave the way for "exhaustive"
blanchet@55698
     1
(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar.ML
blanchet@54440
     2
    Author:     Lorenz Panny, TU Muenchen
blanchet@55698
     3
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@54440
     4
    Copyright   2013
blanchet@54440
     5
blanchet@55698
     6
Corecursor sugar.
blanchet@54440
     7
*)
blanchet@54440
     8
blanchet@55698
     9
signature BNF_GFP_REC_SUGAR =
blanchet@54440
    10
sig
panny@55964
    11
  datatype primcorec_option =
panny@55964
    12
    Option_Sequential |
panny@55964
    13
    Option_Exhaustive
panny@55964
    14
  val add_primcorecursive_cmd: primcorec_option list ->
panny@54968
    15
    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
panny@54968
    16
    Proof.context -> Proof.state
panny@55964
    17
  val add_primcorec_cmd: primcorec_option list ->
panny@54968
    18
    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
panny@54968
    19
    local_theory -> local_theory
blanchet@54440
    20
end;
blanchet@54440
    21
blanchet@55698
    22
structure BNF_GFP_Rec_Sugar : BNF_GFP_REC_SUGAR =
blanchet@54440
    23
struct
blanchet@54440
    24
blanchet@55698
    25
open Ctr_Sugar
blanchet@54440
    26
open BNF_Util
blanchet@55698
    27
open BNF_Def
blanchet@54440
    28
open BNF_FP_Util
blanchet@55698
    29
open BNF_FP_Def_Sugar
blanchet@55695
    30
open BNF_FP_N2M_Sugar
blanchet@54440
    31
open BNF_FP_Rec_Sugar_Util
blanchet@55698
    32
open BNF_GFP_Rec_Sugar_Tactics
blanchet@54440
    33
blanchet@55698
    34
val codeN = "code"
blanchet@55698
    35
val ctrN = "ctr"
blanchet@55698
    36
val discN = "disc"
blanchet@55698
    37
val selN = "sel"
blanchet@54928
    38
blanchet@55597
    39
val nitpicksimp_attrs = @{attributes [nitpick_simp]};
blanchet@54931
    40
val simp_attrs = @{attributes [simp]};
blanchet@55597
    41
val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
blanchet@54931
    42
blanchet@55698
    43
exception Primcorec_Error of string * term list;
blanchet@54440
    44
blanchet@55698
    45
fun primcorec_error str = raise Primcorec_Error (str, []);
blanchet@55698
    46
fun primcorec_error_eqn str eqn = raise Primcorec_Error (str, [eqn]);
blanchet@55698
    47
fun primcorec_error_eqns str eqns = raise Primcorec_Error (str, eqns);
blanchet@54440
    48
panny@55964
    49
datatype primcorec_option =
panny@55964
    50
  Option_Sequential |
panny@55964
    51
  Option_Exhaustive
panny@55964
    52
blanchet@55698
    53
datatype corec_call =
blanchet@55698
    54
  Dummy_No_Corec of int |
blanchet@55698
    55
  No_Corec of int |
blanchet@55698
    56
  Mutual_Corec of int * int * int |
blanchet@55698
    57
  Nested_Corec of int;
panny@54495
    58
blanchet@55698
    59
type basic_corec_ctr_spec =
blanchet@55698
    60
  {ctr: term,
blanchet@55698
    61
   disc: term,
blanchet@55698
    62
   sels: term list};
blanchet@55698
    63
blanchet@55698
    64
type corec_ctr_spec =
blanchet@55698
    65
  {ctr: term,
blanchet@55698
    66
   disc: term,
blanchet@55698
    67
   sels: term list,
blanchet@55698
    68
   pred: int option,
blanchet@55698
    69
   calls: corec_call list,
blanchet@55698
    70
   discI: thm,
blanchet@55698
    71
   sel_thms: thm list,
blanchet@55698
    72
   collapse: thm,
blanchet@55698
    73
   corec_thm: thm,
blanchet@55698
    74
   disc_corec: thm,
blanchet@55698
    75
   sel_corecs: thm list};
blanchet@55698
    76
blanchet@55698
    77
type corec_spec =
blanchet@55698
    78
  {corec: term,
blanchet@55698
    79
   nested_map_idents: thm list,
blanchet@55698
    80
   nested_map_comps: thm list,
blanchet@55698
    81
   ctr_specs: corec_ctr_spec list};
blanchet@55698
    82
blanchet@55698
    83
exception AINT_NO_MAP of term;
blanchet@55698
    84
blanchet@55698
    85
fun not_codatatype ctxt T =
blanchet@55698
    86
  error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
blanchet@55698
    87
fun ill_formed_corec_call ctxt t =
blanchet@55698
    88
  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
blanchet@55698
    89
fun invalid_map ctxt t =
blanchet@55698
    90
  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
blanchet@55698
    91
fun unexpected_corec_call ctxt t =
blanchet@55698
    92
  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
blanchet@55698
    93
blanchet@55698
    94
val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
blanchet@55698
    95
val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
blanchet@55698
    96
blanchet@55698
    97
val conjuncts_s = filter_out (curry (op =) @{const True}) o HOLogic.conjuncts;
blanchet@55698
    98
blanchet@55698
    99
fun s_not @{const True} = @{const False}
blanchet@55698
   100
  | s_not @{const False} = @{const True}
blanchet@55698
   101
  | s_not (@{const Not} $ t) = t
blanchet@55698
   102
  | s_not (@{const conj} $ t $ u) = @{const disj} $ s_not t $ s_not u
blanchet@55698
   103
  | s_not (@{const disj} $ t $ u) = @{const conj} $ s_not t $ s_not u
blanchet@55698
   104
  | s_not t = @{const Not} $ t;
blanchet@55698
   105
blanchet@55698
   106
val s_not_conj = conjuncts_s o s_not o mk_conjs;
blanchet@55698
   107
blanchet@55698
   108
fun propagate_unit_pos u cs = if member (op aconv) cs u then [@{const False}] else cs;
blanchet@55698
   109
blanchet@55698
   110
fun propagate_unit_neg not_u cs = remove (op aconv) not_u cs;
blanchet@55698
   111
blanchet@55698
   112
fun propagate_units css =
blanchet@55698
   113
  (case List.partition (can the_single) css of
blanchet@55698
   114
     ([], _) => css
blanchet@55698
   115
   | ([u] :: uss, css') =>
blanchet@55698
   116
     [u] :: propagate_units (map (propagate_unit_neg (s_not u))
blanchet@55698
   117
       (map (propagate_unit_pos u) (uss @ css'))));
blanchet@55698
   118
blanchet@55698
   119
fun s_conjs cs =
blanchet@55698
   120
  if member (op aconv) cs @{const False} then @{const False}
blanchet@55698
   121
  else mk_conjs (remove (op aconv) @{const True} cs);
blanchet@55698
   122
blanchet@55698
   123
fun s_disjs ds =
blanchet@55698
   124
  if member (op aconv) ds @{const True} then @{const True}
blanchet@55698
   125
  else mk_disjs (remove (op aconv) @{const False} ds);
blanchet@55698
   126
blanchet@55698
   127
fun s_dnf css0 =
blanchet@55698
   128
  let val css = propagate_units css0 in
blanchet@55698
   129
    if null css then
blanchet@55698
   130
      [@{const False}]
blanchet@55698
   131
    else if exists null css then
blanchet@55698
   132
      []
blanchet@55698
   133
    else
blanchet@55698
   134
      map (fn c :: cs => (c, cs)) css
blanchet@55698
   135
      |> AList.coalesce (op =)
blanchet@55698
   136
      |> map (fn (c, css) => c :: s_dnf css)
blanchet@55698
   137
      |> (fn [cs] => cs | css => [s_disjs (map s_conjs css)])
blanchet@55698
   138
  end;
blanchet@55698
   139
blanchet@55698
   140
fun fold_rev_let_if_case ctxt f bound_Ts t =
blanchet@55698
   141
  let
blanchet@55698
   142
    val thy = Proof_Context.theory_of ctxt;
blanchet@55698
   143
blanchet@55698
   144
    fun fld conds t =
blanchet@55698
   145
      (case Term.strip_comb t of
blanchet@55698
   146
        (Const (@{const_name Let}, _), [_, _]) => fld conds (unfold_let t)
blanchet@55698
   147
      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
blanchet@55698
   148
        fld (conds @ conjuncts_s cond) then_branch o fld (conds @ s_not_conj [cond]) else_branch
blanchet@55698
   149
      | (Const (c, _), args as _ :: _ :: _) =>
blanchet@55698
   150
        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
blanchet@55698
   151
          if n >= 0 andalso n < length args then
blanchet@55698
   152
            (case fastype_of1 (bound_Ts, nth args n) of
blanchet@55698
   153
              Type (s, Ts) =>
blanchet@55698
   154
              (case dest_case ctxt s Ts t of
blanchet@55698
   155
                NONE => apsnd (f conds t)
blanchet@55698
   156
              | SOME (conds', branches) =>
blanchet@55698
   157
                apfst (cons s) o fold_rev (uncurry fld)
blanchet@55698
   158
                  (map (append conds o conjuncts_s) conds' ~~ branches))
blanchet@55698
   159
            | _ => apsnd (f conds t))
blanchet@55698
   160
          else
blanchet@55698
   161
            apsnd (f conds t)
blanchet@55698
   162
        end
blanchet@55698
   163
      | _ => apsnd (f conds t))
blanchet@55698
   164
  in
blanchet@55698
   165
    fld [] t o pair []
blanchet@55698
   166
  end;
blanchet@55698
   167
blanchet@55698
   168
fun case_of ctxt = ctr_sugar_of ctxt #> Option.map (fst o dest_Const o #casex);
blanchet@55698
   169
blanchet@55698
   170
fun massage_let_if_case ctxt has_call massage_leaf =
blanchet@55698
   171
  let
blanchet@55698
   172
    val thy = Proof_Context.theory_of ctxt;
blanchet@55698
   173
blanchet@55698
   174
    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
blanchet@55698
   175
blanchet@55698
   176
    fun massage_abs bound_Ts 0 t = massage_rec bound_Ts t
blanchet@55698
   177
      | massage_abs bound_Ts m (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) (m - 1) t)
blanchet@55698
   178
      | massage_abs bound_Ts m t =
blanchet@55698
   179
        let val T = domain_type (fastype_of1 (bound_Ts, t)) in
blanchet@55698
   180
          Abs (Name.uu, T, massage_abs (T :: bound_Ts) (m - 1) (incr_boundvars 1 t $ Bound 0))
blanchet@55698
   181
        end
blanchet@55698
   182
    and massage_rec bound_Ts t =
blanchet@55698
   183
      let val typof = curry fastype_of1 bound_Ts in
blanchet@55698
   184
        (case Term.strip_comb t of
blanchet@55698
   185
          (Const (@{const_name Let}, _), [_, _]) => massage_rec bound_Ts (unfold_let t)
blanchet@55698
   186
        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
blanchet@55698
   187
          let val branches' = map (massage_rec bound_Ts) branches in
blanchet@55698
   188
            Term.list_comb (If_const (typof (hd branches')) $ tap check_no_call obj, branches')
blanchet@55698
   189
          end
blanchet@55698
   190
        | (Const (c, _), args as _ :: _ :: _) =>
blanchet@55698
   191
          (case try strip_fun_type (Sign.the_const_type thy c) of
blanchet@55698
   192
            SOME (gen_branch_Ts, gen_body_fun_T) =>
blanchet@55698
   193
            let
blanchet@55698
   194
              val gen_branch_ms = map num_binder_types gen_branch_Ts;
blanchet@55698
   195
              val n = length gen_branch_ms;
blanchet@55698
   196
            in
blanchet@55698
   197
              if n < length args then
blanchet@55698
   198
                (case gen_body_fun_T of
blanchet@55698
   199
                  Type (_, [Type (T_name, _), _]) =>
blanchet@55698
   200
                  if case_of ctxt T_name = SOME c then
blanchet@55698
   201
                    let
blanchet@55698
   202
                      val (branches, obj_leftovers) = chop n args;
blanchet@55698
   203
                      val branches' = map2 (massage_abs bound_Ts) gen_branch_ms branches;
blanchet@55698
   204
                      val branch_Ts' = map typof branches';
blanchet@55698
   205
                      val body_T' = snd (strip_typeN (hd gen_branch_ms) (hd branch_Ts'));
blanchet@55698
   206
                      val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers ---> body_T');
blanchet@55698
   207
                    in
blanchet@55724
   208
                      Term.list_comb (casex',
blanchet@55724
   209
                        branches' @ tap (List.app check_no_call) obj_leftovers)
blanchet@55698
   210
                    end
blanchet@55698
   211
                  else
blanchet@55698
   212
                    massage_leaf bound_Ts t
blanchet@55698
   213
                | _ => massage_leaf bound_Ts t)
blanchet@55698
   214
              else
blanchet@55698
   215
                massage_leaf bound_Ts t
blanchet@55698
   216
            end
blanchet@55698
   217
          | NONE => massage_leaf bound_Ts t)
blanchet@55698
   218
        | _ => massage_leaf bound_Ts t)
blanchet@55698
   219
      end
blanchet@55698
   220
  in
blanchet@55698
   221
    massage_rec
blanchet@55698
   222
  end;
blanchet@55698
   223
blanchet@55698
   224
val massage_mutual_corec_call = massage_let_if_case;
blanchet@55698
   225
blanchet@55698
   226
fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
blanchet@55698
   227
blanchet@55698
   228
fun massage_nested_corec_call ctxt has_call raw_massage_call bound_Ts U t =
blanchet@55698
   229
  let
blanchet@55698
   230
    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
blanchet@55698
   231
blanchet@55698
   232
    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd);
blanchet@55698
   233
blanchet@55698
   234
    fun massage_mutual_call bound_Ts U T t =
blanchet@55698
   235
      if has_call t then
blanchet@55698
   236
        (case try dest_sumT U of
blanchet@55698
   237
          SOME (U1, U2) => if U1 = T then raw_massage_call bound_Ts T U2 t else invalid_map ctxt t
blanchet@55698
   238
        | NONE => invalid_map ctxt t)
blanchet@55698
   239
      else
blanchet@55698
   240
        build_map_Inl (T, U) $ t;
blanchet@55698
   241
blanchet@55698
   242
    fun massage_mutual_fun bound_Ts U T t =
blanchet@55698
   243
      (case t of
blanchet@55698
   244
        Const (@{const_name comp}, _) $ t1 $ t2 =>
blanchet@55698
   245
        mk_comp bound_Ts (massage_mutual_fun bound_Ts U T t1, tap check_no_call t2)
blanchet@55698
   246
      | _ =>
blanchet@55698
   247
        let
blanchet@55698
   248
          val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
blanchet@55698
   249
            domain_type (fastype_of1 (bound_Ts, t)));
blanchet@55698
   250
        in
blanchet@55698
   251
          Term.lambda var (massage_mutual_call bound_Ts U T (t $ var))
blanchet@55698
   252
        end);
blanchet@55698
   253
blanchet@55698
   254
    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
blanchet@55698
   255
        (case try (dest_map ctxt s) t of
blanchet@55698
   256
          SOME (map0, fs) =>
blanchet@55698
   257
          let
blanchet@55698
   258
            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
blanchet@55698
   259
            val map' = mk_map (length fs) dom_Ts Us map0;
blanchet@55698
   260
            val fs' =
blanchet@55698
   261
              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
blanchet@55698
   262
          in
blanchet@55698
   263
            Term.list_comb (map', fs')
blanchet@55698
   264
          end
blanchet@55698
   265
        | NONE => raise AINT_NO_MAP t)
blanchet@55698
   266
      | massage_map _ _ _ t = raise AINT_NO_MAP t
blanchet@55698
   267
    and massage_map_or_map_arg bound_Ts U T t =
blanchet@55698
   268
      if T = U then
blanchet@55698
   269
        tap check_no_call t
blanchet@55698
   270
      else
blanchet@55698
   271
        massage_map bound_Ts U T t
blanchet@55698
   272
        handle AINT_NO_MAP _ => massage_mutual_fun bound_Ts U T t;
blanchet@55698
   273
blanchet@55698
   274
    fun massage_call bound_Ts U T =
blanchet@55698
   275
      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
blanchet@55698
   276
        if has_call t then
blanchet@55729
   277
          (case t of
blanchet@55729
   278
            Const (@{const_name prod_case}, _) $ t' =>
blanchet@55729
   279
            let
blanchet@55729
   280
              val U' = curried_type U;
blanchet@55729
   281
              val T' = curried_type T;
blanchet@55729
   282
            in
blanchet@55729
   283
              Const (@{const_name prod_case}, U' --> U) $ massage_call bound_Ts U' T' t'
blanchet@55729
   284
            end
blanchet@55729
   285
          | t1 $ t2 =>
blanchet@55729
   286
            (if has_call t2 then
blanchet@55729
   287
              massage_mutual_call bound_Ts U T t
blanchet@55729
   288
            else
blanchet@55729
   289
              massage_map bound_Ts U T t1 $ t2
blanchet@55729
   290
              handle AINT_NO_MAP _ => massage_mutual_call bound_Ts U T t)
blanchet@55729
   291
          | Abs (s, T', t') =>
blanchet@55729
   292
            Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
blanchet@55729
   293
          | _ => massage_mutual_call bound_Ts U T t)
blanchet@55698
   294
        else
blanchet@55698
   295
          build_map_Inl (T, U) $ t) bound_Ts;
blanchet@55698
   296
blanchet@55698
   297
    val T = fastype_of1 (bound_Ts, t);
blanchet@55698
   298
  in
blanchet@55698
   299
    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
blanchet@55698
   300
  end;
blanchet@55698
   301
blanchet@55698
   302
val fold_rev_corec_call = fold_rev_let_if_case;
blanchet@55698
   303
blanchet@55698
   304
fun expand_to_ctr_term ctxt s Ts t =
blanchet@55698
   305
  (case ctr_sugar_of ctxt s of
blanchet@55698
   306
    SOME {ctrs, casex, ...} =>
blanchet@55698
   307
    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
blanchet@55698
   308
  | NONE => raise Fail "expand_to_ctr_term");
blanchet@55698
   309
blanchet@55698
   310
fun expand_corec_code_rhs ctxt has_call bound_Ts t =
blanchet@55698
   311
  (case fastype_of1 (bound_Ts, t) of
blanchet@55698
   312
    Type (s, Ts) =>
blanchet@55698
   313
    massage_let_if_case ctxt has_call (fn _ => fn t =>
blanchet@55698
   314
      if can (dest_ctr ctxt s) t then t else expand_to_ctr_term ctxt s Ts t) bound_Ts t
blanchet@55698
   315
  | _ => raise Fail "expand_corec_code_rhs");
blanchet@55698
   316
blanchet@55698
   317
fun massage_corec_code_rhs ctxt massage_ctr =
blanchet@55698
   318
  massage_let_if_case ctxt (K false)
blanchet@55698
   319
    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
blanchet@55698
   320
blanchet@55698
   321
fun fold_rev_corec_code_rhs ctxt f =
blanchet@55698
   322
  snd ooo fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb);
blanchet@55698
   323
blanchet@55698
   324
fun case_thms_of_term ctxt bound_Ts t =
blanchet@55698
   325
  let
blanchet@55698
   326
    val (caseT_names, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t ();
blanchet@55698
   327
    val ctr_sugars = map (the o ctr_sugar_of ctxt) caseT_names;
blanchet@55698
   328
  in
blanchet@55698
   329
    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #sel_splits ctr_sugars,
blanchet@55698
   330
     maps #sel_split_asms ctr_sugars)
blanchet@55698
   331
  end;
blanchet@55698
   332
blanchet@55698
   333
fun basic_corec_specs_of ctxt res_T =
blanchet@55698
   334
  (case res_T of
blanchet@55698
   335
    Type (T_name, _) =>
blanchet@55698
   336
    (case Ctr_Sugar.ctr_sugar_of ctxt T_name of
blanchet@55698
   337
      NONE => not_codatatype ctxt res_T
blanchet@55698
   338
    | SOME {ctrs, discs, selss, ...} =>
blanchet@55698
   339
      let
blanchet@55698
   340
        val thy = Proof_Context.theory_of ctxt;
blanchet@55724
   341
blanchet@55698
   342
        val gfpT = body_type (fastype_of (hd ctrs));
blanchet@55698
   343
        val As_rho = tvar_subst thy [gfpT] [res_T];
blanchet@55698
   344
        val substA = Term.subst_TVars As_rho;
blanchet@55698
   345
blanchet@55698
   346
        fun mk_spec ctr disc sels = {ctr = substA ctr, disc = substA disc, sels = map substA sels};
blanchet@55698
   347
      in
blanchet@55698
   348
        map3 mk_spec ctrs discs selss
blanchet@55698
   349
      end)
blanchet@55698
   350
  | _ => not_codatatype ctxt res_T);
blanchet@55698
   351
blanchet@55698
   352
fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
blanchet@55698
   353
  let
blanchet@55698
   354
    val thy = Proof_Context.theory_of lthy;
blanchet@55698
   355
blanchet@55698
   356
    val ((missing_res_Ts, perm0_kks,
blanchet@55698
   357
          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
blanchet@55698
   358
            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy') =
blanchet@55698
   359
      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy;
blanchet@55698
   360
blanchet@55698
   361
    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
blanchet@55698
   362
blanchet@55698
   363
    val indices = map #index fp_sugars;
blanchet@55698
   364
    val perm_indices = map #index perm_fp_sugars;
blanchet@55698
   365
blanchet@55698
   366
    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
blanchet@55698
   367
    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
blanchet@55698
   368
    val perm_gfpTs = map (body_type o fastype_of o hd) perm_ctrss;
blanchet@55698
   369
blanchet@55698
   370
    val nn0 = length res_Ts;
blanchet@55698
   371
    val nn = length perm_gfpTs;
blanchet@55698
   372
    val kks = 0 upto nn - 1;
blanchet@55698
   373
    val perm_ns = map length perm_ctr_Tsss;
blanchet@55698
   374
blanchet@55698
   375
    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
blanchet@55698
   376
      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
blanchet@55698
   377
    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
blanchet@55698
   378
      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
blanchet@55698
   379
blanchet@55698
   380
    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
blanchet@55698
   381
    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
blanchet@55698
   382
    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
blanchet@55698
   383
blanchet@55698
   384
    val fun_arg_hs =
blanchet@55698
   385
      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
blanchet@55698
   386
blanchet@55698
   387
    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
blanchet@55698
   388
    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
blanchet@55698
   389
blanchet@55698
   390
    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
blanchet@55698
   391
blanchet@55698
   392
    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
blanchet@55698
   393
    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
blanchet@55698
   394
    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
blanchet@55698
   395
blanchet@55698
   396
    val f_Tssss = unpermute perm_f_Tssss;
blanchet@55698
   397
    val gfpTs = unpermute perm_gfpTs;
blanchet@55698
   398
    val Cs = unpermute perm_Cs;
blanchet@55698
   399
blanchet@55698
   400
    val As_rho = tvar_subst thy (take nn0 gfpTs) res_Ts;
blanchet@55698
   401
    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
blanchet@55698
   402
blanchet@55698
   403
    val substA = Term.subst_TVars As_rho;
blanchet@55698
   404
    val substAT = Term.typ_subst_TVars As_rho;
blanchet@55698
   405
    val substCT = Term.typ_subst_TVars Cs_rho;
blanchet@55698
   406
blanchet@55698
   407
    val perm_Cs' = map substCT perm_Cs;
blanchet@55698
   408
blanchet@55698
   409
    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
blanchet@55698
   410
        (if exists_subtype_in Cs T then Nested_Corec
blanchet@55698
   411
         else if nullary then Dummy_No_Corec
blanchet@55698
   412
         else No_Corec) g_i
blanchet@55698
   413
      | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i');
blanchet@55698
   414
blanchet@55698
   415
    fun mk_ctr_spec ctr disc sels p_ho q_iss f_iss f_Tss discI sel_thms collapse corec_thm
blanchet@55698
   416
        disc_corec sel_corecs =
blanchet@55698
   417
      let val nullary = not (can dest_funT (fastype_of ctr)) in
blanchet@55698
   418
        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_ho,
blanchet@55698
   419
         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
blanchet@55698
   420
         collapse = collapse, corec_thm = corec_thm, disc_corec = disc_corec,
blanchet@55698
   421
         sel_corecs = sel_corecs}
blanchet@55698
   422
      end;
blanchet@55698
   423
blanchet@55698
   424
    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) p_is q_isss f_isss f_Tsss coiter_thmsss
blanchet@55698
   425
        disc_coitersss sel_coiterssss =
blanchet@55698
   426
      let
blanchet@55698
   427
        val ctrs = #ctrs (nth ctr_sugars index);
blanchet@55698
   428
        val discs = #discs (nth ctr_sugars index);
blanchet@55698
   429
        val selss = #selss (nth ctr_sugars index);
blanchet@55698
   430
        val p_ios = map SOME p_is @ [NONE];
blanchet@55698
   431
        val discIs = #discIs (nth ctr_sugars index);
blanchet@55698
   432
        val sel_thmss = #sel_thmss (nth ctr_sugars index);
blanchet@55698
   433
        val collapses = #collapses (nth ctr_sugars index);
blanchet@55698
   434
        val corec_thms = co_rec_of (nth coiter_thmsss index);
blanchet@55698
   435
        val disc_corecs = co_rec_of (nth disc_coitersss index);
blanchet@55698
   436
        val sel_corecss = co_rec_of (nth sel_coiterssss index);
blanchet@55698
   437
      in
blanchet@55698
   438
        map13 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss collapses
blanchet@55698
   439
          corec_thms disc_corecs sel_corecss
blanchet@55698
   440
      end;
blanchet@55698
   441
blanchet@55698
   442
    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
blanchet@55698
   443
          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
blanchet@55698
   444
        p_is q_isss f_isss f_Tsss =
blanchet@55698
   445
      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
blanchet@55698
   446
       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
blanchet@55698
   447
       nested_map_comps = map map_comp_of_bnf nested_bnfs,
blanchet@55698
   448
       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
blanchet@55698
   449
         disc_coitersss sel_coiterssss};
blanchet@55698
   450
  in
blanchet@55698
   451
    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
blanchet@55698
   452
      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
blanchet@55698
   453
      strong_co_induct_of coinduct_thmss), lthy')
blanchet@55698
   454
  end;
blanchet@55698
   455
panny@54495
   456
val undef_const = Const (@{const_name undefined}, dummyT);
panny@54494
   457
panny@54538
   458
val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
panny@54857
   459
fun abstract vs =
panny@54857
   460
  let fun a n (t $ u) = a n t $ a n u
panny@54857
   461
        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
panny@54857
   462
        | a n t = let val idx = find_index (equal t) vs in
panny@54857
   463
            if idx < 0 then t else Bound (n + idx) end
panny@54857
   464
  in a 0 end;
blanchet@55723
   465
blanchet@55723
   466
fun mk_prod1 bound_Ts (t, u) =
blanchet@55723
   467
  HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;
blanchet@55723
   468
fun mk_tuple1 bound_Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 bound_Ts));
blanchet@54440
   469
blanchet@55605
   470
type coeqn_data_disc = {
blanchet@54440
   471
  fun_name: string,
panny@54857
   472
  fun_T: typ,
panny@54538
   473
  fun_args: term list,
panny@54857
   474
  ctr: term,
panny@54478
   475
  ctr_no: int, (*###*)
panny@54857
   476
  disc: term,
panny@54791
   477
  prems: term list,
panny@54959
   478
  auto_gen: bool,
panny@55549
   479
  maybe_ctr_rhs: term option,
panny@55549
   480
  maybe_code_rhs: term option,
blanchet@54440
   481
  user_eqn: term
blanchet@54440
   482
};
blanchet@55138
   483
blanchet@55605
   484
type coeqn_data_sel = {
blanchet@54440
   485
  fun_name: string,
panny@54857
   486
  fun_T: typ,
panny@54538
   487
  fun_args: term list,
panny@54478
   488
  ctr: term,
panny@54478
   489
  sel: term,
blanchet@54440
   490
  rhs_term: term,
blanchet@54440
   491
  user_eqn: term
blanchet@54440
   492
};
blanchet@55138
   493
blanchet@55605
   494
datatype coeqn_data =
blanchet@55605
   495
  Disc of coeqn_data_disc |
blanchet@55605
   496
  Sel of coeqn_data_sel;
blanchet@54440
   497
panny@55612
   498
fun dissect_coeqn_disc seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
panny@55612
   499
    maybe_ctr_rhs maybe_code_rhs prems' concl matchedsss =
blanchet@54440
   500
  let
blanchet@55724
   501
    fun find_subterm p =
blanchet@55724
   502
      let (* FIXME \<exists>? *)
blanchet@55724
   503
        fun find (t as u $ v) = if p t then SOME t else merge_options (find u, find v)
blanchet@55724
   504
          | find t = if p t then SOME t else NONE;
blanchet@55724
   505
      in find end;
blanchet@54440
   506
panny@54791
   507
    val applied_fun = concl
panny@54791
   508
      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
panny@54791
   509
      |> the
blanchet@55698
   510
      handle Option.Option => primcorec_error_eqn "malformed discriminator formula" concl;
panny@54857
   511
    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
blanchet@55661
   512
    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
blanchet@54440
   513
panny@55612
   514
    val discs = map #disc basic_ctr_specs;
panny@55612
   515
    val ctrs = map #ctr basic_ctr_specs;
panny@54791
   516
    val not_disc = head_of concl = @{term Not};
panny@54538
   517
    val _ = not_disc andalso length ctrs <> 2 andalso
blanchet@55698
   518
      primcorec_error_eqn "negated discriminator for a type with \<noteq> 2 constructors" concl;
panny@55612
   519
    val disc' = find_subterm (member (op =) discs o head_of) concl;
blanchet@55661
   520
    val eq_ctr0 = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd)
blanchet@54440
   521
        |> (fn SOME t => let val n = find_index (equal t) ctrs in
blanchet@54440
   522
          if n >= 0 then SOME n else NONE end | _ => NONE);
panny@55612
   523
    val _ = is_some disc' orelse is_some eq_ctr0 orelse
blanchet@55698
   524
      primcorec_error_eqn "no discriminator in equation" concl;
blanchet@54440
   525
    val ctr_no' =
panny@55612
   526
      if is_none disc' then the eq_ctr0 else find_index (equal (head_of (the disc'))) discs;
blanchet@54440
   527
    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
panny@55612
   528
    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
blanchet@54440
   529
panny@54791
   530
    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
panny@54857
   531
    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
panny@54857
   532
    val prems = map (abstract (List.rev fun_args)) prems';
panny@54857
   533
    val real_prems =
blanchet@55519
   534
      (if catch_all orelse seq then maps s_not_conj matchedss else []) @
panny@54791
   535
      (if catch_all then [] else prems);
blanchet@54440
   536
panny@54857
   537
    val matchedsss' = AList.delete (op =) fun_name matchedsss
panny@55517
   538
      |> cons (fun_name, if seq then matchedss @ [prems] else matchedss @ [real_prems]);
panny@54791
   539
panny@54791
   540
    val user_eqn =
panny@55549
   541
      (real_prems, concl)
panny@55549
   542
      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
panny@55549
   543
      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
blanchet@54440
   544
  in
panny@54478
   545
    (Disc {
blanchet@54440
   546
      fun_name = fun_name,
panny@54857
   547
      fun_T = fun_T,
panny@54538
   548
      fun_args = fun_args,
panny@54857
   549
      ctr = ctr,
blanchet@54440
   550
      ctr_no = ctr_no,
panny@55612
   551
      disc = disc,
panny@54791
   552
      prems = real_prems,
panny@54959
   553
      auto_gen = catch_all,
panny@55549
   554
      maybe_ctr_rhs = maybe_ctr_rhs,
panny@55549
   555
      maybe_code_rhs = maybe_code_rhs,
panny@54791
   556
      user_eqn = user_eqn
panny@54857
   557
    }, matchedsss')
blanchet@54440
   558
  end;
blanchet@54440
   559
blanchet@55661
   560
fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
blanchet@55661
   561
    maybe_of_spec eqn =
blanchet@54440
   562
  let
blanchet@54440
   563
    val (lhs, rhs) = HOLogic.dest_eq eqn
blanchet@54440
   564
      handle TERM _ =>
blanchet@55698
   565
        primcorec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
blanchet@54440
   566
    val sel = head_of lhs;
panny@54857
   567
    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
blanchet@54440
   568
      handle TERM _ =>
blanchet@55698
   569
        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
panny@55612
   570
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
blanchet@55724
   571
      handle Option.Option =>
blanchet@55724
   572
        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
panny@55612
   573
    val {ctr, ...} =
blanchet@55661
   574
      (case maybe_of_spec of
blanchet@55661
   575
        SOME of_spec => the (find_first (equal of_spec o #ctr) basic_ctr_specs)
blanchet@55661
   576
      | NONE => filter (exists (equal sel) o #sels) basic_ctr_specs |> the_single
blanchet@55698
   577
          handle List.Empty => primcorec_error_eqn "ambiguous selector - use \"of\"" eqn);
panny@54791
   578
    val user_eqn = drop_All eqn';
blanchet@54440
   579
  in
panny@54478
   580
    Sel {
blanchet@54440
   581
      fun_name = fun_name,
panny@54857
   582
      fun_T = fun_T,
panny@54538
   583
      fun_args = fun_args,
panny@55612
   584
      ctr = ctr,
panny@54478
   585
      sel = sel,
blanchet@54440
   586
      rhs_term = rhs,
panny@54791
   587
      user_eqn = user_eqn
blanchet@54440
   588
    }
blanchet@54440
   589
  end;
blanchet@54440
   590
panny@55612
   591
fun dissect_coeqn_ctr seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
panny@55612
   592
    maybe_code_rhs prems concl matchedsss =
blanchet@55047
   593
  let
panny@55517
   594
    val (lhs, rhs) = HOLogic.dest_eq concl;
panny@55549
   595
    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
blanchet@55661
   596
    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
blanchet@55526
   597
    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
panny@55612
   598
    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) basic_ctr_specs)
blanchet@55698
   599
      handle Option.Option => primcorec_error_eqn "not a constructor" ctr;
panny@54478
   600
panny@55517
   601
    val disc_concl = betapply (disc, lhs);
panny@55612
   602
    val (maybe_eqn_data_disc, matchedsss') = if length basic_ctr_specs = 1
panny@54857
   603
      then (NONE, matchedsss)
panny@55612
   604
      else apfst SOME (dissect_coeqn_disc seq fun_names basic_ctr_specss
panny@55549
   605
          (SOME (abstract (List.rev fun_args) rhs)) maybe_code_rhs prems disc_concl matchedsss);
blanchet@54440
   606
blanchet@55611
   607
    val sel_concls = sels ~~ ctr_args
blanchet@54440
   608
      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
blanchet@54440
   609
blanchet@54993
   610
(*
panny@55517
   611
val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
panny@55517
   612
 (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
panny@55549
   613
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
panny@55549
   614
 "\nfor premise(s)\n    \<cdot> " ^
panny@55549
   615
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
blanchet@54993
   616
*)
blanchet@54440
   617
panny@55612
   618
    val eqns_data_sel =
panny@55612
   619
      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn' (SOME ctr)) sel_concls;
blanchet@54440
   620
  in
panny@54857
   621
    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
blanchet@54440
   622
  end;
blanchet@54440
   623
panny@55612
   624
fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss =
panny@55517
   625
  let
panny@55517
   626
    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
panny@55549
   627
    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
blanchet@55661
   628
    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
panny@55517
   629
panny@55517
   630
    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
panny@55612
   631
        if member ((op =) o apsnd #ctr) basic_ctr_specs ctr
panny@55517
   632
        then cons (ctr, cs)
blanchet@55698
   633
        else primcorec_error_eqn "not a constructor" ctr) [] rhs' []
panny@55517
   634
      |> AList.group (op =);
panny@55517
   635
blanchet@55520
   636
    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
panny@55517
   637
    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
panny@55517
   638
        binder_types (fastype_of ctr)
panny@55517
   639
        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
panny@55517
   640
          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
panny@55517
   641
        |> curry list_comb ctr
panny@55517
   642
        |> curry HOLogic.mk_eq lhs);
panny@55517
   643
  in
panny@55612
   644
    fold_map2 (dissect_coeqn_ctr false fun_names basic_ctr_specss eqn'
panny@55549
   645
        (SOME (abstract (List.rev fun_args) rhs)))
panny@55549
   646
      ctr_premss ctr_concls matchedsss
panny@55517
   647
  end;
panny@55517
   648
panny@55612
   649
fun dissect_coeqn lthy seq has_call fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
blanchet@55661
   650
    eqn' maybe_of_spec matchedsss =
blanchet@54440
   651
  let
panny@54791
   652
    val eqn = drop_All eqn'
blanchet@55698
   653
      handle TERM _ => primcorec_error_eqn "malformed function equation" eqn';
panny@55517
   654
    val (prems, concl) = Logic.strip_horn eqn
panny@54478
   655
      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
blanchet@54440
   656
panny@55517
   657
    val head = concl
blanchet@54440
   658
      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
blanchet@54440
   659
      |> head_of;
blanchet@54440
   660
blanchet@55661
   661
    val maybe_rhs = concl |> perhaps (try HOLogic.dest_not) |> try (snd o HOLogic.dest_eq);
blanchet@54440
   662
panny@55612
   663
    val discs = maps (map #disc) basic_ctr_specss;
panny@55612
   664
    val sels = maps (maps #sels) basic_ctr_specss;
panny@55612
   665
    val ctrs = maps (map #ctr) basic_ctr_specss;
blanchet@54440
   666
  in
blanchet@54440
   667
    if member (op =) discs head orelse
blanchet@54440
   668
      is_some maybe_rhs andalso
blanchet@54440
   669
        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
panny@55612
   670
      dissect_coeqn_disc seq fun_names basic_ctr_specss NONE NONE prems concl matchedsss
blanchet@54440
   671
      |>> single
blanchet@54440
   672
    else if member (op =) sels head then
blanchet@55661
   673
      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn' maybe_of_spec concl], matchedsss)
panny@55517
   674
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
blanchet@55526
   675
      member (op =) ctrs (head_of (unfold_let (the maybe_rhs))) then
panny@55612
   676
      dissect_coeqn_ctr seq fun_names basic_ctr_specss eqn' NONE prems concl matchedsss
panny@55517
   677
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
panny@55517
   678
      null prems then
panny@55612
   679
      dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss
panny@55517
   680
      |>> flat
blanchet@54440
   681
    else
blanchet@55698
   682
      primcorec_error_eqn "malformed function equation" eqn
blanchet@54440
   683
  end;
blanchet@54440
   684
blanchet@55139
   685
fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
blanchet@55605
   686
    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
panny@54791
   687
  if is_none (#pred (nth ctr_specs ctr_no)) then I else
blanchet@55520
   688
    s_conjs prems
panny@54791
   689
    |> curry subst_bounds (List.rev fun_args)
panny@54791
   690
    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
panny@54791
   691
    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
blanchet@54440
   692
blanchet@55605
   693
fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
panny@54857
   694
  find_first (equal sel o #sel) sel_eqns
panny@54857
   695
  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
panny@54857
   696
  |> the_default undef_const
panny@54548
   697
  |> K;
panny@54497
   698
blanchet@55605
   699
fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
blanchet@55660
   700
  (case find_first (equal sel o #sel) sel_eqns of
blanchet@55660
   701
    NONE => (I, I, I)
blanchet@55660
   702
  | SOME {fun_args, rhs_term, ... } =>
panny@55013
   703
    let
panny@55549
   704
      val bound_Ts = List.rev (map fastype_of fun_args);
blanchet@55659
   705
      fun rewrite_stop _ t = if has_call t then @{term False} else @{term True};
blanchet@55659
   706
      fun rewrite_end _ t = if has_call t then undef_const else t;
blanchet@55659
   707
      fun rewrite_cont bound_Ts t =
panny@55036
   708
        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
blanchet@55554
   709
      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
panny@55549
   710
        |> abs_tuple fun_args;
panny@55013
   711
    in
blanchet@55659
   712
      (massage rewrite_stop, massage rewrite_end, massage rewrite_cont)
blanchet@55660
   713
    end);
panny@54497
   714
blanchet@55605
   715
fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
blanchet@55660
   716
  (case find_first (equal sel o #sel) sel_eqns of
blanchet@55660
   717
    NONE => I
blanchet@55660
   718
  | SOME {fun_args, rhs_term, ...} =>
panny@55036
   719
    let
panny@55549
   720
      val bound_Ts = List.rev (map fastype_of fun_args);
panny@55036
   721
      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
panny@55036
   722
        | rewrite bound_Ts U T (t as _ $ _) =
panny@55036
   723
          let val (u, vs) = strip_comb t in
panny@55036
   724
            if is_Free u andalso has_call u then
panny@55036
   725
              Inr_const U T $ mk_tuple1 bound_Ts vs
blanchet@55723
   726
            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
panny@55036
   727
              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
panny@55036
   728
            else
panny@55036
   729
              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
panny@55036
   730
          end
panny@55036
   731
        | rewrite _ U T t =
panny@55036
   732
          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
panny@55036
   733
      fun massage t =
panny@55549
   734
        rhs_term
blanchet@55554
   735
        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
panny@54872
   736
        |> abs_tuple fun_args;
panny@55036
   737
    in
panny@55036
   738
      massage
blanchet@55660
   739
    end);
panny@54497
   740
blanchet@55605
   741
fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
blanchet@55139
   742
    (ctr_spec : corec_ctr_spec) =
blanchet@55660
   743
  (case filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns of
blanchet@55660
   744
    [] => I
blanchet@55660
   745
  | sel_eqns =>
blanchet@55660
   746
    let
blanchet@55660
   747
      val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
blanchet@55660
   748
      val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
blanchet@55660
   749
      val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
blanchet@55660
   750
      val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
blanchet@55660
   751
    in
blanchet@55660
   752
      I
blanchet@55660
   753
      #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
blanchet@55660
   754
      #> fold (fn (sel, (q, g, h)) =>
blanchet@55660
   755
        let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
blanchet@55660
   756
          nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
blanchet@55660
   757
      #> fold (fn (sel, n) => nth_map n
blanchet@55660
   758
        (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
blanchet@55660
   759
    end);
blanchet@54440
   760
blanchet@55605
   761
fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
blanchet@55605
   762
    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
panny@54791
   763
  let
panny@55612
   764
    val corecs = map #corec corec_specs;
panny@55612
   765
    val ctr_specss = map #ctr_specs corec_specs;
panny@54497
   766
    val corec_args = hd corecs
panny@54497
   767
      |> fst o split_last o binder_types o fastype_of
panny@54497
   768
      |> map (Const o pair @{const_name undefined})
panny@54857
   769
      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
panny@54497
   770
      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
panny@54872
   771
    fun currys [] t = t
panny@54872
   772
      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
panny@54872
   773
          |> fold_rev (Term.abs o pair Name.uu) Ts;
panny@54538
   774
blanchet@54993
   775
(*
panny@54497
   776
val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
panny@54548
   777
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
blanchet@54993
   778
*)
blanchet@54440
   779
panny@54791
   780
    val exclss' =
panny@54857
   781
      disc_eqnss
panny@54959
   782
      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
panny@54791
   783
        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
panny@54791
   784
        #> maps (uncurry (map o pair)
panny@54959
   785
          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
blanchet@55520
   786
              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
panny@54791
   787
            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
panny@54791
   788
            ||> Logic.list_implies
panny@54791
   789
            ||> curry Logic.list_all (map dest_Free fun_args))))
blanchet@54440
   790
  in
blanchet@54440
   791
    map (list_comb o rpair corec_args) corecs
blanchet@54440
   792
    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
blanchet@54440
   793
    |> map2 currys arg_Tss
blanchet@54440
   794
    |> Syntax.check_terms lthy
blanchet@55607
   795
    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
blanchet@55607
   796
      bs mxs
panny@54791
   797
    |> rpair exclss'
blanchet@54440
   798
  end;
blanchet@54440
   799
blanchet@55139
   800
fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
blanchet@55605
   801
    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
panny@54857
   802
  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
panny@54857
   803
    let
panny@54857
   804
      val n = 0 upto length ctr_specs
panny@54857
   805
        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
panny@54859
   806
      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
panny@54859
   807
        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
panny@54857
   808
      val extra_disc_eqn = {
panny@54857
   809
        fun_name = Binding.name_of fun_binding,
panny@54857
   810
        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
panny@54859
   811
        fun_args = fun_args,
panny@54857
   812
        ctr = #ctr (nth ctr_specs n),
panny@54857
   813
        ctr_no = n,
panny@54857
   814
        disc = #disc (nth ctr_specs n),
blanchet@55519
   815
        prems = maps (s_not_conj o #prems) disc_eqns,
panny@54959
   816
        auto_gen = true,
panny@55549
   817
        maybe_ctr_rhs = NONE,
panny@55549
   818
        maybe_code_rhs = NONE,
panny@54857
   819
        user_eqn = undef_const};
panny@54857
   820
    in
panny@54857
   821
      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
panny@54857
   822
    end;
panny@54857
   823
blanchet@55695
   824
fun find_corec_calls ctxt has_call basic_ctr_specs ({ctr, sel, rhs_term, ...} : coeqn_data_sel) =
panny@55612
   825
  let
panny@55612
   826
    val sel_no = find_first (equal ctr o #ctr) basic_ctr_specs
panny@55612
   827
      |> find_index (equal sel) o #sels o the;
blanchet@55695
   828
    fun find t = if has_call t then snd (fold_rev_corec_call ctxt (K cons) [] t []) else [];
panny@55612
   829
  in
panny@55612
   830
    find rhs_term
panny@55612
   831
    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
panny@55612
   832
  end;
panny@55612
   833
panny@55964
   834
fun add_primcorec_ursive maybe_tac opts fixes specs maybe_of_specs lthy =
blanchet@54440
   835
  let
blanchet@55724
   836
    val thy = Proof_Context.theory_of lthy;
blanchet@55724
   837
traytel@54489
   838
    val (bs, mxs) = map_split (apfst fst) fixes;
blanchet@54440
   839
    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
blanchet@54440
   840
blanchet@55724
   841
    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ arg_Ts) of
blanchet@55724
   842
        [] => ()
blanchet@55724
   843
      | (b, _) :: _ => primcorec_error ("type of " ^ Binding.print b ^ " contains top sort"));
blanchet@55724
   844
panny@55964
   845
    val seq = member (op =) opts Option_Sequential;
panny@55964
   846
    val exhaustive = member (op =) opts Option_Exhaustive;
panny@55964
   847
panny@55612
   848
    val fun_names = map Binding.name_of bs;
panny@55612
   849
    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
panny@55612
   850
    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
panny@55612
   851
    val eqns_data =
panny@55612
   852
      fold_map2 (dissect_coeqn lthy seq has_call fun_names basic_ctr_specss) (map snd specs)
blanchet@55661
   853
        maybe_of_specs []
panny@55612
   854
      |> flat o fst;
panny@55612
   855
panny@55612
   856
    val callssss =
panny@55612
   857
      map_filter (try (fn Sel x => x)) eqns_data
panny@55612
   858
      |> partition_eq ((op =) o pairself #fun_name)
panny@55612
   859
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
panny@55613
   860
      |> map (flat o snd)
blanchet@55695
   861
      |> map2 (fold o find_corec_calls lthy has_call) basic_ctr_specss
panny@55612
   862
      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
panny@55612
   863
        (ctr, map (K []) sels))) basic_ctr_specss);
panny@55612
   864
panny@55612
   865
(*
panny@55612
   866
val _ = tracing ("callssss = " ^ @{make_string} callssss);
panny@55612
   867
*)
blanchet@54440
   868
blanchet@54967
   869
    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
blanchet@54934
   870
          strong_coinduct_thms), lthy') =
blanchet@54931
   871
      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
blanchet@54967
   872
    val actual_nn = length bs;
blanchet@54967
   873
    val corec_specs = take actual_nn corec_specs'; (*###*)
blanchet@55630
   874
    val ctr_specss = map #ctr_specs corec_specs;
blanchet@54440
   875
panny@54857
   876
    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
panny@54791
   877
      |> partition_eq ((op =) o pairself #fun_name)
panny@54857
   878
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
panny@54791
   879
      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
panny@54857
   880
    val _ = disc_eqnss' |> map (fn x =>
panny@54857
   881
      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
blanchet@55698
   882
        primcorec_error_eqns "excess discriminator formula in definition"
panny@54857
   883
          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
panny@54791
   884
panny@54791
   885
    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
panny@54791
   886
      |> partition_eq ((op =) o pairself #fun_name)
panny@54857
   887
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
panny@54791
   888
      |> map (flat o snd);
panny@54791
   889
panny@54497
   890
    val arg_Tss = map (binder_types o snd o fst) fixes;
panny@54859
   891
    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
panny@54791
   892
    val (defs, exclss') =
blanchet@55605
   893
      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
panny@54791
   894
blanchet@55060
   895
    fun excl_tac (c, c', a) =
blanchet@55629
   896
      if a orelse c = c' orelse seq then SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
blanchet@55629
   897
      else maybe_tac;
panny@54959
   898
blanchet@54993
   899
(*
panny@54959
   900
val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
panny@54959
   901
 space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
blanchet@54993
   902
*)
panny@54959
   903
panny@54959
   904
    val exclss'' = exclss' |> map (map (fn (idx, t) =>
blanchet@55629
   905
      (idx, (Option.map (Goal.prove lthy [] [] t #> Thm.close_derivation) (excl_tac idx), t))));
panny@54791
   906
    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
blanchet@55632
   907
    val (goal_idxss, goalss) = exclss''
panny@54791
   908
      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
panny@54791
   909
      |> split_list o map split_list;
panny@54791
   910
panny@55964
   911
    val exhaustive_props = map (mk_disjs o map (mk_conjs o #prems)) disc_eqnss;
panny@55964
   912
panny@54791
   913
    fun prove thmss' def_thms' lthy =
panny@54791
   914
      let
panny@54791
   915
        val def_thms = map (snd o snd) def_thms';
panny@54791
   916
blanchet@55632
   917
        val exclss' = map (op ~~) (goal_idxss ~~ thmss');
panny@54791
   918
        fun mk_exclsss excls n =
panny@54791
   919
          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
panny@54959
   920
          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
panny@54791
   921
        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
panny@54791
   922
          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
panny@54791
   923
blanchet@55139
   924
        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
blanchet@55605
   925
            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
blanchet@55724
   926
          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then
blanchet@55724
   927
            []
blanchet@55724
   928
          else
panny@54857
   929
            let
panny@54859
   930
              val {disc_corec, ...} = nth ctr_specs ctr_no;
panny@54857
   931
              val k = 1 + ctr_no;
panny@54857
   932
              val m = length prems;
panny@54857
   933
              val t =
panny@54857
   934
                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
panny@54857
   935
                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
panny@54857
   936
                |> HOLogic.mk_Trueprop
panny@54857
   937
                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
panny@54857
   938
                |> curry Logic.list_all (map dest_Free fun_args);
panny@54857
   939
            in
panny@55549
   940
              if prems = [@{term False}] then [] else
panny@54857
   941
              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
panny@54857
   942
              |> K |> Goal.prove lthy [] [] t
blanchet@55628
   943
              |> Thm.close_derivation
panny@54857
   944
              |> pair (#disc (nth ctr_specs ctr_no))
panny@54857
   945
              |> single
panny@54857
   946
            end;
panny@54857
   947
blanchet@55731
   948
        fun prove_sel ({nested_map_idents, nested_map_comps, ctr_specs, ...} : corec_spec)
blanchet@55731
   949
            (disc_eqns : coeqn_data_disc list) exclsss
blanchet@55605
   950
            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : coeqn_data_sel) =
panny@54791
   951
          let
blanchet@55046
   952
            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
panny@54857
   953
            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
blanchet@55519
   954
            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
panny@54857
   955
                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
panny@54857
   956
            val sel_corec = find_index (equal sel) (#sels ctr_spec)
panny@54857
   957
              |> nth (#sel_corecs ctr_spec);
panny@54791
   958
            val k = 1 + ctr_no;
panny@54791
   959
            val m = length prems;
panny@54791
   960
            val t =
panny@54857
   961
              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
panny@54857
   962
              |> curry betapply sel
panny@54857
   963
              |> rpair (abstract (List.rev fun_args) rhs_term)
panny@54857
   964
              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
panny@54791
   965
              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
panny@54857
   966
              |> curry Logic.list_all (map dest_Free fun_args);
blanchet@55062
   967
            val (distincts, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
panny@54791
   968
          in
blanchet@55731
   969
            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_map_idents
blanchet@55731
   970
              nested_map_comps sel_corec k m exclsss
panny@54791
   971
            |> K |> Goal.prove lthy [] [] t
blanchet@55628
   972
            |> Thm.close_derivation
panny@54857
   973
            |> pair sel
panny@54791
   974
          end;
panny@54791
   975
blanchet@55605
   976
        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
blanchet@55605
   977
            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
panny@55549
   978
          (* don't try to prove theorems when some sel_eqns are missing *)
panny@54857
   979
          if not (exists (equal ctr o #ctr) disc_eqns)
panny@54859
   980
              andalso not (exists (equal ctr o #ctr) sel_eqns)
panny@55549
   981
            orelse
panny@54857
   982
              filter (equal ctr o #ctr) sel_eqns
panny@54857
   983
              |> fst o finds ((op =) o apsnd #sel) sels
panny@54857
   984
              |> exists (null o snd)
panny@54857
   985
          then [] else
panny@54857
   986
            let
panny@55549
   987
              val (fun_name, fun_T, fun_args, prems, maybe_rhs) =
panny@54859
   988
                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
panny@55549
   989
                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
panny@55549
   990
                  #maybe_ctr_rhs x))
panny@55549
   991
                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], NONE))
panny@54859
   992
                |> the o merge_options;
panny@54857
   993
              val m = length prems;
panny@55549
   994
              val t = (if is_some maybe_rhs then the maybe_rhs else
panny@55549
   995
                  filter (equal ctr o #ctr) sel_eqns
panny@55549
   996
                  |> fst o finds ((op =) o apsnd #sel) sels
panny@55549
   997
                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
panny@55549
   998
                  |> curry list_comb ctr)
panny@54857
   999
                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
panny@54857
  1000
                  map Bound (length fun_args - 1 downto 0)))
panny@54857
  1001
                |> HOLogic.mk_Trueprop
panny@54857
  1002
                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
panny@54857
  1003
                |> curry Logic.list_all (map dest_Free fun_args);
blanchet@54928
  1004
              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
blanchet@54928
  1005
              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
panny@54857
  1006
            in
panny@55549
  1007
              if prems = [@{term False}] then [] else
panny@55549
  1008
                mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
panny@55549
  1009
                |> K |> Goal.prove lthy [] [] t
blanchet@55628
  1010
                |> Thm.close_derivation
panny@55549
  1011
                |> pair ctr
panny@55549
  1012
                |> single
panny@55013
  1013
            end;
panny@54857
  1014
blanchet@55630
  1015
        fun prove_code disc_eqns sel_eqns ctr_alist ctr_specs =
panny@55550
  1016
          let
panny@55964
  1017
            val fun_data =
panny@55549
  1018
              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
panny@55549
  1019
               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
panny@55549
  1020
              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #maybe_code_rhs x))
panny@55549
  1021
              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, NONE))
panny@55964
  1022
              |> merge_options;
panny@55964
  1023
          in
panny@55964
  1024
            (case fun_data of
panny@55964
  1025
              NONE => []
panny@55964
  1026
            | SOME (fun_name, fun_T, fun_args, maybe_rhs) =>
panny@55964
  1027
              let
panny@55964
  1028
                val bound_Ts = List.rev (map fastype_of fun_args);
panny@55549
  1029
panny@55964
  1030
                val lhs = list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
panny@55964
  1031
                val maybe_rhs_info =
panny@55964
  1032
                  (case maybe_rhs of
panny@55964
  1033
                    SOME rhs =>
panny@55964
  1034
                    let
panny@55964
  1035
                      val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
panny@55964
  1036
                      val cond_ctrs =
panny@55964
  1037
                        fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
panny@55964
  1038
                      val ctr_thms = map (the o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
panny@55964
  1039
                    in SOME (rhs, raw_rhs, ctr_thms) end
panny@55964
  1040
                  | NONE =>
panny@55964
  1041
                    let
panny@55964
  1042
                      fun prove_code_ctr {ctr, sels, ...} =
panny@55964
  1043
                        if not (exists (equal ctr o fst) ctr_alist) then NONE else
panny@55964
  1044
                          let
panny@55964
  1045
                            val prems = find_first (equal ctr o #ctr) disc_eqns
panny@55964
  1046
                              |> Option.map #prems |> the_default [];
panny@55964
  1047
                            val t =
panny@55964
  1048
                              filter (equal ctr o #ctr) sel_eqns
panny@55964
  1049
                              |> fst o finds ((op =) o apsnd #sel) sels
panny@55964
  1050
                              |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
panny@55964
  1051
                                #-> abstract)
panny@55964
  1052
                              |> curry list_comb ctr;
panny@55964
  1053
                          in
panny@55964
  1054
                            SOME (prems, t)
panny@55964
  1055
                          end;
panny@55964
  1056
                      val maybe_ctr_conds_argss = map prove_code_ctr ctr_specs;
panny@55964
  1057
                    in
panny@55964
  1058
                      if exists is_none maybe_ctr_conds_argss then NONE else
panny@55964
  1059
                        let
panny@55964
  1060
                          val rhs = fold_rev (fn SOME (prems, u) => fn t => mk_If (s_conjs prems) u t)
panny@55964
  1061
                            maybe_ctr_conds_argss
panny@55964
  1062
                            (Const (@{const_name Code.abort}, @{typ String.literal} -->
panny@55964
  1063
                                (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
panny@55964
  1064
                              HOLogic.mk_literal fun_name $
panny@55964
  1065
                              absdummy @{typ unit} (incr_boundvars 1 lhs));
panny@55964
  1066
                        in SOME (rhs, rhs, map snd ctr_alist) end
panny@55964
  1067
                    end);
panny@55964
  1068
              in
panny@55964
  1069
                (case maybe_rhs_info of
panny@55964
  1070
                  NONE => []
panny@55964
  1071
                | SOME (rhs, raw_rhs, ctr_thms) =>
panny@55964
  1072
                  let
panny@55964
  1073
                    val ms = map (Logic.count_prems o prop_of) ctr_thms;
panny@55964
  1074
                    val (raw_t, t) = (raw_rhs, rhs)
panny@55964
  1075
                      |> pairself
panny@55964
  1076
                        (curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
panny@55964
  1077
                          map Bound (length fun_args - 1 downto 0)))
panny@55964
  1078
                        #> HOLogic.mk_Trueprop
panny@55964
  1079
                        #> curry Logic.list_all (map dest_Free fun_args));
panny@55964
  1080
                    val (distincts, discIs, sel_splits, sel_split_asms) =
panny@55964
  1081
                      case_thms_of_term lthy bound_Ts raw_rhs;
blanchet@55625
  1082
panny@55964
  1083
                    val raw_code_thm = mk_primcorec_raw_code_of_ctr_tac lthy distincts discIs sel_splits
panny@55964
  1084
                        sel_split_asms ms ctr_thms
panny@55964
  1085
                      |> K |> Goal.prove lthy [] [] raw_t
panny@55964
  1086
                      |> Thm.close_derivation;
panny@55964
  1087
                  in
panny@55964
  1088
                    mk_primcorec_code_of_raw_code_tac lthy distincts sel_splits raw_code_thm
panny@55964
  1089
                    |> K |> Goal.prove lthy [] [] t
panny@55964
  1090
                    |> Thm.close_derivation
panny@55964
  1091
                    |> single
panny@55964
  1092
                  end)
blanchet@55625
  1093
              end)
blanchet@55625
  1094
          end;
panny@55549
  1095
blanchet@54928
  1096
        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
blanchet@54928
  1097
        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
blanchet@54928
  1098
        val disc_thmss = map (map snd) disc_alists;
blanchet@54928
  1099
        val sel_thmss = map (map snd) sel_alists;
panny@55549
  1100
panny@55549
  1101
        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
blanchet@55630
  1102
          ctr_specss;
panny@55549
  1103
        val ctr_thmss = map (map snd) ctr_alists;
panny@55549
  1104
blanchet@55630
  1105
        val code_thmss = map4 prove_code disc_eqnss sel_eqnss ctr_alists ctr_specss;
blanchet@54928
  1106
blanchet@55167
  1107
        val simp_thmss = map2 append disc_thmss sel_thmss
blanchet@54932
  1108
blanchet@54934
  1109
        val common_name = mk_common_name fun_names;
blanchet@54934
  1110
blanchet@54928
  1111
        val notes =
blanchet@54967
  1112
          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
blanchet@55597
  1113
           (codeN, code_thmss, code_nitpicksimp_attrs),
blanchet@54934
  1114
           (ctrN, ctr_thmss, []),
blanchet@54928
  1115
           (discN, disc_thmss, simp_attrs),
blanchet@54932
  1116
           (selN, sel_thmss, simp_attrs),
blanchet@54934
  1117
           (simpsN, simp_thmss, []),
blanchet@54967
  1118
           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
blanchet@54928
  1119
          |> maps (fn (thmN, thmss, attrs) =>
blanchet@54928
  1120
            map2 (fn fun_name => fn thms =>
blanchet@54928
  1121
                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
blanchet@54967
  1122
              fun_names (take actual_nn thmss))
blanchet@54928
  1123
          |> filter_out (null o fst o hd o snd);
blanchet@54934
  1124
blanchet@54934
  1125
        val common_notes =
blanchet@54967
  1126
          [(coinductN, if n2m then [coinduct_thm] else [], []),
blanchet@54967
  1127
           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
blanchet@54934
  1128
          |> filter_out (null o #2)
blanchet@54934
  1129
          |> map (fn (thmN, thms, attrs) =>
blanchet@54934
  1130
            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
panny@54791
  1131
      in
blanchet@55167
  1132
        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
panny@54791
  1133
      end;
panny@54959
  1134
panny@54959
  1135
    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
blanchet@54440
  1136
  in
blanchet@55629
  1137
    (goalss, after_qed, lthy')
panny@54959
  1138
  end;
blanchet@54440
  1139
panny@55964
  1140
fun add_primcorec_ursive_cmd maybe_tac opts (raw_fixes, raw_specs') lthy =
blanchet@54440
  1141
  let
blanchet@55661
  1142
    val (raw_specs, maybe_of_specs) =
blanchet@55661
  1143
      split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
panny@54968
  1144
    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
blanchet@54440
  1145
  in
panny@55964
  1146
    add_primcorec_ursive maybe_tac opts fixes specs maybe_of_specs lthy
blanchet@55698
  1147
    handle ERROR str => primcorec_error str
blanchet@54440
  1148
  end
blanchet@55698
  1149
  handle Primcorec_Error (str, eqns) =>
blanchet@54440
  1150
    if null eqns
blanchet@54440
  1151
    then error ("primcorec error:\n  " ^ str)
blanchet@54440
  1152
    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
panny@54959
  1153
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
panny@54959
  1154
blanchet@55629
  1155
val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
blanchet@55629
  1156
  lthy
blanchet@55629
  1157
  |> Proof.theorem NONE after_qed goalss
blanchet@55629
  1158
  |> Proof.refine (Method.primitive_text I)
blanchet@55629
  1159
  |> Seq.hd) ooo add_primcorec_ursive_cmd NONE;
blanchet@55629
  1160
blanchet@55629
  1161
val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
blanchet@55629
  1162
  lthy
blanchet@55629
  1163
  |> after_qed (map (fn [] => []
blanchet@55698
  1164
      | _ => primcorec_error "need exclusiveness proofs - use primcorecursive instead of primcorec")
blanchet@55629
  1165
    goalss)) ooo add_primcorec_ursive_cmd (SOME (fn {context = ctxt, ...} => auto_tac ctxt));
blanchet@54440
  1166
blanchet@54440
  1167
end;