src/HOL/Tools/Old_Datatype/old_datatype.ML
changeset 58112 8081087096ad
parent 58111 82db9ad610b9
child 58114 4e5a43b0e7dd
equal deleted inserted replaced
58111:82db9ad610b9 58112:8081087096ad
       
     1 (*  Title:      HOL/Tools/Old_Datatype/old_datatype.ML
       
     2     Author:     Stefan Berghofer, TU Muenchen
       
     3 
       
     4 Datatype package: definitional introduction of datatypes
       
     5 with proof of characteristic theorems: injectivity / distinctness
       
     6 of constructors and induction.  Main interface to datatypes
       
     7 after full bootstrap of datatype package.
       
     8 *)
       
     9 
       
    10 signature OLD_DATATYPE =
       
    11 sig
       
    12   val distinct_lemma: thm
       
    13   type spec =
       
    14     (binding * (string * sort) list * mixfix) *
       
    15     (binding * typ list * mixfix) list
       
    16   type spec_cmd =
       
    17     (binding * (string * string option) list * mixfix) *
       
    18     (binding * string list * mixfix) list
       
    19   val read_specs: spec_cmd list -> theory -> spec list * Proof.context
       
    20   val check_specs: spec list -> theory -> spec list * Proof.context
       
    21   val add_datatype: Old_Datatype_Aux.config -> spec list -> theory -> string list * theory
       
    22   val add_datatype_cmd: Old_Datatype_Aux.config -> spec_cmd list -> theory -> string list * theory
       
    23   val spec_cmd: spec_cmd parser
       
    24 end;
       
    25 
       
    26 structure Old_Datatype : OLD_DATATYPE =
       
    27 struct
       
    28 
       
    29 (** auxiliary **)
       
    30 
       
    31 val distinct_lemma = @{lemma "f x \<noteq> f y ==> x \<noteq> y" by iprover};
       
    32 val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
       
    33 
       
    34 fun exh_thm_of (dt_info : Old_Datatype_Aux.info Symtab.table) tname =
       
    35   #exhaust (the (Symtab.lookup dt_info tname));
       
    36 
       
    37 val In0_inject = @{thm In0_inject};
       
    38 val In1_inject = @{thm In1_inject};
       
    39 val Scons_inject = @{thm Scons_inject};
       
    40 val Leaf_inject = @{thm Leaf_inject};
       
    41 val In0_eq = @{thm In0_eq};
       
    42 val In1_eq = @{thm In1_eq};
       
    43 val In0_not_In1 = @{thm In0_not_In1};
       
    44 val In1_not_In0 = @{thm In1_not_In0};
       
    45 val Lim_inject = @{thm Lim_inject};
       
    46 val Inl_inject = @{thm Inl_inject};
       
    47 val Inr_inject = @{thm Inr_inject};
       
    48 val Suml_inject = @{thm Suml_inject};
       
    49 val Sumr_inject = @{thm Sumr_inject};
       
    50 
       
    51 val datatype_injI =
       
    52   @{lemma "(!!x. ALL y. f x = f y --> x = y) ==> inj f" by (simp add: inj_on_def)};
       
    53 
       
    54 
       
    55 (** proof of characteristic theorems **)
       
    56 
       
    57 fun representation_proofs (config : Old_Datatype_Aux.config)
       
    58     (dt_info : Old_Datatype_Aux.info Symtab.table) descr types_syntax constr_syntax case_names_induct
       
    59     thy =
       
    60   let
       
    61     val descr' = flat descr;
       
    62     val new_type_names = map (Binding.name_of o fst) types_syntax;
       
    63     val big_name = space_implode "_" new_type_names;
       
    64     val thy1 = Sign.add_path big_name thy;
       
    65     val big_rec_name = "rep_set_" ^ big_name;
       
    66     val rep_set_names' =
       
    67       if length descr' = 1 then [big_rec_name]
       
    68       else map (prefix (big_rec_name ^ "_") o string_of_int) (1 upto length descr');
       
    69     val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
       
    70 
       
    71     val tyvars = map (fn (_, (_, Ts, _)) => map Old_Datatype_Aux.dest_DtTFree Ts) (hd descr);
       
    72     val leafTs' = Old_Datatype_Aux.get_nonrec_types descr';
       
    73     val branchTs = Old_Datatype_Aux.get_branching_types descr';
       
    74     val branchT =
       
    75       if null branchTs then HOLogic.unitT
       
    76       else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) branchTs;
       
    77     val arities = remove (op =) 0 (Old_Datatype_Aux.get_arities descr');
       
    78     val unneeded_vars =
       
    79       subtract (op =) (fold Term.add_tfreesT (leafTs' @ branchTs) []) (hd tyvars);
       
    80     val leafTs = leafTs' @ map TFree unneeded_vars;
       
    81     val recTs = Old_Datatype_Aux.get_rec_types descr';
       
    82     val (newTs, oldTs) = chop (length (hd descr)) recTs;
       
    83     val sumT =
       
    84       if null leafTs then HOLogic.unitT
       
    85       else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) leafTs;
       
    86     val Univ_elT = HOLogic.mk_setT (Type (@{type_name Old_Datatype.node}, [sumT, branchT]));
       
    87     val UnivT = HOLogic.mk_setT Univ_elT;
       
    88     val UnivT' = Univ_elT --> HOLogic.boolT;
       
    89     val Collect = Const (@{const_name Collect}, UnivT' --> UnivT);
       
    90 
       
    91     val In0 = Const (@{const_name Old_Datatype.In0}, Univ_elT --> Univ_elT);
       
    92     val In1 = Const (@{const_name Old_Datatype.In1}, Univ_elT --> Univ_elT);
       
    93     val Leaf = Const (@{const_name Old_Datatype.Leaf}, sumT --> Univ_elT);
       
    94     val Lim = Const (@{const_name Old_Datatype.Lim}, (branchT --> Univ_elT) --> Univ_elT);
       
    95 
       
    96     (* make injections needed for embedding types in leaves *)
       
    97 
       
    98     fun mk_inj T' x =
       
    99       let
       
   100         fun mk_inj' T n i =
       
   101           if n = 1 then x
       
   102           else
       
   103             let
       
   104               val n2 = n div 2;
       
   105               val Type (_, [T1, T2]) = T;
       
   106             in
       
   107               if i <= n2
       
   108               then Const (@{const_name Inl}, T1 --> T) $ mk_inj' T1 n2 i
       
   109               else Const (@{const_name Inr}, T2 --> T) $ mk_inj' T2 (n - n2) (i - n2)
       
   110             end;
       
   111       in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) end;
       
   112 
       
   113     (* make injections for constructors *)
       
   114 
       
   115     fun mk_univ_inj ts = Balanced_Tree.access
       
   116       {left = fn t => In0 $ t,
       
   117         right = fn t => In1 $ t,
       
   118         init =
       
   119           if ts = [] then Const (@{const_name undefined}, Univ_elT)
       
   120           else foldr1 (HOLogic.mk_binop @{const_name Old_Datatype.Scons}) ts};
       
   121 
       
   122     (* function spaces *)
       
   123 
       
   124     fun mk_fun_inj T' x =
       
   125       let
       
   126         fun mk_inj T n i =
       
   127           if n = 1 then x
       
   128           else
       
   129             let
       
   130               val n2 = n div 2;
       
   131               val Type (_, [T1, T2]) = T;
       
   132               fun mkT U = (U --> Univ_elT) --> T --> Univ_elT;
       
   133             in
       
   134               if i <= n2 then Const (@{const_name Sum_Type.Suml}, mkT T1) $ mk_inj T1 n2 i
       
   135               else Const (@{const_name Sum_Type.Sumr}, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
       
   136             end;
       
   137       in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) end;
       
   138 
       
   139     fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t;
       
   140 
       
   141     (************** generate introduction rules for representing set **********)
       
   142 
       
   143     val _ = Old_Datatype_Aux.message config "Constructing representing sets ...";
       
   144 
       
   145     (* make introduction rule for a single constructor *)
       
   146 
       
   147     fun make_intr s n (i, (_, cargs)) =
       
   148       let
       
   149         fun mk_prem dt (j, prems, ts) =
       
   150           (case Old_Datatype_Aux.strip_dtyp dt of
       
   151             (dts, Old_Datatype_Aux.DtRec k) =>
       
   152               let
       
   153                 val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr') dts;
       
   154                 val free_t =
       
   155                   Old_Datatype_Aux.app_bnds (Old_Datatype_Aux.mk_Free "x" (Ts ---> Univ_elT) j)
       
   156                     (length Ts)
       
   157               in
       
   158                 (j + 1, Logic.list_all (map (pair "x") Ts,
       
   159                   HOLogic.mk_Trueprop
       
   160                     (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems,
       
   161                 mk_lim free_t Ts :: ts)
       
   162               end
       
   163           | _ =>
       
   164               let val T = Old_Datatype_Aux.typ_of_dtyp descr' dt
       
   165               in (j + 1, prems, (Leaf $ mk_inj T (Old_Datatype_Aux.mk_Free "x" T j)) :: ts) end);
       
   166 
       
   167         val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []);
       
   168         val concl = HOLogic.mk_Trueprop (Free (s, UnivT') $ mk_univ_inj ts n i);
       
   169       in Logic.list_implies (prems, concl) end;
       
   170 
       
   171     val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
       
   172       map (make_intr rep_set_name (length constrs))
       
   173         ((1 upto length constrs) ~~ constrs)) (descr' ~~ rep_set_names');
       
   174 
       
   175     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
       
   176       thy1
       
   177       |> Sign.map_naming Name_Space.conceal
       
   178       |> Inductive.add_inductive_global
       
   179           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
       
   180            coind = false, no_elim = true, no_ind = false, skip_mono = true}
       
   181           (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
       
   182           (map (fn x => (Attrib.empty_binding, x)) intr_ts) []
       
   183       ||> Sign.restore_naming thy1;
       
   184 
       
   185     (********************************* typedef ********************************)
       
   186 
       
   187     val (typedefs, thy3) = thy2
       
   188       |> Sign.parent_path
       
   189       |> fold_map
       
   190         (fn (((name, mx), tvs), c) =>
       
   191           Typedef.add_typedef_global (name, tvs, mx)
       
   192             (Collect $ Const (c, UnivT')) NONE
       
   193             (rtac exI 1 THEN rtac CollectI 1 THEN
       
   194               QUIET_BREADTH_FIRST (has_fewer_prems 1)
       
   195               (resolve_tac rep_intrs 1)))
       
   196         (types_syntax ~~ tyvars ~~ take (length newTs) rep_set_names)
       
   197       ||> Sign.add_path big_name;
       
   198 
       
   199     (*********************** definition of constructors ***********************)
       
   200 
       
   201     val big_rep_name = big_name ^ "_Rep_";
       
   202     val rep_names' = map (fn i => big_rep_name ^ string_of_int i) (1 upto length (flat (tl descr)));
       
   203     val all_rep_names =
       
   204       map (#Rep_name o #1 o #2) typedefs @
       
   205       map (Sign.full_bname thy3) rep_names';
       
   206 
       
   207     (* isomorphism declarations *)
       
   208 
       
   209     val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
       
   210       (oldTs ~~ rep_names');
       
   211 
       
   212     (* constructor definitions *)
       
   213 
       
   214     fun make_constr_def (typedef: Typedef.info) T n
       
   215         ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) =
       
   216       let
       
   217         fun constr_arg dt (j, l_args, r_args) =
       
   218           let
       
   219             val T = Old_Datatype_Aux.typ_of_dtyp descr' dt;
       
   220             val free_t = Old_Datatype_Aux.mk_Free "x" T j;
       
   221           in
       
   222             (case (Old_Datatype_Aux.strip_dtyp dt, strip_type T) of
       
   223               ((_, Old_Datatype_Aux.DtRec m), (Us, U)) =>
       
   224                 (j + 1, free_t :: l_args, mk_lim
       
   225                   (Const (nth all_rep_names m, U --> Univ_elT) $
       
   226                     Old_Datatype_Aux.app_bnds free_t (length Us)) Us :: r_args)
       
   227             | _ => (j + 1, free_t :: l_args, (Leaf $ mk_inj T free_t) :: r_args))
       
   228           end;
       
   229 
       
   230         val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []);
       
   231         val constrT = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs ---> T;
       
   232         val ({Abs_name, Rep_name, ...}, _) = typedef;
       
   233         val lhs = list_comb (Const (cname, constrT), l_args);
       
   234         val rhs = mk_univ_inj r_args n i;
       
   235         val def = Logic.mk_equals (lhs, Const (Abs_name, Univ_elT --> T) $ rhs);
       
   236         val def_name = Thm.def_name (Long_Name.base_name cname);
       
   237         val eqn =
       
   238           HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (Rep_name, T --> Univ_elT) $ lhs, rhs));
       
   239         val ([def_thm], thy') =
       
   240           thy
       
   241           |> Sign.add_consts [(cname', constrT, mx)]
       
   242           |> (Global_Theory.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
       
   243 
       
   244       in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
       
   245 
       
   246     (* constructor definitions for datatype *)
       
   247 
       
   248     fun dt_constr_defs (((((_, (_, _, constrs)), tname), typedef: Typedef.info), T), constr_syntax)
       
   249         (thy, defs, eqns, rep_congs, dist_lemmas) =
       
   250       let
       
   251         val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
       
   252         val rep_const = cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT));
       
   253         val cong' = cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong;
       
   254         val dist = cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma;
       
   255         val (thy', defs', eqns', _) =
       
   256           fold (make_constr_def typedef T (length constrs))
       
   257             (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
       
   258       in
       
   259         (Sign.parent_path thy', defs', eqns @ [eqns'],
       
   260           rep_congs @ [cong'], dist_lemmas @ [dist])
       
   261       end;
       
   262 
       
   263     val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) =
       
   264       fold dt_constr_defs
       
   265         (hd descr ~~ new_type_names ~~ map #2 typedefs ~~ newTs ~~ constr_syntax)
       
   266         (thy3 |> Sign.add_consts iso_decls |> Sign.parent_path, [], [], [], []);
       
   267 
       
   268 
       
   269     (*********** isomorphisms for new types (introduced by typedef) ***********)
       
   270 
       
   271     val _ = Old_Datatype_Aux.message config "Proving isomorphism properties ...";
       
   272 
       
   273     val collect_simp = rewrite_rule (Proof_Context.init_global thy4) [mk_meta_eq mem_Collect_eq];
       
   274 
       
   275     val newT_iso_axms = typedefs |> map (fn (_, (_, {Abs_inverse, Rep_inverse, Rep, ...})) =>
       
   276       (collect_simp Abs_inverse, Rep_inverse, collect_simp Rep));
       
   277 
       
   278     val newT_iso_inj_thms = typedefs |> map (fn (_, (_, {Abs_inject, Rep_inject, ...})) =>
       
   279       (collect_simp Abs_inject RS iffD1, Rep_inject RS iffD1));
       
   280 
       
   281     (********* isomorphisms between existing types and "unfolded" types *******)
       
   282 
       
   283     (*---------------------------------------------------------------------*)
       
   284     (* isomorphisms are defined using primrec-combinators:                 *)
       
   285     (* generate appropriate functions for instantiating primrec-combinator *)
       
   286     (*                                                                     *)
       
   287     (*   e.g.  Rep_dt_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y))    *)
       
   288     (*                                                                     *)
       
   289     (* also generate characteristic equations for isomorphisms             *)
       
   290     (*                                                                     *)
       
   291     (*   e.g.  Rep_dt_i (cons h t) = In1 (Scons (Rep_dt_j h) (Rep_dt_i t)) *)
       
   292     (*---------------------------------------------------------------------*)
       
   293 
       
   294     fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
       
   295       let
       
   296         val argTs = map (Old_Datatype_Aux.typ_of_dtyp descr') cargs;
       
   297         val T = nth recTs k;
       
   298         val rep_const = Const (nth all_rep_names k, T --> Univ_elT);
       
   299         val constr = Const (cname, argTs ---> T);
       
   300 
       
   301         fun process_arg ks' dt (i2, i2', ts, Ts) =
       
   302           let
       
   303             val T' = Old_Datatype_Aux.typ_of_dtyp descr' dt;
       
   304             val (Us, U) = strip_type T'
       
   305           in
       
   306             (case Old_Datatype_Aux.strip_dtyp dt of
       
   307               (_, Old_Datatype_Aux.DtRec j) =>
       
   308                 if member (op =) ks' j then
       
   309                   (i2 + 1, i2' + 1, ts @ [mk_lim (Old_Datatype_Aux.app_bnds
       
   310                      (Old_Datatype_Aux.mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
       
   311                    Ts @ [Us ---> Univ_elT])
       
   312                 else
       
   313                   (i2 + 1, i2', ts @ [mk_lim
       
   314                      (Const (nth all_rep_names j, U --> Univ_elT) $
       
   315                         Old_Datatype_Aux.app_bnds
       
   316                           (Old_Datatype_Aux.mk_Free "x" T' i2) (length Us)) Us], Ts)
       
   317             | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (Old_Datatype_Aux.mk_Free "x" T' i2)], Ts))
       
   318           end;
       
   319 
       
   320         val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []);
       
   321         val xs = map (uncurry (Old_Datatype_Aux.mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
       
   322         val ys = map (uncurry (Old_Datatype_Aux.mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
       
   323         val f = fold_rev lambda (xs @ ys) (mk_univ_inj ts n i);
       
   324 
       
   325         val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []);
       
   326         val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   327           (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
       
   328 
       
   329       in (fs @ [f], eqns @ [eqn], i + 1) end;
       
   330 
       
   331     (* define isomorphisms for all mutually recursive datatypes in list ds *)
       
   332 
       
   333     fun make_iso_defs ds (thy, char_thms) =
       
   334       let
       
   335         val ks = map fst ds;
       
   336         val (_, (tname, _, _)) = hd ds;
       
   337         val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
       
   338 
       
   339         fun process_dt (k, (_, _, constrs)) (fs, eqns, isos) =
       
   340           let
       
   341             val (fs', eqns', _) = fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
       
   342             val iso = (nth recTs k, nth all_rep_names k);
       
   343           in (fs', eqns', isos @ [iso]) end;
       
   344 
       
   345         val (fs, eqns, isos) = fold process_dt ds ([], [], []);
       
   346         val fTs = map fastype_of fs;
       
   347         val defs =
       
   348           map (fn (rec_name, (T, iso_name)) =>
       
   349             (Binding.name (Thm.def_name (Long_Name.base_name iso_name)),
       
   350               Logic.mk_equals (Const (iso_name, T --> Univ_elT),
       
   351                 list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
       
   352         val (def_thms, thy') =
       
   353           (Global_Theory.add_defs false o map Thm.no_attributes) defs thy;
       
   354 
       
   355         (* prove characteristic equations *)
       
   356 
       
   357         val rewrites = def_thms @ map mk_meta_eq rec_rewrites;
       
   358         val char_thms' =
       
   359           map (fn eqn => Goal.prove_sorry_global thy' [] [] eqn
       
   360             (fn {context = ctxt, ...} => EVERY [rewrite_goals_tac ctxt rewrites, rtac refl 1])) eqns;
       
   361 
       
   362       in (thy', char_thms' @ char_thms) end;
       
   363 
       
   364     val (thy5, iso_char_thms) =
       
   365       fold_rev make_iso_defs (tl descr) (Sign.add_path big_name thy4, []);
       
   366 
       
   367     (* prove isomorphism properties *)
       
   368 
       
   369     fun mk_funs_inv thy thm =
       
   370       let
       
   371         val prop = Thm.prop_of thm;
       
   372         val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
       
   373           (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop;
       
   374         val used = Term.add_tfree_names a [];
       
   375 
       
   376         fun mk_thm i =
       
   377           let
       
   378             val Ts = map (TFree o rpair @{sort type}) (Name.variant_list used (replicate i "'t"));
       
   379             val f = Free ("f", Ts ---> U);
       
   380           in
       
   381             Goal.prove_sorry_global thy [] []
       
   382               (Logic.mk_implies
       
   383                 (HOLogic.mk_Trueprop (HOLogic.list_all
       
   384                    (map (pair "x") Ts, S $ Old_Datatype_Aux.app_bnds f i)),
       
   385                  HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (Term.abs o pair "x") Ts
       
   386                    (r $ (a $ Old_Datatype_Aux.app_bnds f i)), f))))
       
   387               (fn _ => EVERY [REPEAT_DETERM_N i (rtac @{thm ext} 1),
       
   388                  REPEAT (etac allE 1), rtac thm 1, atac 1])
       
   389           end
       
   390       in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
       
   391 
       
   392     (* prove  inj Rep_dt_i  and  Rep_dt_i x : rep_set_dt_i *)
       
   393 
       
   394     val fun_congs =
       
   395       map (fn T => make_elim (Drule.instantiate' [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
       
   396 
       
   397     fun prove_iso_thms ds (inj_thms, elem_thms) =
       
   398       let
       
   399         val (_, (tname, _, _)) = hd ds;
       
   400         val induct = #induct (the (Symtab.lookup dt_info tname));
       
   401 
       
   402         fun mk_ind_concl (i, _) =
       
   403           let
       
   404             val T = nth recTs i;
       
   405             val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
       
   406             val rep_set_name = nth rep_set_names i;
       
   407             val concl1 =
       
   408               HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
       
   409                 HOLogic.mk_eq (Rep_t $ Old_Datatype_Aux.mk_Free "x" T i, Rep_t $ Bound 0) $
       
   410                   HOLogic.mk_eq (Old_Datatype_Aux.mk_Free "x" T i, Bound 0));
       
   411             val concl2 = Const (rep_set_name, UnivT') $ (Rep_t $ Old_Datatype_Aux.mk_Free "x" T i);
       
   412           in (concl1, concl2) end;
       
   413 
       
   414         val (ind_concl1, ind_concl2) = split_list (map mk_ind_concl ds);
       
   415 
       
   416         val rewrites = map mk_meta_eq iso_char_thms;
       
   417         val inj_thms' = map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) inj_thms;
       
   418 
       
   419         val inj_thm =
       
   420           Goal.prove_sorry_global thy5 [] []
       
   421             (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj ind_concl1))
       
   422             (fn {context = ctxt, ...} => EVERY
       
   423               [(Old_Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
       
   424                REPEAT (EVERY
       
   425                  [rtac allI 1, rtac impI 1,
       
   426                   Old_Datatype_Aux.exh_tac (exh_thm_of dt_info) 1,
       
   427                   REPEAT (EVERY
       
   428                     [hyp_subst_tac ctxt 1,
       
   429                      rewrite_goals_tac ctxt rewrites,
       
   430                      REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
       
   431                      (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
       
   432                      ORELSE (EVERY
       
   433                        [REPEAT (eresolve_tac (Scons_inject ::
       
   434                           map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
       
   435                         REPEAT (cong_tac 1), rtac refl 1,
       
   436                         REPEAT (atac 1 ORELSE (EVERY
       
   437                           [REPEAT (rtac @{thm ext} 1),
       
   438                            REPEAT (eresolve_tac (mp :: allE ::
       
   439                              map make_elim (Suml_inject :: Sumr_inject ::
       
   440                                Lim_inject :: inj_thms') @ fun_congs) 1),
       
   441                            atac 1]))])])])]);
       
   442 
       
   443         val inj_thms'' = map (fn r => r RS datatype_injI) (Old_Datatype_Aux.split_conj_thm inj_thm);
       
   444 
       
   445         val elem_thm =
       
   446           Goal.prove_sorry_global thy5 [] []
       
   447             (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj ind_concl2))
       
   448             (fn {context = ctxt, ...} =>
       
   449               EVERY [
       
   450                 (Old_Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
       
   451                 rewrite_goals_tac ctxt rewrites,
       
   452                 REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
       
   453                   ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
       
   454 
       
   455       in (inj_thms'' @ inj_thms, elem_thms @ Old_Datatype_Aux.split_conj_thm elem_thm) end;
       
   456 
       
   457     val (iso_inj_thms_unfolded, iso_elem_thms) =
       
   458       fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms);
       
   459     val iso_inj_thms =
       
   460       map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
       
   461 
       
   462     (* prove  rep_set_dt_i x --> x : range Rep_dt_i *)
       
   463 
       
   464     fun mk_iso_t (((set_name, iso_name), i), T) =
       
   465       let val isoT = T --> Univ_elT in
       
   466         HOLogic.imp $
       
   467           (Const (set_name, UnivT') $ Old_Datatype_Aux.mk_Free "x" Univ_elT i) $
       
   468             (if i < length newTs then @{term True}
       
   469              else HOLogic.mk_mem (Old_Datatype_Aux.mk_Free "x" Univ_elT i,
       
   470                Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
       
   471                  Const (iso_name, isoT) $ Const (@{const_abbrev UNIV}, HOLogic.mk_setT T)))
       
   472       end;
       
   473 
       
   474     val iso_t = HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj (map mk_iso_t
       
   475       (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
       
   476 
       
   477     (* all the theorems are proved by one single simultaneous induction *)
       
   478 
       
   479     val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq})) iso_inj_thms_unfolded;
       
   480 
       
   481     val iso_thms =
       
   482       if length descr = 1 then []
       
   483       else
       
   484         drop (length newTs) (Old_Datatype_Aux.split_conj_thm
       
   485           (Goal.prove_sorry_global thy5 [] [] iso_t (fn {context = ctxt, ...} => EVERY
       
   486              [(Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW
       
   487                  Object_Logic.atomize_prems_tac ctxt) 1,
       
   488               REPEAT (rtac TrueI 1),
       
   489               rewrite_goals_tac ctxt (mk_meta_eq @{thm choice_eq} ::
       
   490                 Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
       
   491               rewrite_goals_tac ctxt (map Thm.symmetric range_eqs),
       
   492               REPEAT (EVERY
       
   493                 [REPEAT (eresolve_tac ([rangeE, @{thm ex1_implies_ex} RS exE] @
       
   494                    maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
       
   495                  TRY (hyp_subst_tac ctxt 1),
       
   496                  rtac (sym RS range_eqI) 1,
       
   497                  resolve_tac iso_char_thms 1])])));
       
   498 
       
   499     val Abs_inverse_thms' =
       
   500       map #1 newT_iso_axms @
       
   501       map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp])
       
   502         iso_inj_thms_unfolded iso_thms;
       
   503 
       
   504     val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
       
   505 
       
   506     (******************* freeness theorems for constructors *******************)
       
   507 
       
   508     val _ = Old_Datatype_Aux.message config "Proving freeness of constructors ...";
       
   509 
       
   510     (* prove theorem  Rep_i (Constr_j ...) = Inj_j ...  *)
       
   511 
       
   512     fun prove_constr_rep_thm eqn =
       
   513       let
       
   514         val inj_thms = map fst newT_iso_inj_thms;
       
   515         val rewrites = @{thm o_def} :: constr_defs @ map (mk_meta_eq o #2) newT_iso_axms;
       
   516       in
       
   517         Goal.prove_sorry_global thy5 [] [] eqn
       
   518         (fn {context = ctxt, ...} => EVERY
       
   519           [resolve_tac inj_thms 1,
       
   520            rewrite_goals_tac ctxt rewrites,
       
   521            rtac refl 3,
       
   522            resolve_tac rep_intrs 2,
       
   523            REPEAT (resolve_tac iso_elem_thms 1)])
       
   524       end;
       
   525 
       
   526     (*--------------------------------------------------------------*)
       
   527     (* constr_rep_thms and rep_congs are used to prove distinctness *)
       
   528     (* of constructors.                                             *)
       
   529     (*--------------------------------------------------------------*)
       
   530 
       
   531     val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
       
   532 
       
   533     val dist_rewrites =
       
   534       map (fn (rep_thms, dist_lemma) =>
       
   535         dist_lemma :: (rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
       
   536           (constr_rep_thms ~~ dist_lemmas);
       
   537 
       
   538     fun prove_distinct_thms dist_rewrites' =
       
   539       let
       
   540         fun prove [] = []
       
   541           | prove (t :: ts) =
       
   542               let
       
   543                 val dist_thm = Goal.prove_sorry_global thy5 [] [] t (fn {context = ctxt, ...} =>
       
   544                   EVERY [simp_tac (put_simpset HOL_ss ctxt addsimps dist_rewrites') 1])
       
   545               in dist_thm :: Drule.zero_var_indexes (dist_thm RS not_sym) :: prove ts end;
       
   546       in prove end;
       
   547 
       
   548     val distinct_thms =
       
   549       map2 (prove_distinct_thms) dist_rewrites (Old_Datatype_Prop.make_distincts descr);
       
   550 
       
   551     (* prove injectivity of constructors *)
       
   552 
       
   553     fun prove_constr_inj_thm rep_thms t =
       
   554       let
       
   555         val inj_thms = Scons_inject ::
       
   556           map make_elim
       
   557             (iso_inj_thms @
       
   558               [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
       
   559                Lim_inject, Suml_inject, Sumr_inject])
       
   560       in
       
   561         Goal.prove_sorry_global thy5 [] [] t
       
   562           (fn {context = ctxt, ...} => EVERY
       
   563             [rtac iffI 1,
       
   564              REPEAT (etac conjE 2), hyp_subst_tac ctxt 2, rtac refl 2,
       
   565              dresolve_tac rep_congs 1, dtac @{thm box_equals} 1,
       
   566              REPEAT (resolve_tac rep_thms 1),
       
   567              REPEAT (eresolve_tac inj_thms 1),
       
   568              REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac @{thm ext} 1),
       
   569                REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
       
   570                atac 1]))])
       
   571       end;
       
   572 
       
   573     val constr_inject =
       
   574       map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
       
   575         (Old_Datatype_Prop.make_injs descr ~~ constr_rep_thms);
       
   576 
       
   577     val ((constr_inject', distinct_thms'), thy6) =
       
   578       thy5
       
   579       |> Sign.parent_path
       
   580       |> Old_Datatype_Aux.store_thmss "inject" new_type_names constr_inject
       
   581       ||>> Old_Datatype_Aux.store_thmss "distinct" new_type_names distinct_thms;
       
   582 
       
   583     (*************************** induction theorem ****************************)
       
   584 
       
   585     val _ = Old_Datatype_Aux.message config "Proving induction rule for datatypes ...";
       
   586 
       
   587     val Rep_inverse_thms =
       
   588       map (fn (_, iso, _) => iso RS subst) newT_iso_axms @
       
   589       map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded;
       
   590     val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded;
       
   591 
       
   592     fun mk_indrule_lemma (i, _) T =
       
   593       let
       
   594         val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ Old_Datatype_Aux.mk_Free "x" T i;
       
   595         val Abs_t =
       
   596           if i < length newTs then
       
   597             Const (#Abs_name (#1 (#2 (nth typedefs i))), Univ_elT --> T)
       
   598           else
       
   599             Const (@{const_name the_inv_into},
       
   600               [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $
       
   601             HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT);
       
   602         val prem =
       
   603           HOLogic.imp $
       
   604             (Const (nth rep_set_names i, UnivT') $ Rep_t) $
       
   605               (Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t));
       
   606         val concl =
       
   607           Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $
       
   608             Old_Datatype_Aux.mk_Free "x" T i;
       
   609       in (prem, concl) end;
       
   610 
       
   611     val (indrule_lemma_prems, indrule_lemma_concls) =
       
   612       split_list (map2 mk_indrule_lemma descr' recTs);
       
   613 
       
   614     val cert = cterm_of thy6;
       
   615 
       
   616     val indrule_lemma =
       
   617       Goal.prove_sorry_global thy6 [] []
       
   618         (Logic.mk_implies
       
   619           (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems),
       
   620            HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls)))
       
   621         (fn _ =>
       
   622           EVERY
       
   623            [REPEAT (etac conjE 1),
       
   624             REPEAT (EVERY
       
   625               [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
       
   626                etac mp 1, resolve_tac iso_elem_thms 1])]);
       
   627 
       
   628     val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
       
   629     val frees =
       
   630       if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))]
       
   631       else map (Free o apfst fst o dest_Var) Ps;
       
   632     val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
       
   633 
       
   634     val dt_induct_prop = Old_Datatype_Prop.make_ind descr;
       
   635     val dt_induct =
       
   636       Goal.prove_sorry_global thy6 []
       
   637       (Logic.strip_imp_prems dt_induct_prop)
       
   638       (Logic.strip_imp_concl dt_induct_prop)
       
   639       (fn {context = ctxt, prems, ...} =>
       
   640         EVERY
       
   641           [rtac indrule_lemma' 1,
       
   642            (Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW
       
   643               Object_Logic.atomize_prems_tac ctxt) 1,
       
   644            EVERY (map (fn (prem, r) => (EVERY
       
   645              [REPEAT (eresolve_tac Abs_inverse_thms 1),
       
   646               simp_tac (put_simpset HOL_basic_ss ctxt
       
   647                 addsimps (Thm.symmetric r :: Rep_inverse_thms')) 1,
       
   648               DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
       
   649                   (prems ~~ (constr_defs @ map mk_meta_eq iso_char_thms)))]);
       
   650 
       
   651     val ([(_, [dt_induct'])], thy7) =
       
   652       thy6
       
   653       |> Global_Theory.note_thmss ""
       
   654         [((Binding.qualify true big_name (Binding.name "induct"), [case_names_induct]),
       
   655           [([dt_induct], [])])];
       
   656   in
       
   657     ((constr_inject', distinct_thms', dt_induct'), thy7)
       
   658   end;
       
   659 
       
   660 
       
   661 
       
   662 (** datatype definition **)
       
   663 
       
   664 (* specifications *)
       
   665 
       
   666 type spec = (binding * (string * sort) list * mixfix) * (binding * typ list * mixfix) list;
       
   667 
       
   668 type spec_cmd =
       
   669   (binding * (string * string option) list * mixfix) * (binding * string list * mixfix) list;
       
   670 
       
   671 local
       
   672 
       
   673 fun parse_spec ctxt ((b, args, mx), constrs) =
       
   674   ((b, map (apsnd (Typedecl.read_constraint ctxt)) args, mx),
       
   675     constrs |> map (fn (c, Ts, mx') => (c, map (Syntax.parse_typ ctxt) Ts, mx')));
       
   676 
       
   677 fun check_specs ctxt (specs: spec list) =
       
   678   let
       
   679     fun prep_spec ((tname, args, mx), constrs) tys =
       
   680       let
       
   681         val (args', tys1) = chop (length args) tys;
       
   682         val (constrs', tys3) = (constrs, tys1) |-> fold_map (fn (cname, cargs, mx') => fn tys2 =>
       
   683           let val (cargs', tys3) = chop (length cargs) tys2;
       
   684           in ((cname, cargs', mx'), tys3) end);
       
   685       in (((tname, map dest_TFree args', mx), constrs'), tys3) end;
       
   686 
       
   687     val all_tys =
       
   688       specs |> maps (fn ((_, args, _), cs) => map TFree args @ maps #2 cs)
       
   689       |> Syntax.check_typs ctxt;
       
   690 
       
   691   in #1 (fold_map prep_spec specs all_tys) end;
       
   692 
       
   693 fun prep_specs parse raw_specs thy =
       
   694   let
       
   695     val ctxt = thy
       
   696       |> Sign.add_types_global (map (fn ((b, args, mx), _) => (b, length args, mx)) raw_specs)
       
   697       |> Proof_Context.init_global
       
   698       |> fold (fn ((_, args, _), _) => fold (fn (a, _) =>
       
   699           Variable.declare_typ (TFree (a, dummyS))) args) raw_specs;
       
   700     val specs = check_specs ctxt (map (parse ctxt) raw_specs);
       
   701   in (specs, ctxt) end;
       
   702 
       
   703 in
       
   704 
       
   705 val read_specs = prep_specs parse_spec;
       
   706 val check_specs = prep_specs (K I);
       
   707 
       
   708 end;
       
   709 
       
   710 
       
   711 (* main commands *)
       
   712 
       
   713 fun gen_add_datatype prep_specs config raw_specs thy =
       
   714   let
       
   715     val _ = Theory.requires thy (Context.theory_name @{theory}) "datatype definitions";
       
   716 
       
   717     val (dts, spec_ctxt) = prep_specs raw_specs thy;
       
   718     val ((_, tyvars, _), _) :: _ = dts;
       
   719     val string_of_tyvar = Syntax.string_of_typ spec_ctxt o TFree;
       
   720 
       
   721     val (new_dts, types_syntax) = dts |> map (fn ((tname, tvs, mx), _) =>
       
   722       let val full_tname = Sign.full_name thy tname in
       
   723         (case duplicates (op =) tvs of
       
   724           [] =>
       
   725             if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
       
   726             else error "Mutually recursive datatypes must have same type parameters"
       
   727         | dups =>
       
   728             error ("Duplicate parameter(s) for datatype " ^ Binding.print tname ^
       
   729               " : " ^ commas (map string_of_tyvar dups)))
       
   730       end) |> split_list;
       
   731     val dt_names = map fst new_dts;
       
   732 
       
   733     val _ =
       
   734       (case duplicates (op =) (map fst new_dts) of
       
   735         [] => ()
       
   736       | dups => error ("Duplicate datatypes: " ^ commas_quote dups));
       
   737 
       
   738     fun prep_dt_spec ((tname, tvs, _), constrs) (dts', constr_syntax, i) =
       
   739       let
       
   740         fun prep_constr (cname, cargs, mx) (constrs, constr_syntax') =
       
   741           let
       
   742             val _ =
       
   743               (case subtract (op =) tvs (fold Term.add_tfreesT cargs []) of
       
   744                 [] => ()
       
   745               | vs => error ("Extra type variables on rhs: " ^ commas (map string_of_tyvar vs)));
       
   746             val c = Sign.full_name_path thy (Binding.name_of tname) cname;
       
   747           in
       
   748             (constrs @ [(c, map (Old_Datatype_Aux.dtyp_of_typ new_dts) cargs)],
       
   749               constr_syntax' @ [(cname, mx)])
       
   750           end handle ERROR msg =>
       
   751             cat_error msg ("The error above occurred in constructor " ^ Binding.print cname ^
       
   752               " of datatype " ^ Binding.print tname);
       
   753 
       
   754         val (constrs', constr_syntax') = fold prep_constr constrs ([], []);
       
   755       in
       
   756         (case duplicates (op =) (map fst constrs') of
       
   757           [] =>
       
   758             (dts' @ [(i, (Sign.full_name thy tname, map Old_Datatype_Aux.DtTFree tvs, constrs'))],
       
   759               constr_syntax @ [constr_syntax'], i + 1)
       
   760         | dups =>
       
   761             error ("Duplicate constructors " ^ commas_quote dups ^
       
   762               " in datatype " ^ Binding.print tname))
       
   763       end;
       
   764 
       
   765     val (dts', constr_syntax, i) = fold prep_dt_spec dts ([], [], 0);
       
   766 
       
   767     val dt_info = Old_Datatype_Data.get_all thy;
       
   768     val (descr, _) = Old_Datatype_Aux.unfold_datatypes spec_ctxt dts' dt_info dts' i;
       
   769     val _ =
       
   770       Old_Datatype_Aux.check_nonempty descr
       
   771         handle (exn as Old_Datatype_Aux.Datatype_Empty s) =>
       
   772           if #strict config then error ("Nonemptiness check failed for datatype " ^ quote s)
       
   773           else reraise exn;
       
   774 
       
   775     val _ =
       
   776       Old_Datatype_Aux.message config
       
   777         ("Constructing datatype(s) " ^ commas_quote (map (Binding.name_of o #1 o #1) dts));
       
   778   in
       
   779     thy
       
   780     |> representation_proofs config dt_info descr types_syntax constr_syntax
       
   781       (Old_Datatype_Data.mk_case_names_induct (flat descr))
       
   782     |-> (fn (inject, distinct, induct) =>
       
   783       Old_Rep_Datatype.derive_datatype_props config dt_names descr induct inject distinct)
       
   784   end;
       
   785 
       
   786 val add_datatype = gen_add_datatype check_specs;
       
   787 val add_datatype_cmd = gen_add_datatype read_specs;
       
   788 
       
   789 
       
   790 (* outer syntax *)
       
   791 
       
   792 val spec_cmd =
       
   793   Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
       
   794   (@{keyword "="} |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat Parse.typ -- Parse.opt_mixfix))
       
   795   >> (fn (((vs, t), mx), cons) => ((t, vs, mx), map Parse.triple1 cons));
       
   796 
       
   797 val _ =
       
   798   Outer_Syntax.command @{command_spec "datatype"} "define inductive datatypes"
       
   799     (Parse.and_list1 spec_cmd
       
   800       >> (Toplevel.theory o (snd oo add_datatype_cmd Old_Datatype_Aux.default_config)));
       
   801 
       
   802 end;