src/HOL/Tools/Datatype/rep_datatype.ML
changeset 45907 4b41967bd77e
parent 45901 cea7cd0c7e99
child 45909 6fe61da4c467
equal deleted inserted replaced
45906:0aaeb5520f2f 45907:4b41967bd77e
     1 (*  Title:      HOL/Tools/Datatype/rep_datatype.ML
     1 (*  Title:      HOL/Tools/Datatype/rep_datatype.ML
     2     Author:     Stefan Berghofer, TU Muenchen
     2     Author:     Stefan Berghofer, TU Muenchen
     3 
     3 
     4 Representation of existing types as datatypes.
     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).
     5 *)
     8 *)
     6 
     9 
     7 signature REP_DATATYPE =
    10 signature REP_DATATYPE =
     8 sig
    11 sig
     9   val derive_datatype_props : Datatype_Aux.config -> string list -> Datatype_Aux.descr list ->
    12   val derive_datatype_props : Datatype_Aux.config -> string list -> Datatype_Aux.descr list ->
    14     string list -> theory -> Proof.state
    17     string list -> theory -> Proof.state
    15 end;
    18 end;
    16 
    19 
    17 structure Rep_Datatype: REP_DATATYPE =
    20 structure Rep_Datatype: REP_DATATYPE =
    18 struct
    21 struct
       
    22 
       
    23 type config = Datatype_Aux.config;
       
    24 type descr = Datatype_Aux.descr;
       
    25 
       
    26 
       
    27 
       
    28 (** derived definitions and proofs **)
       
    29 
       
    30 (* case distinction theorems *)
       
    31 
       
    32 fun prove_casedist_thms (config : config) new_type_names descr induct case_names_exhausts thy =
       
    33   let
       
    34     val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
       
    35 
       
    36     val descr' = flat descr;
       
    37     val recTs = Datatype_Aux.get_rec_types descr';
       
    38     val newTs = take (length (hd descr)) recTs;
       
    39 
       
    40     val maxidx = Thm.maxidx_of induct;
       
    41     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
       
    42 
       
    43     fun prove_casedist_thm (i, (T, t)) =
       
    44       let
       
    45         val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
       
    46           Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
       
    47         val P =
       
    48           Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx + 1), T), Bound 0) $
       
    49             Var (("P", 0), HOLogic.boolT));
       
    50         val insts = take i dummyPs @ (P :: drop (i + 1) dummyPs);
       
    51         val cert = cterm_of thy;
       
    52         val insts' = map cert induct_Ps ~~ map cert insts;
       
    53         val induct' =
       
    54           refl RS
       
    55             (nth (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i RSN (2, rev_mp));
       
    56       in
       
    57         Skip_Proof.prove_global thy []
       
    58           (Logic.strip_imp_prems t)
       
    59           (Logic.strip_imp_concl t)
       
    60           (fn {prems, ...} =>
       
    61             EVERY
       
    62               [rtac induct' 1,
       
    63                REPEAT (rtac TrueI 1),
       
    64                REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
       
    65                REPEAT (rtac TrueI 1)])
       
    66       end;
       
    67 
       
    68     val casedist_thms =
       
    69       map_index prove_casedist_thm (newTs ~~ Datatype_Prop.make_casedists descr);
       
    70   in
       
    71     thy
       
    72     |> Datatype_Aux.store_thms_atts "exhaust" new_type_names
       
    73         (map single case_names_exhausts) casedist_thms
       
    74   end;
       
    75 
       
    76 
       
    77 (* primrec combinators *)
       
    78 
       
    79 fun prove_primrec_thms (config : config) new_type_names descr
       
    80     injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
       
    81   let
       
    82     val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
       
    83 
       
    84     val big_name = space_implode "_" new_type_names;
       
    85     val thy0 = Sign.add_path big_name thy;
       
    86 
       
    87     val descr' = flat descr;
       
    88     val recTs = Datatype_Aux.get_rec_types descr';
       
    89     val used = fold Term.add_tfree_namesT recTs [];
       
    90     val newTs = take (length (hd descr)) recTs;
       
    91 
       
    92     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
       
    93 
       
    94     val big_rec_name' = big_name ^ "_rec_set";
       
    95     val rec_set_names' =
       
    96       if length descr' = 1 then [big_rec_name']
       
    97       else map (prefix (big_rec_name' ^ "_") o string_of_int) (1 upto length descr');
       
    98     val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
       
    99 
       
   100     val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr used;
       
   101 
       
   102     val rec_set_Ts =
       
   103       map (fn (T1, T2) => (reccomb_fn_Ts @ [T1, T2]) ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
       
   104 
       
   105     val rec_fns =
       
   106       map (uncurry (Datatype_Aux.mk_Free "f")) (reccomb_fn_Ts ~~ (1 upto length reccomb_fn_Ts));
       
   107     val rec_sets' =
       
   108       map (fn c => list_comb (Free c, rec_fns)) (rec_set_names' ~~ rec_set_Ts);
       
   109     val rec_sets =
       
   110       map (fn c => list_comb (Const c, rec_fns)) (rec_set_names ~~ rec_set_Ts);
       
   111 
       
   112     (* introduction rules for graph of primrec function *)
       
   113 
       
   114     fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) =
       
   115       let
       
   116         fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
       
   117           let val free1 = Datatype_Aux.mk_Free "x" U j in
       
   118             (case (Datatype_Aux.strip_dtyp dt, strip_type U) of
       
   119               ((_, Datatype_Aux.DtRec m), (Us, _)) =>
       
   120                 let
       
   121                   val free2 = Datatype_Aux.mk_Free "y" (Us ---> nth rec_result_Ts m) k;
       
   122                   val i = length Us;
       
   123                 in
       
   124                   (j + 1, k + 1,
       
   125                     HOLogic.mk_Trueprop (HOLogic.list_all
       
   126                       (map (pair "x") Us, nth rec_sets' m $
       
   127                         Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems,
       
   128                     free1 :: t1s, free2 :: t2s)
       
   129                 end
       
   130             | _ => (j + 1, k, prems, free1 :: t1s, t2s))
       
   131           end;
       
   132 
       
   133         val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
       
   134         val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []);
       
   135 
       
   136       in
       
   137         (rec_intr_ts @
       
   138           [Logic.list_implies (prems, HOLogic.mk_Trueprop
       
   139             (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
       
   140               list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
       
   141       end;
       
   142 
       
   143     val (rec_intr_ts, _) =
       
   144       fold (fn ((d, T), set_name) =>
       
   145         fold (make_rec_intr T set_name) (#3 (snd d))) (descr' ~~ recTs ~~ rec_sets') ([], 0);
       
   146 
       
   147     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
       
   148       thy0
       
   149       |> Sign.map_naming Name_Space.conceal
       
   150       |> Inductive.add_inductive_global
       
   151           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
       
   152             coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false}
       
   153           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
       
   154           (map dest_Free rec_fns)
       
   155           (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) []
       
   156       ||> Sign.restore_naming thy0
       
   157       ||> Theory.checkpoint;
       
   158 
       
   159     (* prove uniqueness and termination of primrec combinators *)
       
   160 
       
   161     val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ...";
       
   162 
       
   163     fun mk_unique_tac ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
       
   164       let
       
   165         val distinct_tac =
       
   166           if i < length newTs then
       
   167             full_simp_tac (HOL_ss addsimps (nth dist_rewrites i)) 1
       
   168           else full_simp_tac (HOL_ss addsimps (flat other_dist_rewrites)) 1;
       
   169 
       
   170         val inject =
       
   171           map (fn r => r RS iffD1)
       
   172             (if i < length newTs then nth constr_inject i else injects_of tname);
       
   173 
       
   174         fun mk_unique_constr_tac n (cname, cargs) (tac, intr :: intrs, j) =
       
   175           let
       
   176             val k = length (filter Datatype_Aux.is_rec_type cargs);
       
   177           in
       
   178             (EVERY
       
   179               [DETERM tac,
       
   180                 REPEAT (etac ex1E 1), rtac ex1I 1,
       
   181                 DEPTH_SOLVE_1 (ares_tac [intr] 1),
       
   182                 REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
       
   183                 etac elim 1,
       
   184                 REPEAT_DETERM_N j distinct_tac,
       
   185                 TRY (dresolve_tac inject 1),
       
   186                 REPEAT (etac conjE 1), hyp_subst_tac 1,
       
   187                 REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
       
   188                 TRY (hyp_subst_tac 1),
       
   189                 rtac refl 1,
       
   190                 REPEAT_DETERM_N (n - j - 1) distinct_tac],
       
   191               intrs, j + 1)
       
   192           end;
       
   193 
       
   194         val (tac', intrs', _) =
       
   195           fold (mk_unique_constr_tac (length constrs)) constrs (tac, intrs, 0);
       
   196       in (tac', intrs') end;
       
   197 
       
   198     val rec_unique_thms =
       
   199       let
       
   200         val rec_unique_ts =
       
   201           map (fn (((set_t, T1), T2), i) =>
       
   202             Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
       
   203               absfree ("y", T2) (set_t $ Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
       
   204                 (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
       
   205         val cert = cterm_of thy1;
       
   206         val insts =
       
   207           map (fn ((i, T), t) => absfree ("x" ^ string_of_int i, T) t)
       
   208             ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
       
   209         val induct' = cterm_instantiate (map cert induct_Ps ~~ map cert insts) induct;
       
   210         val (tac, _) =
       
   211           fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
       
   212             (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1 THEN
       
   213                 rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs));
       
   214       in
       
   215         Datatype_Aux.split_conj_thm (Skip_Proof.prove_global thy1 [] []
       
   216           (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts)) (K tac))
       
   217       end;
       
   218 
       
   219     val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
       
   220 
       
   221     (* define primrec combinators *)
       
   222 
       
   223     val big_reccomb_name = space_implode "_" new_type_names ^ "_rec";
       
   224     val reccomb_names =
       
   225       map (Sign.full_bname thy1)
       
   226         (if length descr' = 1 then [big_reccomb_name]
       
   227          else map (prefix (big_reccomb_name ^ "_") o string_of_int) (1 upto length descr'));
       
   228     val reccombs =
       
   229       map (fn ((name, T), T') => Const (name, reccomb_fn_Ts @ [T] ---> T'))
       
   230         (reccomb_names ~~ recTs ~~ rec_result_Ts);
       
   231 
       
   232     val (reccomb_defs, thy2) =
       
   233       thy1
       
   234       |> Sign.add_consts_i (map (fn ((name, T), T') =>
       
   235             (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
       
   236             (reccomb_names ~~ recTs ~~ rec_result_Ts))
       
   237       |> (Global_Theory.add_defs false o map Thm.no_attributes)
       
   238           (map
       
   239             (fn ((((name, comb), set), T), T') =>
       
   240               (Binding.name (Long_Name.base_name name ^ "_def"),
       
   241                 Logic.mk_equals (comb, fold_rev lambda rec_fns (absfree ("x", T)
       
   242                  (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
       
   243                    (set $ Free ("x", T) $ Free ("y", T')))))))
       
   244             (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
       
   245       ||> Sign.parent_path
       
   246       ||> Theory.checkpoint;
       
   247 
       
   248 
       
   249     (* prove characteristic equations for primrec combinators *)
       
   250 
       
   251     val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ...";
       
   252 
       
   253     val rec_thms =
       
   254       map (fn t =>
       
   255         Skip_Proof.prove_global thy2 [] [] t
       
   256           (fn _ => EVERY
       
   257             [rewrite_goals_tac reccomb_defs,
       
   258              rtac @{thm the1_equality} 1,
       
   259              resolve_tac rec_unique_thms 1,
       
   260              resolve_tac rec_intrs 1,
       
   261              REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
       
   262        (Datatype_Prop.make_primrecs reccomb_names descr thy2);
       
   263   in
       
   264     thy2
       
   265     |> Sign.add_path (space_implode "_" new_type_names)
       
   266     |> Global_Theory.note_thmss ""
       
   267       [((Binding.name "recs", [Nitpick_Simps.add]), [(rec_thms, [])])]
       
   268     ||> Sign.parent_path
       
   269     ||> Theory.checkpoint
       
   270     |-> (fn thms => pair (reccomb_names, maps #2 thms))
       
   271   end;
       
   272 
       
   273 
       
   274 (* case combinators *)
       
   275 
       
   276 fun prove_case_thms (config : config) new_type_names descr reccomb_names primrec_thms thy =
       
   277   let
       
   278     val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
       
   279 
       
   280     val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
       
   281 
       
   282     val descr' = flat descr;
       
   283     val recTs = Datatype_Aux.get_rec_types descr';
       
   284     val used = fold Term.add_tfree_namesT recTs [];
       
   285     val newTs = take (length (hd descr)) recTs;
       
   286     val T' = TFree (singleton (Name.variant_list used) "'t", HOLogic.typeS);
       
   287 
       
   288     fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' dt) ---> T';
       
   289 
       
   290     val case_dummy_fns =
       
   291       map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
       
   292         let
       
   293           val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
       
   294           val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
       
   295         in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
       
   296 
       
   297     val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
       
   298 
       
   299     (* define case combinators via primrec combinators *)
       
   300 
       
   301     val (case_defs, thy2) =
       
   302       fold (fn ((((i, (_, _, constrs)), T), name), recname) => fn (defs, thy) =>
       
   303           let
       
   304             val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
       
   305               let
       
   306                 val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
       
   307                 val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
       
   308                 val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
       
   309                 val frees = take (length cargs) frees';
       
   310                 val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
       
   311               in
       
   312                 (free, fold_rev (absfree o dest_Free) frees' (list_comb (free, frees)))
       
   313               end) (constrs ~~ (1 upto length constrs)));
       
   314 
       
   315             val caseT = map (snd o dest_Free) fns1 @ [T] ---> T';
       
   316             val fns = flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns);
       
   317             val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
       
   318             val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
       
   319             val def =
       
   320               (Binding.name (Long_Name.base_name name ^ "_def"),
       
   321                 Logic.mk_equals (Const (name, caseT),
       
   322                   fold_rev lambda fns1
       
   323                     (list_comb (reccomb,
       
   324                       flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
       
   325             val ([def_thm], thy') =
       
   326               thy
       
   327               |> Sign.declare_const_global decl |> snd
       
   328               |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
       
   329 
       
   330           in (defs @ [def_thm], thy') end)
       
   331         (hd descr ~~ newTs ~~ case_names ~~ take (length newTs) reccomb_names) ([], thy1)
       
   332       ||> Theory.checkpoint;
       
   333 
       
   334     val case_thms =
       
   335       (map o map) (fn t =>
       
   336           Skip_Proof.prove_global thy2 [] [] t
       
   337             (fn _ =>
       
   338               EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))
       
   339         (Datatype_Prop.make_cases case_names descr thy2);
       
   340   in
       
   341     thy2
       
   342     |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms)
       
   343     |> Sign.parent_path
       
   344     |> Datatype_Aux.store_thmss "cases" new_type_names case_thms
       
   345     |-> (fn thmss => pair (thmss, case_names))
       
   346   end;
       
   347 
       
   348 
       
   349 (* case splitting *)
       
   350 
       
   351 fun prove_split_thms (config : config)
       
   352     new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
       
   353   let
       
   354     val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
       
   355 
       
   356     val descr' = flat descr;
       
   357     val recTs = Datatype_Aux.get_rec_types descr';
       
   358     val newTs = take (length (hd descr)) recTs;
       
   359 
       
   360     fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) =
       
   361       let
       
   362         val cert = cterm_of thy;
       
   363         val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
       
   364         val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion;
       
   365         val tac =
       
   366           EVERY [rtac exhaustion' 1,
       
   367             ALLGOALS (asm_simp_tac (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))];
       
   368       in
       
   369         (Skip_Proof.prove_global thy [] [] t1 (K tac),
       
   370          Skip_Proof.prove_global thy [] [] t2 (K tac))
       
   371       end;
       
   372 
       
   373     val split_thm_pairs =
       
   374       map prove_split_thms
       
   375         (Datatype_Prop.make_splits case_names descr thy ~~ constr_inject ~~
       
   376           dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
       
   377 
       
   378     val (split_thms, split_asm_thms) = split_list split_thm_pairs
       
   379 
       
   380   in
       
   381     thy
       
   382     |> Datatype_Aux.store_thms "split" new_type_names split_thms
       
   383     ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
       
   384     |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
       
   385   end;
       
   386 
       
   387 fun prove_weak_case_congs new_type_names case_names descr thy =
       
   388   let
       
   389     fun prove_weak_case_cong t =
       
   390      Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
       
   391        (fn {prems, ...} => EVERY [rtac (hd prems RS arg_cong) 1]);
       
   392 
       
   393     val weak_case_congs =
       
   394       map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs case_names descr thy);
       
   395 
       
   396   in thy |> Datatype_Aux.store_thms "weak_case_cong" new_type_names weak_case_congs end;
       
   397 
       
   398 
       
   399 (* additional theorems for TFL *)
       
   400 
       
   401 fun prove_nchotomys (config : config) new_type_names descr casedist_thms thy =
       
   402   let
       
   403     val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
       
   404 
       
   405     fun prove_nchotomy (t, exhaustion) =
       
   406       let
       
   407         (* For goal i, select the correct disjunct to attack, then prove it *)
       
   408         fun tac i 0 = EVERY [TRY (rtac disjI1 i), hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
       
   409           | tac i n = rtac disjI2 i THEN tac i (n - 1);
       
   410       in
       
   411         Skip_Proof.prove_global thy [] [] t
       
   412           (fn _ =>
       
   413             EVERY [rtac allI 1,
       
   414              Datatype_Aux.exh_tac (K exhaustion) 1,
       
   415              ALLGOALS (fn i => tac i (i - 1))])
       
   416       end;
       
   417 
       
   418     val nchotomys =
       
   419       map prove_nchotomy (Datatype_Prop.make_nchotomys descr ~~ casedist_thms);
       
   420 
       
   421   in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
       
   422 
       
   423 fun prove_case_congs new_type_names case_names descr nchotomys case_thms thy =
       
   424   let
       
   425     fun prove_case_cong ((t, nchotomy), case_rewrites) =
       
   426       let
       
   427         val Const ("==>", _) $ tm $ _ = t;
       
   428         val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma) = tm;
       
   429         val cert = cterm_of thy;
       
   430         val nchotomy' = nchotomy RS spec;
       
   431         val [v] = Term.add_vars (concl_of nchotomy') [];
       
   432         val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy';
       
   433       in
       
   434         Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
       
   435           (fn {prems, ...} =>
       
   436             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites)) in
       
   437               EVERY [
       
   438                 simp_tac (HOL_ss addsimps [hd prems]) 1,
       
   439                 cut_facts_tac [nchotomy''] 1,
       
   440                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
       
   441                 REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
       
   442             end)
       
   443       end;
       
   444 
       
   445     val case_congs =
       
   446       map prove_case_cong
       
   447         (Datatype_Prop.make_case_congs case_names descr thy ~~ nchotomys ~~ case_thms);
       
   448 
       
   449   in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
       
   450 
       
   451 
       
   452 
       
   453 (** derive datatype props **)
       
   454 
       
   455 local
    19 
   456 
    20 fun make_dt_info descr induct inducts rec_names rec_rewrites
   457 fun make_dt_info descr induct inducts rec_names rec_rewrites
    21     (index, (((((((((((_, (tname, _, _))), inject), distinct),
   458     (index, (((((((((((_, (tname, _, _))), inject), distinct),
    22       exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
   459       exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
    23         (split, split_asm))) =
   460         (split, split_asm))) =
    37     case_cong = case_cong,
   474     case_cong = case_cong,
    38     weak_case_cong = weak_case_cong,
   475     weak_case_cong = weak_case_cong,
    39     split = split,
   476     split = split,
    40     split_asm = split_asm});
   477     split_asm = split_asm});
    41 
   478 
       
   479 in
       
   480 
    42 fun derive_datatype_props config dt_names descr induct inject distinct thy1 =
   481 fun derive_datatype_props config dt_names descr induct inject distinct thy1 =
    43   let
   482   let
    44     val thy2 = thy1 |> Theory.checkpoint;
   483     val thy2 = thy1 |> Theory.checkpoint;
    45     val flat_descr = flat descr;
   484     val flat_descr = flat descr;
    46     val new_type_names = map Long_Name.base_name dt_names;
   485     val new_type_names = map Long_Name.base_name dt_names;
    47     val _ =
   486     val _ =
    48       Datatype_Aux.message config
   487       Datatype_Aux.message config
    49         ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
   488         ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
    50 
   489 
    51     val (exhaust, thy3) = thy2
   490     val (exhaust, thy3) = thy2
    52       |> Datatype_Abs_Proofs.prove_casedist_thms config new_type_names
   491       |> prove_casedist_thms config new_type_names descr induct
    53         descr induct (Datatype_Data.mk_case_names_exhausts flat_descr dt_names);
   492         (Datatype_Data.mk_case_names_exhausts flat_descr dt_names);
    54     val (nchotomys, thy4) = thy3
   493     val (nchotomys, thy4) = thy3
    55       |> Datatype_Abs_Proofs.prove_nchotomys config new_type_names descr exhaust;
   494       |> prove_nchotomys config new_type_names descr exhaust;
    56     val ((rec_names, rec_rewrites), thy5) = thy4
   495     val ((rec_names, rec_rewrites), thy5) = thy4
    57       |> Datatype_Abs_Proofs.prove_primrec_thms
   496       |> prove_primrec_thms config new_type_names descr
    58         config new_type_names descr (#inject o the o Symtab.lookup (Datatype_Data.get_all thy4))
   497         (#inject o the o Symtab.lookup (Datatype_Data.get_all thy4)) inject
    59         inject (distinct, Datatype_Data.all_distincts thy2 (Datatype_Aux.get_rec_types flat_descr))
   498         (distinct, Datatype_Data.all_distincts thy2 (Datatype_Aux.get_rec_types flat_descr)) induct;
    60         induct;
       
    61     val ((case_rewrites, case_names), thy6) = thy5
   499     val ((case_rewrites, case_names), thy6) = thy5
    62       |> Datatype_Abs_Proofs.prove_case_thms config new_type_names descr rec_names rec_rewrites;
   500       |> prove_case_thms config new_type_names descr rec_names rec_rewrites;
    63     val (case_congs, thy7) = thy6
   501     val (case_congs, thy7) = thy6
    64       |> Datatype_Abs_Proofs.prove_case_congs new_type_names case_names descr
   502       |> prove_case_congs new_type_names case_names descr nchotomys case_rewrites;
    65         nchotomys case_rewrites;
       
    66     val (weak_case_congs, thy8) = thy7
   503     val (weak_case_congs, thy8) = thy7
    67       |> Datatype_Abs_Proofs.prove_weak_case_congs new_type_names case_names descr;
   504       |> prove_weak_case_congs new_type_names case_names descr;
    68     val (splits, thy9) = thy8
   505     val (splits, thy9) = thy8
    69       |> Datatype_Abs_Proofs.prove_split_thms
   506       |> prove_split_thms config new_type_names case_names descr
    70         config new_type_names case_names descr inject distinct exhaust case_rewrites;
   507         inject distinct exhaust case_rewrites;
    71 
   508 
    72     val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
   509     val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
    73     val dt_infos =
   510     val dt_infos =
    74       map_index
   511       map_index
    75         (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
   512         (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
   104     |> Datatype_Data.interpretation_data (config, dt_names)
   541     |> Datatype_Data.interpretation_data (config, dt_names)
   105     |> Datatype_Case.add_case_tr' case_names
   542     |> Datatype_Case.add_case_tr' case_names
   106     |> pair dt_names
   543     |> pair dt_names
   107   end;
   544   end;
   108 
   545 
       
   546 end;
       
   547 
   109 
   548 
   110 
   549 
   111 (** declare existing type as datatype **)
   550 (** declare existing type as datatype **)
   112 
   551 
   113 local
   552 local