src/HOL/Tools/Datatype/rep_datatype.ML
changeset 58112 8081087096ad
parent 58111 82db9ad610b9
child 58113 ab6220d6cf70
equal deleted inserted replaced
58111:82db9ad610b9 58112:8081087096ad
     1 (*  Title:      HOL/Tools/Datatype/rep_datatype.ML
       
     2     Author:     Stefan Berghofer, TU Muenchen
       
     3 
       
     4 Representation of existing types as datatypes: proofs and definitions
       
     5 independent of concrete representation of datatypes (i.e. requiring
       
     6 only abstract properties: injectivity / distinctness of constructors
       
     7 and induction).
       
     8 *)
       
     9 
       
    10 signature REP_DATATYPE =
       
    11 sig
       
    12   val derive_datatype_props : Datatype_Aux.config -> string list -> Datatype_Aux.descr list ->
       
    13     thm -> thm list list -> thm list list -> theory -> string list * theory
       
    14   val rep_datatype : Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
       
    15     term list -> theory -> Proof.state
       
    16   val rep_datatype_cmd : Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
       
    17     string list -> theory -> Proof.state
       
    18 end;
       
    19 
       
    20 structure Rep_Datatype: REP_DATATYPE =
       
    21 struct
       
    22 
       
    23 (** derived definitions and proofs **)
       
    24 
       
    25 (* case distinction theorems *)
       
    26 
       
    27 fun prove_casedist_thms (config : Datatype_Aux.config)
       
    28     new_type_names descr induct case_names_exhausts thy =
       
    29   let
       
    30     val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
       
    31 
       
    32     val descr' = flat descr;
       
    33     val recTs = Datatype_Aux.get_rec_types descr';
       
    34     val newTs = take (length (hd descr)) recTs;
       
    35 
       
    36     val maxidx = Thm.maxidx_of induct;
       
    37     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
       
    38 
       
    39     fun prove_casedist_thm (i, (T, t)) =
       
    40       let
       
    41         val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
       
    42           Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
       
    43         val P =
       
    44           Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx + 1), T), Bound 0) $
       
    45             Var (("P", 0), HOLogic.boolT));
       
    46         val insts = take i dummyPs @ (P :: drop (i + 1) dummyPs);
       
    47         val cert = cterm_of thy;
       
    48         val insts' = map cert induct_Ps ~~ map cert insts;
       
    49         val induct' =
       
    50           refl RS
       
    51             (nth (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i RSN (2, rev_mp));
       
    52       in
       
    53         Goal.prove_sorry_global thy []
       
    54           (Logic.strip_imp_prems t)
       
    55           (Logic.strip_imp_concl t)
       
    56           (fn {prems, ...} =>
       
    57             EVERY
       
    58               [rtac induct' 1,
       
    59                REPEAT (rtac TrueI 1),
       
    60                REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
       
    61                REPEAT (rtac TrueI 1)])
       
    62       end;
       
    63 
       
    64     val casedist_thms =
       
    65       map_index prove_casedist_thm (newTs ~~ Datatype_Prop.make_casedists descr);
       
    66   in
       
    67     thy
       
    68     |> Datatype_Aux.store_thms_atts "exhaust" new_type_names
       
    69         (map single case_names_exhausts) casedist_thms
       
    70   end;
       
    71 
       
    72 
       
    73 (* primrec combinators *)
       
    74 
       
    75 fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr
       
    76     injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
       
    77   let
       
    78     val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
       
    79 
       
    80     val big_name = space_implode "_" new_type_names;
       
    81     val thy0 = Sign.add_path big_name thy;
       
    82 
       
    83     val descr' = flat descr;
       
    84     val recTs = Datatype_Aux.get_rec_types descr';
       
    85     val used = fold Term.add_tfree_namesT recTs [];
       
    86     val newTs = take (length (hd descr)) recTs;
       
    87 
       
    88     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
       
    89 
       
    90     val big_rec_name' = "rec_set_" ^ big_name;
       
    91     val rec_set_names' =
       
    92       if length descr' = 1 then [big_rec_name']
       
    93       else map (prefix (big_rec_name' ^ "_") o string_of_int) (1 upto length descr');
       
    94     val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
       
    95 
       
    96     val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr used;
       
    97 
       
    98     val rec_set_Ts =
       
    99       map (fn (T1, T2) => (reccomb_fn_Ts @ [T1, T2]) ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
       
   100 
       
   101     val rec_fns =
       
   102       map (uncurry (Datatype_Aux.mk_Free "f")) (reccomb_fn_Ts ~~ (1 upto length reccomb_fn_Ts));
       
   103     val rec_sets' =
       
   104       map (fn c => list_comb (Free c, rec_fns)) (rec_set_names' ~~ rec_set_Ts);
       
   105     val rec_sets =
       
   106       map (fn c => list_comb (Const c, rec_fns)) (rec_set_names ~~ rec_set_Ts);
       
   107 
       
   108     (* introduction rules for graph of primrec function *)
       
   109 
       
   110     fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) =
       
   111       let
       
   112         fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
       
   113           let val free1 = Datatype_Aux.mk_Free "x" U j in
       
   114             (case (Datatype_Aux.strip_dtyp dt, strip_type U) of
       
   115               ((_, Datatype_Aux.DtRec m), (Us, _)) =>
       
   116                 let
       
   117                   val free2 = Datatype_Aux.mk_Free "y" (Us ---> nth rec_result_Ts m) k;
       
   118                   val i = length Us;
       
   119                 in
       
   120                   (j + 1, k + 1,
       
   121                     HOLogic.mk_Trueprop (HOLogic.list_all
       
   122                       (map (pair "x") Us, nth rec_sets' m $
       
   123                         Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems,
       
   124                     free1 :: t1s, free2 :: t2s)
       
   125                 end
       
   126             | _ => (j + 1, k, prems, free1 :: t1s, t2s))
       
   127           end;
       
   128 
       
   129         val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
       
   130         val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []);
       
   131 
       
   132       in
       
   133         (rec_intr_ts @
       
   134           [Logic.list_implies (prems, HOLogic.mk_Trueprop
       
   135             (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
       
   136               list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
       
   137       end;
       
   138 
       
   139     val (rec_intr_ts, _) =
       
   140       fold (fn ((d, T), set_name) =>
       
   141         fold (make_rec_intr T set_name) (#3 (snd d))) (descr' ~~ recTs ~~ rec_sets') ([], 0);
       
   142 
       
   143     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
       
   144       thy0
       
   145       |> Sign.map_naming Name_Space.conceal
       
   146       |> Inductive.add_inductive_global
       
   147           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
       
   148             coind = false, no_elim = false, no_ind = true, skip_mono = true}
       
   149           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
       
   150           (map dest_Free rec_fns)
       
   151           (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) []
       
   152       ||> Sign.restore_naming thy0;
       
   153 
       
   154     (* prove uniqueness and termination of primrec combinators *)
       
   155 
       
   156     val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ...";
       
   157 
       
   158     fun mk_unique_tac ctxt ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
       
   159       let
       
   160         val distinct_tac =
       
   161           if i < length newTs then
       
   162             full_simp_tac (put_simpset HOL_ss ctxt addsimps (nth dist_rewrites i)) 1
       
   163           else full_simp_tac (put_simpset HOL_ss ctxt addsimps (flat other_dist_rewrites)) 1;
       
   164 
       
   165         val inject =
       
   166           map (fn r => r RS iffD1)
       
   167             (if i < length newTs then nth constr_inject i else injects_of tname);
       
   168 
       
   169         fun mk_unique_constr_tac n (cname, cargs) (tac, intr :: intrs, j) =
       
   170           let
       
   171             val k = length (filter Datatype_Aux.is_rec_type cargs);
       
   172           in
       
   173             (EVERY
       
   174               [DETERM tac,
       
   175                 REPEAT (etac @{thm ex1E} 1), rtac @{thm ex1I} 1,
       
   176                 DEPTH_SOLVE_1 (ares_tac [intr] 1),
       
   177                 REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
       
   178                 etac elim 1,
       
   179                 REPEAT_DETERM_N j distinct_tac,
       
   180                 TRY (dresolve_tac inject 1),
       
   181                 REPEAT (etac conjE 1), hyp_subst_tac ctxt 1,
       
   182                 REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
       
   183                 TRY (hyp_subst_tac ctxt 1),
       
   184                 rtac refl 1,
       
   185                 REPEAT_DETERM_N (n - j - 1) distinct_tac],
       
   186               intrs, j + 1)
       
   187           end;
       
   188 
       
   189         val (tac', intrs', _) =
       
   190           fold (mk_unique_constr_tac (length constrs)) constrs (tac, intrs, 0);
       
   191       in (tac', intrs') end;
       
   192 
       
   193     val rec_unique_thms =
       
   194       let
       
   195         val rec_unique_ts =
       
   196           map (fn (((set_t, T1), T2), i) =>
       
   197             Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
       
   198               absfree ("y", T2) (set_t $ Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
       
   199                 (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
       
   200         val cert = cterm_of thy1;
       
   201         val insts =
       
   202           map (fn ((i, T), t) => absfree ("x" ^ string_of_int i, T) t)
       
   203             ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
       
   204         val induct' = cterm_instantiate (map cert induct_Ps ~~ map cert insts) induct;
       
   205       in
       
   206         Datatype_Aux.split_conj_thm (Goal.prove_sorry_global thy1 [] []
       
   207           (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts))
       
   208           (fn {context = ctxt, ...} =>
       
   209             #1 (fold (mk_unique_tac ctxt) (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
       
   210               (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1 THEN
       
   211                   rewrite_goals_tac ctxt [mk_meta_eq @{thm choice_eq}], rec_intrs)))))
       
   212       end;
       
   213 
       
   214     val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
       
   215 
       
   216     (* define primrec combinators *)
       
   217 
       
   218     val big_reccomb_name = "rec_" ^ space_implode "_" new_type_names;
       
   219     val reccomb_names =
       
   220       map (Sign.full_bname thy1)
       
   221         (if length descr' = 1 then [big_reccomb_name]
       
   222          else map (prefix (big_reccomb_name ^ "_") o string_of_int) (1 upto length descr'));
       
   223     val reccombs =
       
   224       map (fn ((name, T), T') => Const (name, reccomb_fn_Ts @ [T] ---> T'))
       
   225         (reccomb_names ~~ recTs ~~ rec_result_Ts);
       
   226 
       
   227     val (reccomb_defs, thy2) =
       
   228       thy1
       
   229       |> Sign.add_consts (map (fn ((name, T), T') =>
       
   230             (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
       
   231             (reccomb_names ~~ recTs ~~ rec_result_Ts))
       
   232       |> (Global_Theory.add_defs false o map Thm.no_attributes)
       
   233           (map
       
   234             (fn ((((name, comb), set), T), T') =>
       
   235               (Binding.name (Thm.def_name (Long_Name.base_name name)),
       
   236                 Logic.mk_equals (comb, fold_rev lambda rec_fns (absfree ("x", T)
       
   237                  (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
       
   238                    (set $ Free ("x", T) $ Free ("y", T')))))))
       
   239             (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
       
   240       ||> Sign.parent_path;
       
   241 
       
   242 
       
   243     (* prove characteristic equations for primrec combinators *)
       
   244 
       
   245     val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ...";
       
   246 
       
   247     val rec_thms =
       
   248       map (fn t =>
       
   249         Goal.prove_sorry_global thy2 [] [] t
       
   250           (fn {context = ctxt, ...} => EVERY
       
   251             [rewrite_goals_tac ctxt reccomb_defs,
       
   252              rtac @{thm the1_equality} 1,
       
   253              resolve_tac rec_unique_thms 1,
       
   254              resolve_tac rec_intrs 1,
       
   255              REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
       
   256        (Datatype_Prop.make_primrecs reccomb_names descr thy2);
       
   257   in
       
   258     thy2
       
   259     |> Sign.add_path (space_implode "_" new_type_names)
       
   260     |> Global_Theory.note_thmss ""
       
   261       [((Binding.name "rec", [Named_Theorems.add @{named_theorems nitpick_simp}]),
       
   262           [(rec_thms, [])])]
       
   263     ||> Sign.parent_path
       
   264     |-> (fn thms => pair (reccomb_names, maps #2 thms))
       
   265   end;
       
   266 
       
   267 
       
   268 (* case combinators *)
       
   269 
       
   270 fun prove_case_thms (config : Datatype_Aux.config)
       
   271     new_type_names descr reccomb_names primrec_thms thy =
       
   272   let
       
   273     val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
       
   274 
       
   275     val ctxt = Proof_Context.init_global thy;
       
   276     val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
       
   277 
       
   278     val descr' = flat descr;
       
   279     val recTs = Datatype_Aux.get_rec_types descr';
       
   280     val used = fold Term.add_tfree_namesT recTs [];
       
   281     val newTs = take (length (hd descr)) recTs;
       
   282     val T' = TFree (singleton (Name.variant_list used) "'t", @{sort type});
       
   283 
       
   284     fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' dt) ---> T';
       
   285 
       
   286     val case_dummy_fns =
       
   287       map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
       
   288         let
       
   289           val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
       
   290           val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
       
   291         in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
       
   292 
       
   293     val case_names0 = map (fn s => Sign.full_bname thy1 ("case_" ^ s)) new_type_names;
       
   294 
       
   295     (* define case combinators via primrec combinators *)
       
   296 
       
   297     fun def_case ((((i, (_, _, constrs)), T as Type (Tcon, _)), name), recname) (defs, thy) =
       
   298       if is_some (Ctr_Sugar.ctr_sugar_of ctxt Tcon) then
       
   299         (defs, thy)
       
   300       else
       
   301         let
       
   302           val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
       
   303             let
       
   304               val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
       
   305               val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
       
   306               val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
       
   307               val frees = take (length cargs) frees';
       
   308               val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
       
   309             in
       
   310               (free, fold_rev (absfree o dest_Free) frees' (list_comb (free, frees)))
       
   311             end) (constrs ~~ (1 upto length constrs)));
       
   312 
       
   313           val caseT = map (snd o dest_Free) fns1 @ [T] ---> T';
       
   314           val fns = flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns);
       
   315           val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
       
   316           val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
       
   317           val def =
       
   318             (Binding.name (Thm.def_name (Long_Name.base_name name)),
       
   319               Logic.mk_equals (Const (name, caseT),
       
   320                 fold_rev lambda fns1
       
   321                   (list_comb (reccomb,
       
   322                     flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
       
   323           val ([def_thm], thy') =
       
   324             thy
       
   325             |> Sign.declare_const_global decl |> snd
       
   326             |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
       
   327         in (defs @ [def_thm], thy') end;
       
   328 
       
   329     val (case_defs, thy2) =
       
   330       fold def_case (hd descr ~~ newTs ~~ case_names0 ~~ take (length newTs) reccomb_names)
       
   331         ([], thy1);
       
   332 
       
   333     fun prove_case t =
       
   334       Goal.prove_sorry_global thy2 [] [] t (fn {context = ctxt, ...} =>
       
   335         EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]);
       
   336 
       
   337     fun prove_cases (Type (Tcon, _)) ts =
       
   338       (case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
       
   339         SOME {case_thms, ...} => case_thms
       
   340       | NONE => map prove_case ts);
       
   341 
       
   342     val case_thms =
       
   343       map2 prove_cases newTs (Datatype_Prop.make_cases case_names0 descr thy2);
       
   344 
       
   345     fun case_name_of (th :: _) =
       
   346       fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))));
       
   347 
       
   348     val case_names = map case_name_of case_thms;
       
   349   in
       
   350     thy2
       
   351     |> Context.theory_map
       
   352         ((fold o fold) (Named_Theorems.add_thm @{named_theorems nitpick_simp}) case_thms)
       
   353     |> Sign.parent_path
       
   354     |> Datatype_Aux.store_thmss "case" new_type_names case_thms
       
   355     |-> (fn thmss => pair (thmss, case_names))
       
   356   end;
       
   357 
       
   358 
       
   359 (* case splitting *)
       
   360 
       
   361 fun prove_split_thms (config : Datatype_Aux.config)
       
   362     new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
       
   363   let
       
   364     val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
       
   365 
       
   366     val descr' = flat descr;
       
   367     val recTs = Datatype_Aux.get_rec_types descr';
       
   368     val newTs = take (length (hd descr)) recTs;
       
   369 
       
   370     fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) =
       
   371       let
       
   372         val cert = cterm_of thy;
       
   373         val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
       
   374         val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion;
       
   375         fun tac ctxt =
       
   376           EVERY [rtac exhaustion' 1,
       
   377             ALLGOALS (asm_simp_tac
       
   378               (put_simpset HOL_ss ctxt addsimps (dist_rewrites' @ inject @ case_thms')))];
       
   379       in
       
   380         (Goal.prove_sorry_global thy [] [] t1 (tac o #context),
       
   381          Goal.prove_sorry_global thy [] [] t2 (tac o #context))
       
   382       end;
       
   383 
       
   384     val split_thm_pairs =
       
   385       map prove_split_thms
       
   386         (Datatype_Prop.make_splits case_names descr thy ~~ constr_inject ~~
       
   387           dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
       
   388 
       
   389     val (split_thms, split_asm_thms) = split_list split_thm_pairs
       
   390 
       
   391   in
       
   392     thy
       
   393     |> Datatype_Aux.store_thms "split" new_type_names split_thms
       
   394     ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
       
   395     |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
       
   396   end;
       
   397 
       
   398 fun prove_case_cong_weaks new_type_names case_names descr thy =
       
   399   let
       
   400     fun prove_case_cong_weak t =
       
   401      Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
       
   402        (fn {prems, ...} => EVERY [rtac (hd prems RS arg_cong) 1]);
       
   403 
       
   404     val case_cong_weaks =
       
   405       map prove_case_cong_weak (Datatype_Prop.make_case_cong_weaks case_names descr thy);
       
   406 
       
   407   in thy |> Datatype_Aux.store_thms "case_cong_weak" new_type_names case_cong_weaks end;
       
   408 
       
   409 
       
   410 (* additional theorems for TFL *)
       
   411 
       
   412 fun prove_nchotomys (config : Datatype_Aux.config) new_type_names descr casedist_thms thy =
       
   413   let
       
   414     val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
       
   415 
       
   416     fun prove_nchotomy (t, exhaustion) =
       
   417       let
       
   418         (* For goal i, select the correct disjunct to attack, then prove it *)
       
   419         fun tac ctxt i 0 =
       
   420               EVERY [TRY (rtac disjI1 i), hyp_subst_tac ctxt i, REPEAT (rtac exI i), rtac refl i]
       
   421           | tac ctxt i n = rtac disjI2 i THEN tac ctxt i (n - 1);
       
   422       in
       
   423         Goal.prove_sorry_global thy [] [] t
       
   424           (fn {context = ctxt, ...} =>
       
   425             EVERY [rtac allI 1,
       
   426              Datatype_Aux.exh_tac (K exhaustion) 1,
       
   427              ALLGOALS (fn i => tac ctxt i (i - 1))])
       
   428       end;
       
   429 
       
   430     val nchotomys =
       
   431       map prove_nchotomy (Datatype_Prop.make_nchotomys descr ~~ casedist_thms);
       
   432 
       
   433   in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
       
   434 
       
   435 fun prove_case_congs new_type_names case_names descr nchotomys case_thms thy =
       
   436   let
       
   437     fun prove_case_cong ((t, nchotomy), case_rewrites) =
       
   438       let
       
   439         val Const (@{const_name Pure.imp}, _) $ tm $ _ = t;
       
   440         val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma) = tm;
       
   441         val cert = cterm_of thy;
       
   442         val nchotomy' = nchotomy RS spec;
       
   443         val [v] = Term.add_vars (concl_of nchotomy') [];
       
   444         val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy';
       
   445       in
       
   446         Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
       
   447           (fn {context = ctxt, prems, ...} =>
       
   448             let
       
   449               val simplify = asm_simp_tac (put_simpset HOL_ss ctxt addsimps (prems @ case_rewrites))
       
   450             in
       
   451               EVERY [
       
   452                 simp_tac (put_simpset HOL_ss ctxt addsimps [hd prems]) 1,
       
   453                 cut_tac nchotomy'' 1,
       
   454                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
       
   455                 REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
       
   456             end)
       
   457       end;
       
   458 
       
   459     val case_congs =
       
   460       map prove_case_cong
       
   461         (Datatype_Prop.make_case_congs case_names descr thy ~~ nchotomys ~~ case_thms);
       
   462 
       
   463   in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
       
   464 
       
   465 
       
   466 
       
   467 (** derive datatype props **)
       
   468 
       
   469 local
       
   470 
       
   471 fun make_dt_info descr induct inducts rec_names rec_rewrites
       
   472     (index, (((((((((((_, (tname, _, _))), inject), distinct),
       
   473       exhaust), nchotomy), case_name), case_rewrites), case_cong), case_cong_weak),
       
   474         (split, split_asm))) =
       
   475   (tname,
       
   476    {index = index,
       
   477     descr = descr,
       
   478     inject = inject,
       
   479     distinct = distinct,
       
   480     induct = induct,
       
   481     inducts = inducts,
       
   482     exhaust = exhaust,
       
   483     nchotomy = nchotomy,
       
   484     rec_names = rec_names,
       
   485     rec_rewrites = rec_rewrites,
       
   486     case_name = case_name,
       
   487     case_rewrites = case_rewrites,
       
   488     case_cong = case_cong,
       
   489     case_cong_weak = case_cong_weak,
       
   490     split = split,
       
   491     split_asm = split_asm});
       
   492 
       
   493 in
       
   494 
       
   495 fun derive_datatype_props config dt_names descr induct inject distinct thy2 =
       
   496   let
       
   497     val flat_descr = flat descr;
       
   498     val new_type_names = map Long_Name.base_name dt_names;
       
   499     val _ =
       
   500       Datatype_Aux.message config
       
   501         ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
       
   502 
       
   503     val (exhaust, thy3) = thy2
       
   504       |> prove_casedist_thms config new_type_names descr induct
       
   505         (Datatype_Data.mk_case_names_exhausts flat_descr dt_names);
       
   506     val (nchotomys, thy4) = thy3
       
   507       |> prove_nchotomys config new_type_names descr exhaust;
       
   508     val ((rec_names, rec_rewrites), thy5) = thy4
       
   509       |> prove_primrec_thms config new_type_names descr
       
   510         (#inject o the o Symtab.lookup (Datatype_Data.get_all thy4)) inject
       
   511         (distinct, Datatype_Data.all_distincts thy2 (Datatype_Aux.get_rec_types flat_descr)) induct;
       
   512     val ((case_rewrites, case_names), thy6) = thy5
       
   513       |> prove_case_thms config new_type_names descr rec_names rec_rewrites;
       
   514     val (case_congs, thy7) = thy6
       
   515       |> prove_case_congs new_type_names case_names descr nchotomys case_rewrites;
       
   516     val (case_cong_weaks, thy8) = thy7
       
   517       |> prove_case_cong_weaks new_type_names case_names descr;
       
   518     val (splits, thy9) = thy8
       
   519       |> prove_split_thms config new_type_names case_names descr
       
   520         inject distinct exhaust case_rewrites;
       
   521 
       
   522     val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
       
   523     val dt_infos =
       
   524       map_index
       
   525         (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
       
   526         (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
       
   527           case_names ~~ case_rewrites ~~ case_congs ~~ case_cong_weaks ~~ splits);
       
   528     val dt_names = map fst dt_infos;
       
   529     val prfx = Binding.qualify true (space_implode "_" new_type_names);
       
   530     val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
       
   531     val named_rules = flat (map_index (fn (i, tname) =>
       
   532       [((Binding.empty, [Induct.induct_type tname]), [([nth inducts i], [])]),
       
   533        ((Binding.empty, [Induct.cases_type tname]), [([nth exhaust i], [])])]) dt_names);
       
   534     val unnamed_rules = map (fn induct =>
       
   535       ((Binding.empty, [Rule_Cases.inner_rule, Induct.induct_type ""]), [([induct], [])]))
       
   536         (drop (length dt_names) inducts);
       
   537 
       
   538     val ctxt = Proof_Context.init_global thy9;
       
   539     val case_combs =
       
   540       map (Proof_Context.read_const {proper = true, strict = true} ctxt) case_names;
       
   541     val constrss = map (fn (dtname, {descr, index, ...}) =>
       
   542       map (Proof_Context.read_const {proper = true, strict = true} ctxt o fst)
       
   543         (#3 (the (AList.lookup op = descr index)))) dt_infos;
       
   544   in
       
   545     thy9
       
   546     |> Global_Theory.note_thmss ""
       
   547       ([((prfx (Binding.name "simps"), []), [(simps, [])]),
       
   548         ((prfx (Binding.name "inducts"), []), [(inducts, [])]),
       
   549         ((prfx (Binding.name "splits"), []), [(maps (fn (x, y) => [x, y]) splits, [])]),
       
   550         ((Binding.empty, [Simplifier.simp_add]),
       
   551           [(flat case_rewrites @ flat distinct @ rec_rewrites, [])]),
       
   552         ((Binding.empty, [Code.add_default_eqn_attribute]), [(rec_rewrites, [])]),
       
   553         ((Binding.empty, [iff_add]), [(flat inject, [])]),
       
   554         ((Binding.empty, [Classical.safe_elim NONE]),
       
   555           [(map (fn th => th RS notE) (flat distinct), [])]),
       
   556         ((Binding.empty, [Simplifier.cong_add]), [(case_cong_weaks, [])]),
       
   557         ((Binding.empty, [Induct.induct_simp_add]), [(flat (distinct @ inject), [])])] @
       
   558           named_rules @ unnamed_rules)
       
   559     |> snd
       
   560     |> Datatype_Data.register dt_infos
       
   561     |> Context.theory_map (fold2 Case_Translation.register case_combs constrss)
       
   562     |> Datatype_Data.interpretation_data (config, dt_names)
       
   563     |> pair dt_names
       
   564   end;
       
   565 
       
   566 end;
       
   567 
       
   568 
       
   569 
       
   570 (** declare existing type as datatype **)
       
   571 
       
   572 local
       
   573 
       
   574 fun prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct thy1 =
       
   575   let
       
   576     val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
       
   577     val new_type_names = map Long_Name.base_name dt_names;
       
   578     val prfx = Binding.qualify true (space_implode "_" new_type_names);
       
   579     val (((inject, distinct), [(_, [induct])]), thy2) =
       
   580       thy1
       
   581       |> Datatype_Aux.store_thmss "inject" new_type_names raw_inject
       
   582       ||>> Datatype_Aux.store_thmss "distinct" new_type_names raw_distinct
       
   583       ||>> Global_Theory.note_thmss ""
       
   584         [((prfx (Binding.name "induct"), [Datatype_Data.mk_case_names_induct descr]),
       
   585           [([raw_induct], [])])];
       
   586   in
       
   587     thy2
       
   588     |> derive_datatype_props config dt_names [descr] induct inject distinct
       
   589  end;
       
   590 
       
   591 fun gen_rep_datatype prep_term config after_qed raw_ts thy =
       
   592   let
       
   593     val ctxt = Proof_Context.init_global thy;
       
   594 
       
   595     fun constr_of_term (Const (c, T)) = (c, T)
       
   596       | constr_of_term t = error ("Not a constant: " ^ Syntax.string_of_term ctxt t);
       
   597     fun no_constr (c, T) =
       
   598       error ("Bad constructor: " ^ Proof_Context.markup_const ctxt c ^ "::" ^
       
   599         Syntax.string_of_typ ctxt T);
       
   600     fun type_of_constr (cT as (_, T)) =
       
   601       let
       
   602         val frees = Term.add_tfreesT T [];
       
   603         val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type T))
       
   604           handle TYPE _ => no_constr cT
       
   605         val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
       
   606         val _ = if length frees <> length vs then no_constr cT else ();
       
   607       in (tyco, (vs, cT)) end;
       
   608 
       
   609     val raw_cs =
       
   610       AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
       
   611     val _ =
       
   612       (case map_filter (fn (tyco, _) =>
       
   613           if Symtab.defined (Datatype_Data.get_all thy) tyco then SOME tyco else NONE) raw_cs of
       
   614         [] => ()
       
   615       | tycos => error ("Type(s) " ^ commas_quote tycos ^ " already represented inductively"));
       
   616     val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
       
   617     val ms =
       
   618       (case distinct (op =) (map length raw_vss) of
       
   619          [n] => 0 upto n - 1
       
   620       | _ => error "Different types in given constructors");
       
   621     fun inter_sort m =
       
   622       map (fn xs => nth xs m) raw_vss
       
   623       |> foldr1 (Sorts.inter_sort (Sign.classes_of thy));
       
   624     val sorts = map inter_sort ms;
       
   625     val vs = Name.invent_names Name.context Name.aT sorts;
       
   626 
       
   627     fun norm_constr (raw_vs, (c, T)) =
       
   628       (c, map_atyps
       
   629         (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
       
   630 
       
   631     val cs = map (apsnd (map norm_constr)) raw_cs;
       
   632     val dtyps_of_typ = map (Datatype_Aux.dtyp_of_typ (map (rpair vs o fst) cs)) o binder_types;
       
   633     val dt_names = map fst cs;
       
   634 
       
   635     fun mk_spec (i, (tyco, constr)) =
       
   636       (i, (tyco, map Datatype_Aux.DtTFree vs, (map o apsnd) dtyps_of_typ constr));
       
   637     val descr = map_index mk_spec cs;
       
   638     val injs = Datatype_Prop.make_injs [descr];
       
   639     val half_distincts = Datatype_Prop.make_distincts [descr];
       
   640     val ind = Datatype_Prop.make_ind [descr];
       
   641     val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
       
   642 
       
   643     fun after_qed' raw_thms =
       
   644       let
       
   645         val [[[raw_induct]], raw_inject, half_distinct] =
       
   646           unflat rules (map Drule.zero_var_indexes_list raw_thms);
       
   647             (*FIXME somehow dubious*)
       
   648       in
       
   649         Proof_Context.background_theory_result  (* FIXME !? *)
       
   650           (prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct)
       
   651         #-> after_qed
       
   652       end;
       
   653   in
       
   654     ctxt
       
   655     |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
       
   656   end;
       
   657 
       
   658 in
       
   659 
       
   660 val rep_datatype = gen_rep_datatype Sign.cert_term;
       
   661 val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global;
       
   662 
       
   663 end;
       
   664 
       
   665 
       
   666 (* outer syntax *)
       
   667 
       
   668 val _ =
       
   669   Outer_Syntax.command @{command_spec "rep_datatype"} "represent existing types inductively"
       
   670     (Scan.repeat1 Parse.term >> (fn ts =>
       
   671       Toplevel.theory_to_proof (rep_datatype_cmd Datatype_Aux.default_config (K I) ts)));
       
   672 
       
   673 end;