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