src/HOL/Codatatype/Tools/bnf_wrap.ML
author blanchet
Tue, 18 Sep 2012 11:42:22 +0200
changeset 50453 5bc80d96241e
parent 50452 c139da00fb4a
child 50473 9321a9465021
permissions -rw-r--r--
group "simps" together
blanchet@50089
     1
(*  Title:      HOL/Codatatype/Tools/bnf_wrap.ML
blanchet@50032
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@50032
     3
    Copyright   2012
blanchet@50032
     4
blanchet@50089
     5
Wrapping existing datatypes.
blanchet@50032
     6
*)
blanchet@50032
     7
blanchet@50089
     8
signature BNF_WRAP =
blanchet@50032
     9
sig
blanchet@50136
    10
  val mk_half_pairss: 'a list -> ('a * 'a) list list
blanchet@50218
    11
  val mk_ctr: typ list -> term -> term
blanchet@50452
    12
  val base_name_of_ctr: term -> string
blanchet@50214
    13
  val wrap_datatype: ({prems: thm list, context: Proof.context} -> tactic) list list ->
blanchet@50295
    14
    ((bool * term list) * term) *
blanchet@50295
    15
      (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
blanchet@50453
    16
    (term list list * thm list * thm list * thm list * thm list * thm list list) * local_theory
blanchet@50293
    17
  val parse_wrap_options: bool parser
blanchet@50301
    18
  val parse_bound_term: (binding * string) parser
blanchet@50032
    19
end;
blanchet@50032
    20
blanchet@50089
    21
structure BNF_Wrap : BNF_WRAP =
blanchet@50032
    22
struct
blanchet@50032
    23
blanchet@50032
    24
open BNF_Util
blanchet@50089
    25
open BNF_Wrap_Tactics
blanchet@50032
    26
blanchet@50238
    27
val isN = "is_";
blanchet@50238
    28
val unN = "un_";
blanchet@50238
    29
fun mk_unN 1 1 suf = unN ^ suf
blanchet@50238
    30
  | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
blanchet@50061
    31
blanchet@50069
    32
val case_congN = "case_cong";
blanchet@50131
    33
val case_eqN = "case_eq";
blanchet@50069
    34
val casesN = "cases";
blanchet@50133
    35
val collapseN = "collapse";
blanchet@50137
    36
val disc_excludeN = "disc_exclude";
blanchet@50069
    37
val disc_exhaustN = "disc_exhaust";
blanchet@50069
    38
val discsN = "discs";
blanchet@50069
    39
val distinctN = "distinct";
blanchet@50090
    40
val exhaustN = "exhaust";
blanchet@50090
    41
val injectN = "inject";
blanchet@50090
    42
val nchotomyN = "nchotomy";
blanchet@50069
    43
val selsN = "sels";
blanchet@50069
    44
val splitN = "split";
blanchet@50069
    45
val split_asmN = "split_asm";
blanchet@50069
    46
val weak_case_cong_thmsN = "weak_case_cong";
blanchet@50034
    47
blanchet@50351
    48
val std_binding = @{binding _};
blanchet@50315
    49
blanchet@50315
    50
val induct_simp_attrs = @{attributes [induct_simp]};
blanchet@50315
    51
val cong_attrs = @{attributes [cong]};
blanchet@50315
    52
val iff_attrs = @{attributes [iff]};
blanchet@50315
    53
val safe_elim_attrs = @{attributes [elim!]};
blanchet@50315
    54
val simp_attrs = @{attributes [simp]};
blanchet@50061
    55
blanchet@50071
    56
fun pad_list x n xs = xs @ replicate (n - length xs) x;
blanchet@50071
    57
blanchet@50274
    58
fun unflat_lookup eq ys zs = map (map (fn x => nth zs (find_index (curry eq x) ys)));
blanchet@50273
    59
blanchet@50063
    60
fun mk_half_pairss' _ [] = []
blanchet@50071
    61
  | mk_half_pairss' indent (y :: ys) =
blanchet@50071
    62
    indent @ fold_rev (cons o single o pair y) ys (mk_half_pairss' ([] :: indent) ys);
blanchet@50042
    63
blanchet@50063
    64
fun mk_half_pairss ys = mk_half_pairss' [[]] ys;
blanchet@50042
    65
blanchet@50295
    66
fun mk_undefined T = Const (@{const_name undefined}, T);
blanchet@50070
    67
blanchet@50218
    68
fun mk_ctr Ts ctr =
blanchet@50218
    69
  let val Type (_, Ts0) = body_type (fastype_of ctr) in
blanchet@50218
    70
    Term.subst_atomic_types (Ts0 ~~ Ts) ctr
blanchet@50218
    71
  end;
blanchet@50218
    72
blanchet@50315
    73
fun base_name_of_ctr c =
blanchet@50315
    74
  Long_Name.base_name (case head_of c of
blanchet@50315
    75
      Const (s, _) => s
blanchet@50315
    76
    | Free (s, _) => s
blanchet@50315
    77
    | _ => error "Cannot extract name of constructor");
blanchet@50061
    78
blanchet@50452
    79
fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
blanchet@50452
    80
blanchet@50293
    81
fun prepare_wrap_datatype prep_term (((no_dests, raw_ctrs), raw_case),
blanchet@50351
    82
    (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
blanchet@50032
    83
  let
blanchet@50034
    84
    (* TODO: sanity checks on arguments *)
blanchet@50128
    85
    (* TODO: attributes (simp, case_names, etc.) *)
blanchet@50128
    86
    (* TODO: case syntax *)
blanchet@50128
    87
    (* TODO: integration with function package ("size") *)
blanchet@50032
    88
blanchet@50295
    89
    val n = length raw_ctrs;
blanchet@50069
    90
    val ks = 1 upto n;
blanchet@50069
    91
blanchet@50136
    92
    val _ = if n > 0 then () else error "No constructors specified";
blanchet@50136
    93
blanchet@50295
    94
    val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
blanchet@50295
    95
    val case0 = prep_term no_defs_lthy raw_case;
blanchet@50295
    96
    val sel_defaultss =
blanchet@50295
    97
      pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
blanchet@50295
    98
blanchet@50315
    99
    val Type (dataT_name, As0) = body_type (fastype_of (hd ctrs0));
blanchet@50317
   100
    val data_b = Binding.qualified_name dataT_name;
blanchet@50070
   101
blanchet@50070
   102
    val (As, B) =
blanchet@50070
   103
      no_defs_lthy
blanchet@50313
   104
      |> mk_TFrees' (map Type.sort_of_atyp As0)
blanchet@50070
   105
      ||> the_single o fst o mk_TFrees 1;
blanchet@50070
   106
blanchet@50315
   107
    val dataT = Type (dataT_name, As);
blanchet@50070
   108
    val ctrs = map (mk_ctr As) ctrs0;
blanchet@50070
   109
    val ctr_Tss = map (binder_types o fastype_of) ctrs;
blanchet@50070
   110
blanchet@50070
   111
    val ms = map length ctr_Tss;
blanchet@50070
   112
traytel@50449
   113
    val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
blanchet@50135
   114
blanchet@50189
   115
    fun can_really_rely_on_disc k =
traytel@50449
   116
      not (Binding.eq_name (nth raw_disc_bindings' (k - 1), Binding.empty)) orelse
traytel@50449
   117
      nth ms (k - 1) = 0;
blanchet@50167
   118
    fun can_rely_on_disc k =
blanchet@50189
   119
      can_really_rely_on_disc k orelse (k = 1 andalso not (can_really_rely_on_disc 2));
blanchet@50351
   120
    fun can_omit_disc_binding k m =
blanchet@50189
   121
      n = 1 orelse m = 0 orelse (n = 2 andalso can_rely_on_disc (3 - k));
blanchet@50135
   122
blanchet@50351
   123
    val std_disc_binding =
blanchet@50317
   124
      Binding.qualify false (Binding.name_of data_b) o Binding.name o prefix isN o base_name_of_ctr;
blanchet@50135
   125
blanchet@50351
   126
    val disc_bindings =
blanchet@50351
   127
      raw_disc_bindings'
blanchet@50135
   128
      |> map4 (fn k => fn m => fn ctr => fn disc =>
blanchet@50317
   129
        Option.map (Binding.qualify false (Binding.name_of data_b))
traytel@50449
   130
          (if Binding.eq_name (disc, Binding.empty) then
blanchet@50351
   131
             if can_omit_disc_binding k m then NONE else SOME (std_disc_binding ctr)
blanchet@50351
   132
           else if Binding.eq_name (disc, std_binding) then
blanchet@50351
   133
             SOME (std_disc_binding ctr)
blanchet@50317
   134
           else
blanchet@50317
   135
             SOME disc)) ks ms ctrs0;
blanchet@50071
   136
blanchet@50351
   137
    val no_discs = map is_none disc_bindings;
blanchet@50152
   138
    val no_discs_at_all = forall I no_discs;
blanchet@50131
   139
blanchet@50351
   140
    fun std_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
blanchet@50135
   141
blanchet@50351
   142
    val sel_bindingss =
blanchet@50351
   143
      pad_list [] n raw_sel_bindingss
blanchet@50071
   144
      |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
blanchet@50317
   145
        Binding.qualify false (Binding.name_of data_b)
traytel@50449
   146
          (if Binding.eq_name (sel, Binding.empty) orelse Binding.eq_name (sel, std_binding) then
blanchet@50351
   147
            std_sel_binding m l ctr
blanchet@50317
   148
          else
traytel@50449
   149
            sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
blanchet@50035
   150
blanchet@50145
   151
    fun mk_case Ts T =
blanchet@50136
   152
      let
blanchet@50351
   153
        val (bindings, body) = strip_type (fastype_of case0)
blanchet@50351
   154
        val Type (_, Ts0) = List.last bindings
blanchet@50145
   155
      in Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) case0 end;
blanchet@50037
   156
blanchet@50216
   157
    val casex = mk_case As B;
blanchet@50216
   158
    val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
blanchet@50058
   159
blanchet@50379
   160
    val ((((((((xss, xss'), yss), fs), gs), (v, v')), w), (p, p')), names_lthy) = no_defs_lthy |>
blanchet@50379
   161
      mk_Freess' "x" ctr_Tss
blanchet@50040
   162
      ||>> mk_Freess "y" ctr_Tss
blanchet@50216
   163
      ||>> mk_Frees "f" case_Ts
blanchet@50216
   164
      ||>> mk_Frees "g" case_Ts
blanchet@50315
   165
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "v") dataT
blanchet@50315
   166
      ||>> yield_singleton (mk_Frees "w") dataT
blanchet@50058
   167
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
blanchet@50058
   168
blanchet@50058
   169
    val q = Free (fst p', B --> HOLogic.boolT);
blanchet@50035
   170
blanchet@50225
   171
    fun ap_v t = t $ v;
blanchet@50152
   172
    fun mk_v_eq_v () = HOLogic.mk_eq (v, v);
blanchet@50152
   173
blanchet@50040
   174
    val xctrs = map2 (curry Term.list_comb) ctrs xss;
blanchet@50040
   175
    val yctrs = map2 (curry Term.list_comb) ctrs yss;
blanchet@50047
   176
blanchet@50058
   177
    val xfs = map2 (curry Term.list_comb) fs xss;
blanchet@50058
   178
    val xgs = map2 (curry Term.list_comb) gs xss;
blanchet@50058
   179
blanchet@50452
   180
    val eta_fs = map2 eta_expand_arg xss xfs;
blanchet@50452
   181
    val eta_gs = map2 eta_expand_arg xss xgs;
blanchet@50058
   182
blanchet@50216
   183
    val fcase = Term.list_comb (casex, eta_fs);
blanchet@50216
   184
    val gcase = Term.list_comb (casex, eta_gs);
blanchet@50035
   185
blanchet@50040
   186
    val exist_xs_v_eq_ctrs =
blanchet@50040
   187
      map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (v, xctr))) xctrs xss;
blanchet@50037
   188
blanchet@50293
   189
    val unique_disc_no_def = TrueI; (*arbitrary marker*)
blanchet@50293
   190
    val alternate_disc_no_def = FalseE; (*arbitrary marker*)
blanchet@50131
   191
blanchet@50293
   192
    fun alternate_disc_lhs get_disc k =
blanchet@50131
   193
      HOLogic.mk_not
blanchet@50351
   194
        (case nth disc_bindings (k - 1) of
blanchet@50167
   195
          NONE => nth exist_xs_v_eq_ctrs (k - 1)
blanchet@50293
   196
        | SOME b => get_disc b (k - 1) $ v);
blanchet@50131
   197
blanchet@50300
   198
    val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
blanchet@50293
   199
      if no_dests then
blanchet@50300
   200
        (true, [], [], [], [], [], no_defs_lthy)
blanchet@50293
   201
      else
blanchet@50293
   202
        let
blanchet@50315
   203
          fun disc_free b = Free (Binding.name_of b, dataT --> HOLogic.boolT);
blanchet@50040
   204
blanchet@50293
   205
          fun disc_spec b exist_xs_v_eq_ctr = mk_Trueprop_eq (disc_free b $ v, exist_xs_v_eq_ctr);
blanchet@50273
   206
blanchet@50293
   207
          fun alternate_disc k = Term.lambda v (alternate_disc_lhs (K o disc_free) (3 - k));
blanchet@50037
   208
blanchet@50379
   209
          fun mk_default T t =
blanchet@50379
   210
            let
blanchet@50379
   211
              val Ts0 = map TFree (Term.add_tfreesT (fastype_of t) []);
blanchet@50379
   212
              val Ts = map TFree (Term.add_tfreesT T []);
blanchet@50379
   213
            in Term.subst_atomic_types (Ts0 ~~ Ts) t end;
blanchet@50379
   214
blanchet@50295
   215
          fun mk_sel_case_args b proto_sels T =
blanchet@50295
   216
            map2 (fn Ts => fn k =>
blanchet@50295
   217
              (case AList.lookup (op =) proto_sels k of
blanchet@50295
   218
                NONE =>
blanchet@50379
   219
                (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
blanchet@50379
   220
                  NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
blanchet@50379
   221
                | SOME t => mk_default (Ts ---> T) t)
blanchet@50295
   222
              | SOME (xs, x) => fold_rev Term.lambda xs x)) ctr_Tss ks;
blanchet@50131
   223
blanchet@50293
   224
          fun sel_spec b proto_sels =
blanchet@50293
   225
            let
blanchet@50293
   226
              val _ =
blanchet@50293
   227
                (case duplicates (op =) (map fst proto_sels) of
blanchet@50293
   228
                   k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
blanchet@50293
   229
                     " for constructor " ^
blanchet@50293
   230
                     quote (Syntax.string_of_term no_defs_lthy (nth ctrs (k - 1))))
blanchet@50293
   231
                 | [] => ())
blanchet@50293
   232
              val T =
blanchet@50293
   233
                (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
blanchet@50293
   234
                  [T] => T
blanchet@50293
   235
                | T :: T' :: _ => error ("Inconsistent range type for selector " ^
blanchet@50293
   236
                    quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
blanchet@50293
   237
                    " vs. " ^ quote (Syntax.string_of_typ no_defs_lthy T')));
blanchet@50293
   238
            in
blanchet@50315
   239
              mk_Trueprop_eq (Free (Binding.name_of b, dataT --> T) $ v,
blanchet@50295
   240
                Term.list_comb (mk_case As T, mk_sel_case_args b proto_sels T) $ v)
blanchet@50293
   241
            end;
blanchet@50273
   242
blanchet@50351
   243
          val sel_bindings = flat sel_bindingss;
blanchet@50351
   244
          val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
blanchet@50351
   245
          val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
blanchet@50300
   246
blanchet@50351
   247
          val sel_binding_index =
blanchet@50351
   248
            if all_sels_distinct then 1 upto length sel_bindings
blanchet@50351
   249
            else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
blanchet@50300
   250
blanchet@50300
   251
          val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
blanchet@50315
   252
          val sel_infos =
blanchet@50351
   253
            AList.group (op =) (sel_binding_index ~~ proto_sels)
blanchet@50300
   254
            |> sort (int_ord o pairself fst)
blanchet@50351
   255
            |> map snd |> curry (op ~~) uniq_sel_bindings;
blanchet@50351
   256
          val sel_bindings = map fst sel_infos;
blanchet@50273
   257
blanchet@50351
   258
          fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss;
blanchet@50037
   259
blanchet@50293
   260
          val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) =
blanchet@50293
   261
            no_defs_lthy
blanchet@50293
   262
            |> apfst split_list o fold_map4 (fn k => fn m => fn exist_xs_v_eq_ctr =>
blanchet@50293
   263
              fn NONE =>
blanchet@50293
   264
                 if n = 1 then pair (Term.lambda v (mk_v_eq_v ()), unique_disc_no_def)
blanchet@50293
   265
                 else if m = 0 then pair (Term.lambda v exist_xs_v_eq_ctr, refl)
blanchet@50293
   266
                 else pair (alternate_disc k, alternate_disc_no_def)
blanchet@50293
   267
               | SOME b => Specification.definition (SOME (b, NONE, NoSyn),
blanchet@50293
   268
                   ((Thm.def_binding b, []), disc_spec b exist_xs_v_eq_ctr)) #>> apsnd snd)
blanchet@50351
   269
              ks ms exist_xs_v_eq_ctrs disc_bindings
blanchet@50293
   270
            ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
blanchet@50293
   271
              Specification.definition (SOME (b, NONE, NoSyn),
blanchet@50315
   272
                ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
blanchet@50293
   273
            ||> `Local_Theory.restore;
blanchet@50040
   274
blanchet@50293
   275
          val phi = Proof_Context.export_morphism lthy lthy';
blanchet@50040
   276
blanchet@50293
   277
          val disc_defs = map (Morphism.thm phi) raw_disc_defs;
blanchet@50296
   278
          val sel_defs = map (Morphism.thm phi) raw_sel_defs;
blanchet@50296
   279
          val sel_defss = unflat_selss sel_defs;
blanchet@50043
   280
blanchet@50293
   281
          val discs0 = map (Morphism.term phi) raw_discs;
blanchet@50293
   282
          val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
blanchet@50043
   283
blanchet@50293
   284
          fun mk_disc_or_sel Ts c =
blanchet@50293
   285
            Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of c))) ~~ Ts) c;
blanchet@50293
   286
blanchet@50293
   287
          val discs = map (mk_disc_or_sel As) discs0;
blanchet@50293
   288
          val selss = map (map (mk_disc_or_sel As)) selss0;
blanchet@50293
   289
        in
blanchet@50300
   290
          (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy')
blanchet@50293
   291
        end;
blanchet@50040
   292
blanchet@50047
   293
    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
blanchet@50044
   294
blanchet@50035
   295
    val goal_exhaust =
blanchet@50047
   296
      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (v, xctr)]) in
blanchet@50136
   297
        fold_rev Logic.all [p, v] (mk_imp_p (map2 mk_prem xctrs xss))
blanchet@50035
   298
      end;
blanchet@50032
   299
blanchet@50049
   300
    val goal_injectss =
blanchet@50032
   301
      let
blanchet@50049
   302
        fun mk_goal _ _ [] [] = []
blanchet@50040
   303
          | mk_goal xctr yctr xs ys =
blanchet@50136
   304
            [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr),
blanchet@50136
   305
              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
blanchet@50032
   306
      in
blanchet@50049
   307
        map4 mk_goal xctrs yctrs xss yss
blanchet@50032
   308
      end;
blanchet@50032
   309
blanchet@50063
   310
    val goal_half_distinctss =
blanchet@50136
   311
      let
blanchet@50218
   312
        fun mk_goal ((xs, xc), (xs', xc')) =
blanchet@50136
   313
          fold_rev Logic.all (xs @ xs')
blanchet@50218
   314
            (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
blanchet@50136
   315
      in
blanchet@50136
   316
        map (map mk_goal) (mk_half_pairss (xss ~~ xctrs))
blanchet@50136
   317
      end;
blanchet@50034
   318
blanchet@50136
   319
    val goal_cases =
blanchet@50136
   320
      map3 (fn xs => fn xctr => fn xf =>
blanchet@50216
   321
        fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xss xctrs xfs;
blanchet@50040
   322
blanchet@50136
   323
    val goalss = [goal_exhaust] :: goal_injectss @ goal_half_distinctss @ [goal_cases];
blanchet@50034
   324
blanchet@50034
   325
    fun after_qed thmss lthy =
blanchet@50034
   326
      let
blanchet@50063
   327
        val ([exhaust_thm], (inject_thmss, (half_distinct_thmss, [case_thms]))) =
blanchet@50063
   328
          (hd thmss, apsnd (chop (n * n)) (chop n (tl thmss)));
blanchet@50034
   329
blanchet@50453
   330
        val inject_thms = flat inject_thmss;
blanchet@50453
   331
blanchet@50047
   332
        val exhaust_thm' =
blanchet@50047
   333
          let val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As) in
blanchet@50047
   334
            Drule.instantiate' [] [SOME (certify lthy v)]
blanchet@50047
   335
              (Thm.instantiate (Tinst, []) (Drule.zero_var_indexes exhaust_thm))
blanchet@50047
   336
          end;
blanchet@50047
   337
blanchet@50315
   338
        val exhaust_cases = map base_name_of_ctr ctrs;
blanchet@50315
   339
blanchet@50063
   340
        val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss;
blanchet@50063
   341
blanchet@50067
   342
        val (distinct_thmsss', distinct_thmsss) =
blanchet@50063
   343
          map2 (map2 append) (Library.chop_groups n half_distinct_thmss)
blanchet@50067
   344
            (transpose (Library.chop_groups n other_half_distinct_thmss))
blanchet@50067
   345
          |> `transpose;
blanchet@50063
   346
        val distinct_thms = interleave (flat half_distinct_thmss) (flat other_half_distinct_thmss);
blanchet@50034
   347
blanchet@50035
   348
        val nchotomy_thm =
blanchet@50035
   349
          let
blanchet@50035
   350
            val goal =
blanchet@50037
   351
              HOLogic.mk_Trueprop (HOLogic.mk_all (fst v', snd v',
blanchet@50044
   352
                Library.foldr1 HOLogic.mk_disj exist_xs_v_eq_ctrs));
blanchet@50035
   353
          in
blanchet@50035
   354
            Skip_Proof.prove lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
blanchet@50035
   355
          end;
blanchet@50035
   356
blanchet@50296
   357
        val (all_sel_thms, sel_thmss, disc_thms, discI_thms, disc_exclude_thms, disc_exhaust_thms,
blanchet@50296
   358
             collapse_thms, case_eq_thms) =
blanchet@50293
   359
          if no_dests then
blanchet@50296
   360
            ([], [], [], [], [], [], [], [])
blanchet@50131
   361
          else
blanchet@50131
   362
            let
blanchet@50379
   363
              fun make_sel_thm xs' case_thm sel_def =
blanchet@50379
   364
                zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
blanchet@50379
   365
                    (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
blanchet@50296
   366
blanchet@50296
   367
              fun has_undefined_rhs thm =
blanchet@50296
   368
                (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
blanchet@50296
   369
                  Const (@{const_name undefined}, _) => true
blanchet@50296
   370
                | _ => false);
blanchet@50296
   371
blanchet@50379
   372
              val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
blanchet@50296
   373
blanchet@50296
   374
              val all_sel_thms =
blanchet@50300
   375
                (if all_sels_distinct andalso forall null sel_defaultss then
blanchet@50300
   376
                   flat sel_thmss
blanchet@50300
   377
                 else
blanchet@50379
   378
                   map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs
blanchet@50379
   379
                     (xss' ~~ case_thms))
blanchet@50300
   380
                |> filter_out has_undefined_rhs;
blanchet@50043
   381
blanchet@50293
   382
              fun mk_unique_disc_def () =
blanchet@50293
   383
                let
blanchet@50293
   384
                  val m = the_single ms;
blanchet@50293
   385
                  val goal = mk_Trueprop_eq (mk_v_eq_v (), the_single exist_xs_v_eq_ctrs);
blanchet@50293
   386
                in
blanchet@50293
   387
                  Skip_Proof.prove lthy [] [] goal (fn _ => mk_unique_disc_def_tac m exhaust_thm')
blanchet@50293
   388
                  |> singleton (Proof_Context.export names_lthy lthy)
blanchet@50293
   389
                  |> Thm.close_derivation
blanchet@50293
   390
                end;
blanchet@50043
   391
blanchet@50293
   392
              fun mk_alternate_disc_def k =
blanchet@50293
   393
                let
blanchet@50293
   394
                  val goal =
blanchet@50293
   395
                    mk_Trueprop_eq (alternate_disc_lhs (K (nth discs)) (3 - k),
blanchet@50293
   396
                      nth exist_xs_v_eq_ctrs (k - 1));
blanchet@50293
   397
                in
blanchet@50293
   398
                  Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
blanchet@50293
   399
                    mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
blanchet@50293
   400
                      (nth distinct_thms (2 - k)) exhaust_thm')
blanchet@50293
   401
                  |> singleton (Proof_Context.export names_lthy lthy)
blanchet@50293
   402
                  |> Thm.close_derivation
blanchet@50293
   403
                end;
blanchet@50043
   404
blanchet@50293
   405
              val has_alternate_disc_def =
blanchet@50293
   406
                exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
blanchet@50293
   407
blanchet@50293
   408
              val disc_defs' =
blanchet@50293
   409
                map2 (fn k => fn def =>
blanchet@50293
   410
                  if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def ()
blanchet@50293
   411
                  else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k
blanchet@50293
   412
                  else def) ks disc_defs;
blanchet@50293
   413
blanchet@50293
   414
              val discD_thms = map (fn def => def RS iffD1) disc_defs';
blanchet@50293
   415
              val discI_thms =
blanchet@50293
   416
                map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms
blanchet@50293
   417
                  disc_defs';
blanchet@50293
   418
              val not_discI_thms =
blanchet@50293
   419
                map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
blanchet@50293
   420
                    (Local_Defs.unfold lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
blanchet@50293
   421
                  ms disc_defs';
blanchet@50293
   422
blanchet@50293
   423
              val (disc_thmss', disc_thmss) =
blanchet@50293
   424
                let
blanchet@50293
   425
                  fun mk_thm discI _ [] = refl RS discI
blanchet@50293
   426
                    | mk_thm _ not_discI [distinct] = distinct RS not_discI;
blanchet@50293
   427
                  fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
blanchet@50293
   428
                in
blanchet@50293
   429
                  map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
blanchet@50293
   430
                end;
blanchet@50293
   431
blanchet@50293
   432
              val disc_thms = flat (map2 (fn true => K [] | false => I) no_discs disc_thmss);
blanchet@50293
   433
blanchet@50293
   434
              val disc_exclude_thms =
blanchet@50293
   435
                if has_alternate_disc_def then
blanchet@50293
   436
                  []
blanchet@50293
   437
                else
blanchet@50293
   438
                  let
blanchet@50293
   439
                    fun mk_goal [] = []
blanchet@50293
   440
                      | mk_goal [((_, true), (_, true))] = []
blanchet@50293
   441
                      | mk_goal [(((_, disc), _), ((_, disc'), _))] =
blanchet@50293
   442
                        [Logic.all v (Logic.mk_implies (HOLogic.mk_Trueprop (betapply (disc, v)),
blanchet@50293
   443
                           HOLogic.mk_Trueprop (HOLogic.mk_not (betapply (disc', v)))))];
blanchet@50293
   444
                    fun prove tac goal = Skip_Proof.prove lthy [] [] goal (K tac);
blanchet@50293
   445
blanchet@50315
   446
                    val infos = ms ~~ discD_thms ~~ discs ~~ no_discs;
blanchet@50315
   447
                    val half_pairss = mk_half_pairss infos;
blanchet@50293
   448
blanchet@50293
   449
                    val goal_halvess = map mk_goal half_pairss;
blanchet@50293
   450
                    val half_thmss =
blanchet@50293
   451
                      map3 (fn [] => K (K []) | [goal] => fn [((((m, discD), _), _), _)] =>
blanchet@50293
   452
                          fn disc_thm => [prove (mk_half_disc_exclude_tac m discD disc_thm) goal])
blanchet@50293
   453
                        goal_halvess half_pairss (flat disc_thmss');
blanchet@50293
   454
blanchet@50293
   455
                    val goal_other_halvess = map (mk_goal o map swap) half_pairss;
blanchet@50293
   456
                    val other_half_thmss =
blanchet@50293
   457
                      map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
blanchet@50293
   458
                        goal_other_halvess;
blanchet@50293
   459
                  in
blanchet@50293
   460
                    interleave (flat half_thmss) (flat other_half_thmss)
blanchet@50293
   461
                  end;
blanchet@50293
   462
blanchet@50293
   463
              val disc_exhaust_thms =
blanchet@50293
   464
                if has_alternate_disc_def orelse no_discs_at_all then
blanchet@50293
   465
                  []
blanchet@50293
   466
                else
blanchet@50293
   467
                  let
blanchet@50293
   468
                    fun mk_prem disc = mk_imp_p [HOLogic.mk_Trueprop (betapply (disc, v))];
blanchet@50293
   469
                    val goal = fold_rev Logic.all [p, v] (mk_imp_p (map mk_prem discs));
blanchet@50293
   470
                  in
blanchet@50293
   471
                    [Skip_Proof.prove lthy [] [] goal (fn _ =>
blanchet@50293
   472
                       mk_disc_exhaust_tac n exhaust_thm discI_thms)]
blanchet@50293
   473
                  end;
blanchet@50293
   474
blanchet@50293
   475
              val collapse_thms =
blanchet@50293
   476
                if no_dests then
blanchet@50293
   477
                  []
blanchet@50293
   478
                else
blanchet@50293
   479
                  let
blanchet@50293
   480
                    fun mk_goal ctr disc sels =
blanchet@50293
   481
                      let
blanchet@50293
   482
                        val prem = HOLogic.mk_Trueprop (betapply (disc, v));
blanchet@50293
   483
                        val concl =
blanchet@50293
   484
                          mk_Trueprop_eq ((null sels ? swap)
blanchet@50293
   485
                            (Term.list_comb (ctr, map ap_v sels), v));
blanchet@50293
   486
                      in
blanchet@50293
   487
                        if prem aconv concl then NONE
blanchet@50293
   488
                        else SOME (Logic.all v (Logic.mk_implies (prem, concl)))
blanchet@50293
   489
                      end;
blanchet@50293
   490
                    val goals = map3 mk_goal ctrs discs selss;
blanchet@50293
   491
                  in
blanchet@50293
   492
                    map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
blanchet@50293
   493
                      Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
blanchet@50293
   494
                        mk_collapse_tac ctxt m discD sel_thms)
blanchet@50293
   495
                      |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
blanchet@50293
   496
                    |> map_filter I
blanchet@50293
   497
                  end;
blanchet@50293
   498
blanchet@50293
   499
              val case_eq_thms =
blanchet@50293
   500
                if no_dests then
blanchet@50293
   501
                  []
blanchet@50293
   502
                else
blanchet@50293
   503
                  let
blanchet@50293
   504
                    fun mk_body f sels = Term.list_comb (f, map ap_v sels);
blanchet@50293
   505
                    val goal =
blanchet@50293
   506
                      mk_Trueprop_eq (fcase $ v, mk_IfN B (map ap_v discs) (map2 mk_body fs selss));
blanchet@50293
   507
                  in
blanchet@50293
   508
                    [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
blanchet@50293
   509
                      mk_case_eq_tac ctxt n exhaust_thm' case_thms disc_thmss' sel_thmss)]
blanchet@50293
   510
                    |> Proof_Context.export names_lthy lthy
blanchet@50293
   511
                  end;
blanchet@50131
   512
            in
blanchet@50296
   513
              (all_sel_thms, sel_thmss, disc_thms, discI_thms, disc_exclude_thms, disc_exhaust_thms,
blanchet@50293
   514
               collapse_thms, case_eq_thms)
blanchet@50131
   515
            end;
blanchet@50040
   516
blanchet@50048
   517
        val (case_cong_thm, weak_case_cong_thm) =
blanchet@50047
   518
          let
blanchet@50047
   519
            fun mk_prem xctr xs f g =
blanchet@50060
   520
              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (w, xctr),
blanchet@50047
   521
                mk_Trueprop_eq (f, g)));
blanchet@50048
   522
blanchet@50048
   523
            val v_eq_w = mk_Trueprop_eq (v, w);
blanchet@50047
   524
blanchet@50047
   525
            val goal =
blanchet@50048
   526
              Logic.list_implies (v_eq_w :: map4 mk_prem xctrs xss fs gs,
blanchet@50216
   527
                 mk_Trueprop_eq (fcase $ v, gcase $ w));
blanchet@50216
   528
            val goal_weak = Logic.mk_implies (v_eq_w, mk_Trueprop_eq (fcase $ v, fcase $ w));
blanchet@50047
   529
          in
blanchet@50064
   530
            (Skip_Proof.prove lthy [] [] goal (fn _ => mk_case_cong_tac exhaust_thm' case_thms),
blanchet@50048
   531
             Skip_Proof.prove lthy [] [] goal_weak (K (etac arg_cong 1)))
blanchet@50048
   532
            |> pairself (singleton (Proof_Context.export names_lthy lthy))
blanchet@50047
   533
          end;
blanchet@50040
   534
blanchet@50059
   535
        val (split_thm, split_asm_thm) =
blanchet@50058
   536
          let
blanchet@50059
   537
            fun mk_conjunct xctr xs f_xs =
blanchet@50058
   538
              list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (v, xctr), q $ f_xs));
blanchet@50059
   539
            fun mk_disjunct xctr xs f_xs =
blanchet@50059
   540
              list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (v, xctr),
blanchet@50059
   541
                HOLogic.mk_not (q $ f_xs)));
blanchet@50059
   542
blanchet@50216
   543
            val lhs = q $ (fcase $ v);
blanchet@50059
   544
blanchet@50058
   545
            val goal =
blanchet@50059
   546
              mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
blanchet@50059
   547
            val goal_asm =
blanchet@50059
   548
              mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
blanchet@50059
   549
                (map3 mk_disjunct xctrs xss xfs)));
blanchet@50059
   550
blanchet@50059
   551
            val split_thm =
blanchet@50064
   552
              Skip_Proof.prove lthy [] [] goal
blanchet@50067
   553
                (fn _ => mk_split_tac exhaust_thm' case_thms inject_thmss distinct_thmsss)
blanchet@50059
   554
              |> singleton (Proof_Context.export names_lthy lthy)
blanchet@50059
   555
            val split_asm_thm =
blanchet@50059
   556
              Skip_Proof.prove lthy [] [] goal_asm (fn {context = ctxt, ...} =>
blanchet@50059
   557
                mk_split_asm_tac ctxt split_thm)
blanchet@50059
   558
              |> singleton (Proof_Context.export names_lthy lthy)
blanchet@50058
   559
          in
blanchet@50059
   560
            (split_thm, split_asm_thm)
blanchet@50058
   561
          end;
blanchet@50040
   562
blanchet@50452
   563
        val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
blanchet@50315
   564
        val cases_type_attr = Attrib.internal (K (Induct.cases_type dataT_name));
blanchet@50315
   565
blanchet@50067
   566
        val notes =
blanchet@50315
   567
          [(case_congN, [case_cong_thm], []),
blanchet@50315
   568
           (case_eqN, case_eq_thms, []),
blanchet@50315
   569
           (casesN, case_thms, simp_attrs),
blanchet@50315
   570
           (collapseN, collapse_thms, simp_attrs),
blanchet@50315
   571
           (discsN, disc_thms, simp_attrs),
blanchet@50315
   572
           (disc_excludeN, disc_exclude_thms, []),
blanchet@50315
   573
           (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
blanchet@50315
   574
           (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
blanchet@50315
   575
           (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
blanchet@50453
   576
           (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
blanchet@50315
   577
           (nchotomyN, [nchotomy_thm], []),
blanchet@50315
   578
           (selsN, all_sel_thms, simp_attrs),
blanchet@50315
   579
           (splitN, [split_thm], []),
blanchet@50315
   580
           (split_asmN, [split_asm_thm], []),
blanchet@50315
   581
           (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
blanchet@50315
   582
          |> filter_out (null o #2)
blanchet@50315
   583
          |> map (fn (thmN, thms, attrs) =>
blanchet@50317
   584
            ((Binding.qualify true (Binding.name_of data_b) (Binding.name thmN), attrs),
blanchet@50317
   585
             [(thms, [])]));
blanchet@50315
   586
blanchet@50315
   587
        val notes' =
blanchet@50315
   588
          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
blanchet@50315
   589
          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
blanchet@50034
   590
      in
blanchet@50453
   591
        ((selss, inject_thms, distinct_thms, case_thms, discI_thms, sel_thmss),
blanchet@50453
   592
         lthy |> Local_Theory.notes (notes' @ notes) |> snd)
blanchet@50034
   593
      end;
blanchet@50032
   594
  in
blanchet@50136
   595
    (goalss, after_qed, lthy')
blanchet@50032
   596
  end;
blanchet@50032
   597
blanchet@50214
   598
fun wrap_datatype tacss = (fn (goalss, after_qed, lthy) =>
blanchet@50126
   599
  map2 (map2 (Skip_Proof.prove lthy [] [])) goalss tacss
blanchet@50295
   600
  |> (fn thms => after_qed thms lthy)) oo prepare_wrap_datatype (K I);
blanchet@50126
   601
blanchet@50312
   602
val wrap_datatype_cmd = (fn (goalss, after_qed, lthy) =>
blanchet@50312
   603
  Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
blanchet@50312
   604
  prepare_wrap_datatype Syntax.read_term;
blanchet@50312
   605
blanchet@50295
   606
fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
blanchet@50295
   607
blanchet@50295
   608
val parse_bindings = parse_bracket_list Parse.binding;
blanchet@50295
   609
val parse_bindingss = parse_bracket_list parse_bindings;
blanchet@50295
   610
blanchet@50295
   611
val parse_bound_term = (Parse.binding --| @{keyword ":"}) -- Parse.term;
blanchet@50295
   612
val parse_bound_terms = parse_bracket_list parse_bound_term;
blanchet@50295
   613
val parse_bound_termss = parse_bracket_list parse_bound_terms;
blanchet@50032
   614
blanchet@50293
   615
val parse_wrap_options =
blanchet@50293
   616
  Scan.optional (@{keyword "("} |-- (@{keyword "no_dests"} >> K true) --| @{keyword ")"}) false;
blanchet@50293
   617
blanchet@50032
   618
val _ =
blanchet@50089
   619
  Outer_Syntax.local_theory_to_proof @{command_spec "wrap_data"} "wraps an existing datatype"
blanchet@50293
   620
    ((parse_wrap_options -- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) --
blanchet@50295
   621
      Parse.term -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
blanchet@50295
   622
        Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
blanchet@50214
   623
     >> wrap_datatype_cmd);
blanchet@50032
   624
blanchet@50032
   625
end;