# HG changeset patch # User haftmann # Date 1259577769 -3600 # Node ID f94fb13ecbb3739f6ad080bb3546422d6e5aeb0c # Parent e191b400a8e0a92731e853a66743d9d3f2bdfe03 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Datatype.thy --- a/src/HOL/Datatype.thy Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Datatype.thy Mon Nov 30 11:42:49 2009 +0100 @@ -1,16 +1,13 @@ (* Title: HOL/Datatype.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer and Markus Wenzel, TU Muenchen - -Could <*> be generalized to a general summation (Sigma)? *) -header {* Analogues of the Cartesian Product and Disjoint Sum for Datatypes *} +header {* Datatype package: constructing datatypes from Cartesian Products and Disjoint Sums *} theory Datatype imports Product_Type Sum_Type Nat uses - ("Tools/Datatype/datatype_rep_proofs.ML") ("Tools/Datatype/datatype.ML") ("Tools/inductive_realizer.ML") ("Tools/Datatype/datatype_realizer.ML") @@ -520,13 +517,12 @@ hide (open) type node item hide (open) const Push Node Atom Leaf Numb Lim Split Case -use "Tools/Datatype/datatype_rep_proofs.ML" use "Tools/Datatype/datatype.ML" use "Tools/inductive_realizer.ML" setup InductiveRealizer.setup use "Tools/Datatype/datatype_realizer.ML" -setup DatatypeRealizer.setup +setup Datatype_Realizer.setup end diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Inductive.thy Mon Nov 30 11:42:49 2009 +0100 @@ -280,12 +280,12 @@ use "Tools/Datatype/datatype_data.ML" setup Datatype_Data.setup +use "Tools/Datatype/datatype_codegen.ML" +setup Datatype_Codegen.setup + use "Tools/old_primrec.ML" use "Tools/primrec.ML" -use "Tools/Datatype/datatype_codegen.ML" -setup DatatypeCodegen.setup - use "Tools/inductive_codegen.ML" setup InductiveCodegen.setup @@ -301,7 +301,7 @@ fun fun_tr ctxt [cs] = let val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT); - val ft = DatatypeCase.case_tr true Datatype_Data.info_of_constr + val ft = Datatype_Case.case_tr true Datatype_Data.info_of_constr ctxt [x, cs] in lambda x ft end in [("_lam_pats_syntax", fun_tr)] end diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/IsaMakefile Mon Nov 30 11:42:49 2009 +0100 @@ -171,7 +171,6 @@ Tools/Datatype/datatype_data.ML \ Tools/Datatype/datatype_prop.ML \ Tools/Datatype/datatype_realizer.ML \ - Tools/Datatype/datatype_rep_proofs.ML \ Tools/Datatype/datatype.ML \ Tools/dseq.ML \ Tools/Function/context_tree.ML \ diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/List.thy --- a/src/HOL/List.thy Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/List.thy Mon Nov 30 11:42:49 2009 +0100 @@ -366,7 +366,7 @@ val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN $ NilC; val cs = Syntax.const "_case2" $ case1 $ case2 - val ft = DatatypeCase.case_tr false Datatype.info_of_constr + val ft = Datatype_Case.case_tr false Datatype.info_of_constr ctxt [x, cs] in lambda x ft end; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Nominal/nominal_datatype.ML Mon Nov 30 11:42:49 2009 +0100 @@ -37,7 +37,7 @@ val Collect_False_empty = @{thm empty_def [THEN sym, THEN eq_reflection]}; val empty_iff = thm "empty_iff"; -open DatatypeAux; +open Datatype_Aux; open NominalAtoms; (** FIXME: Datatype should export this function **) @@ -56,7 +56,7 @@ fun induct_cases descr = - DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr)); + Datatype_Prop.indexify_names (maps (dt_cases descr) (map #2 descr)); fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i)); @@ -258,7 +258,7 @@ val perm_types = map (fn (i, _) => let val T = nth_dtyp i in permT --> T --> T end) descr; - val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) => + val perm_names' = Datatype_Prop.indexify_names (map (fn (i, _) => "perm_" ^ name_of_typ (nth_dtyp i)) descr); val perm_names = replicate (length new_type_names) "Nominal.perm" @ map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names)); @@ -270,7 +270,7 @@ in map (fn (cname, dts) => let val Ts = map (typ_of_dtyp descr sorts) dts; - val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts); + val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); val args = map Free (names ~~ Ts); val c = Const (cname, Ts ---> T); fun perm_arg (dt, x) = @@ -307,7 +307,7 @@ val _ = warning ("length descr: " ^ string_of_int (length descr)); val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); - val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); + val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types); val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def"; val unfolded_perm_eq_thms = @@ -502,10 +502,10 @@ val _ = warning "representing sets"; - val rep_set_names = DatatypeProp.indexify_names + val rep_set_names = Datatype_Prop.indexify_names (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr); val big_rep_name = - space_implode "_" (DatatypeProp.indexify_names (map_filter + space_implode "_" (Datatype_Prop.indexify_names (map_filter (fn (i, ("Nominal.noption", _, _)) => NONE | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; val _ = warning ("big_rep_name: " ^ big_rep_name); @@ -867,7 +867,7 @@ (* prove distinctness theorems *) - val distinct_props = DatatypeProp.make_distincts descr' sorts; + val distinct_props = Datatype_Prop.make_distincts descr' sorts; val dist_rewrites = map2 (fn rep_thms => fn dist_lemma => dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]) constr_rep_thmss dist_lemmas; @@ -1067,7 +1067,7 @@ val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; - val dt_induct_prop = DatatypeProp.make_ind descr' sorts; + val dt_induct_prop = Datatype_Prop.make_ind descr' sorts; val dt_induct = Goal.prove_global thy8 [] (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) (fn {prems, ...} => EVERY @@ -1085,7 +1085,7 @@ val _ = warning "proving finite support for the new datatype"; - val indnames = DatatypeProp.make_tnames recTs; + val indnames = Datatype_Prop.make_tnames recTs; val abs_supp = PureThy.get_thms thy8 "abs_supp"; val supp_atm = PureThy.get_thms thy8 "supp_atm"; @@ -1120,12 +1120,12 @@ PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>> PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||> Sign.parent_path ||>> - DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> - DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> - DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> - DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>> - DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>> - DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> + Datatype_Aux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> + Datatype_Aux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> + Datatype_Aux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> + Datatype_Aux.store_thmss "inject" new_type_names inject_thms ||>> + Datatype_Aux.store_thmss "supp" new_type_names supp_thms ||>> + Datatype_Aux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> fold (fn (atom, ths) => fn thy => let val class = fs_class_of thy atom; @@ -1145,7 +1145,7 @@ val fsT' = TFree ("'n", HOLogic.typeS); val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) - (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); + (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); fun make_pred fsT i T = Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT); @@ -1167,7 +1167,7 @@ val recs = filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr'' sorts) cargs; val recTs' = map (typ_of_dtyp descr'' sorts) recs; - val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); + val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts); val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = tnames ~~ Ts; val frees' = partition_cargs idxs frees; @@ -1196,7 +1196,7 @@ map (make_ind_prem fsT (fn T => fn t => fn u => fresh_const T fsT $ t $ u) i T) (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs); - val tnames = DatatypeProp.make_tnames recTs; + val tnames = Datatype_Prop.make_tnames recTs; val zs = Name.variant_list tnames (replicate (length descr'') "z"); val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") (map (fn ((((i, _), T), tname), z) => @@ -1220,7 +1220,7 @@ val induct' = Logic.list_implies (ind_prems', ind_concl'); val aux_ind_vars = - (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~ + (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~ map mk_permT dt_atomTs) @ [("z", fsT')]; val aux_ind_Ts = rev (map snd aux_ind_vars); val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") @@ -1418,7 +1418,7 @@ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; - val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used; + val (rec_result_Ts', rec_fn_Ts') = Datatype_Prop.make_primrec_Ts descr' sorts used; val rec_sort = if null dt_atomTs then HOLogic.typeS else Sign.certify_sort thy10 pt_cp_sort; @@ -1664,10 +1664,10 @@ val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns); val fun_tupleT = fastype_of fun_tuple; val rec_unique_frees = - DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs; + Datatype_Prop.indexify_names (replicate (length recTs) "x") ~~ recTs; val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees; val rec_unique_frees' = - DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; + Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; val rec_unique_concls = map (fn ((x, U), R) => Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $ Abs ("y", U, R $ Free x $ Bound 0)) @@ -2048,7 +2048,7 @@ resolve_tac rec_intrs 1, REPEAT (solve (prems @ rec_total_thms) prems 1)]) end) (rec_eq_prems ~~ - DatatypeProp.make_primrecs new_type_names descr' sorts thy12); + Datatype_Prop.make_primrecs new_type_names descr' sorts thy12); val dt_infos = map_index (make_dt_info pdescr sorts induct reccomb_names rec_thms) (descr1 ~~ distinct_thms ~~ inject_thms); diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive.ML Mon Nov 30 11:42:49 2009 +0100 @@ -247,7 +247,7 @@ end) prems); val ind_vars = - (DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~ + (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); @@ -647,7 +647,7 @@ val thss = map (fn atom => let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) in map (fn th => zero_var_indexes (th RS mp)) - (DatatypeAux.split_conj_thm (Goal.prove ctxt' [] [] + (Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => let val (h, ts) = strip_comb p; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive2.ML Mon Nov 30 11:42:49 2009 +0100 @@ -263,7 +263,7 @@ in abs_params params' prem end) prems); val ind_vars = - (DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~ + (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Nominal/nominal_primrec.ML Mon Nov 30 11:42:49 2009 +0100 @@ -21,7 +21,7 @@ structure NominalPrimrec : NOMINAL_PRIMREC = struct -open DatatypeAux; +open Datatype_Aux; exception RecError of string; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype.ML --- a/src/HOL/Tools/Datatype/datatype.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,19 +1,748 @@ -(* Title: HOL/Tools/datatype.ML +(* Title: HOL/Tools/Datatype/datatype.ML Author: Stefan Berghofer, TU Muenchen -Datatype package interface for Isabelle/HOL. +Datatype package: definitional introduction of datatypes +with proof of characteristic theorems: injectivity / distinctness +of constructors and induction. Main interface to datatypes +after full bootstrap of datatype package. *) signature DATATYPE = sig include DATATYPE_DATA - include DATATYPE_REP_PROOFS + val add_datatype : config -> string list -> (string list * binding * mixfix * + (binding * typ list * mixfix) list) list -> theory -> string list * theory + val datatype_cmd : string list -> (string list * binding * mixfix * + (binding * string list * mixfix) list) list -> theory -> theory end; -structure Datatype = +structure Datatype : DATATYPE = struct +(** auxiliary **) + +open Datatype_Aux; open Datatype_Data; -open DatatypeRepProofs; + +val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); + +val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; + +fun exh_thm_of (dt_info : info Symtab.table) tname = + #exhaust (the (Symtab.lookup dt_info tname)); + +val node_name = @{type_name "Datatype.node"}; +val In0_name = @{const_name "Datatype.In0"}; +val In1_name = @{const_name "Datatype.In1"}; +val Scons_name = @{const_name "Datatype.Scons"}; +val Leaf_name = @{const_name "Datatype.Leaf"}; +val Numb_name = @{const_name "Datatype.Numb"}; +val Lim_name = @{const_name "Datatype.Lim"}; +val Suml_name = @{const_name "Sum_Type.Suml"}; +val Sumr_name = @{const_name "Sum_Type.Sumr"}; + +val In0_inject = @{thm In0_inject}; +val In1_inject = @{thm In1_inject}; +val Scons_inject = @{thm Scons_inject}; +val Leaf_inject = @{thm Leaf_inject}; +val In0_eq = @{thm In0_eq}; +val In1_eq = @{thm In1_eq}; +val In0_not_In1 = @{thm In0_not_In1}; +val In1_not_In0 = @{thm In1_not_In0}; +val Lim_inject = @{thm Lim_inject}; +val Inl_inject = @{thm Inl_inject}; +val Inr_inject = @{thm Inr_inject}; +val Suml_inject = @{thm Suml_inject}; +val Sumr_inject = @{thm Sumr_inject}; + + + +(** proof of characteristic theorems **) + +fun representation_proofs (config : config) (dt_info : info Symtab.table) + new_type_names descr sorts types_syntax constr_syntax case_names_induct thy = + let + val descr' = flat descr; + val big_name = space_implode "_" new_type_names; + val thy1 = Sign.add_path big_name thy; + val big_rec_name = big_name ^ "_rep_set"; + val rep_set_names' = + (if length descr' = 1 then [big_rec_name] else + (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) + (1 upto (length descr')))); + val rep_set_names = map (Sign.full_bname thy1) rep_set_names'; + + val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr); + val leafTs' = get_nonrec_types descr' sorts; + val branchTs = get_branching_types descr' sorts; + val branchT = if null branchTs then HOLogic.unitT + else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs; + val arities = remove (op =) 0 (get_arities descr'); + val unneeded_vars = + subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars); + val leafTs = leafTs' @ map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars; + val recTs = get_rec_types descr' sorts; + val (newTs, oldTs) = chop (length (hd descr)) recTs; + val sumT = if null leafTs then HOLogic.unitT + else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) leafTs; + val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT])); + val UnivT = HOLogic.mk_setT Univ_elT; + val UnivT' = Univ_elT --> HOLogic.boolT; + val Collect = Const (@{const_name Collect}, UnivT' --> UnivT); + + val In0 = Const (In0_name, Univ_elT --> Univ_elT); + val In1 = Const (In1_name, Univ_elT --> Univ_elT); + val Leaf = Const (Leaf_name, sumT --> Univ_elT); + val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT); + + (* make injections needed for embedding types in leaves *) + + fun mk_inj T' x = + let + fun mk_inj' T n i = + if n = 1 then x else + let val n2 = n div 2; + val Type (_, [T1, T2]) = T + in + if i <= n2 then + Const (@{const_name "Sum_Type.Inl"}, T1 --> T) $ (mk_inj' T1 n2 i) + else + Const (@{const_name "Sum_Type.Inr"}, T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) + end + in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) + end; + + (* make injections for constructors *) + + fun mk_univ_inj ts = Balanced_Tree.access + {left = fn t => In0 $ t, + right = fn t => In1 $ t, + init = + if ts = [] then Const (@{const_name undefined}, Univ_elT) + else foldr1 (HOLogic.mk_binop Scons_name) ts}; + + (* function spaces *) + + fun mk_fun_inj T' x = + let + fun mk_inj T n i = + if n = 1 then x else + let + val n2 = n div 2; + val Type (_, [T1, T2]) = T; + fun mkT U = (U --> Univ_elT) --> T --> Univ_elT + in + if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i + else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2) + end + in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) + end; + + fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t; + + (************** generate introduction rules for representing set **********) + + val _ = message config "Constructing representing sets ..."; + + (* make introduction rule for a single constructor *) + + fun make_intr s n (i, (_, cargs)) = + let + fun mk_prem dt (j, prems, ts) = + (case strip_dtyp dt of + (dts, DtRec k) => + let + val Ts = map (typ_of_dtyp descr' sorts) dts; + val free_t = + app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts) + in (j + 1, list_all (map (pair "x") Ts, + HOLogic.mk_Trueprop + (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems, + mk_lim free_t Ts :: ts) + end + | _ => + let val T = typ_of_dtyp descr' sorts dt + in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts) + end); + + val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []); + val concl = HOLogic.mk_Trueprop + (Free (s, UnivT') $ mk_univ_inj ts n i) + in Logic.list_implies (prems, concl) + end; + + val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) => + map (make_intr rep_set_name (length constrs)) + ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names'); + + val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) = + thy1 + |> Sign.map_naming Name_Space.conceal + |> Inductive.add_inductive_global + {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, + coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false} + (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') [] + (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] + ||> Sign.restore_naming thy1 + ||> Theory.checkpoint; + + (********************************* typedef ********************************) + + val (typedefs, thy3) = thy2 |> + Sign.parent_path |> + fold_map (fn ((((name, mx), tvs), c), name') => + Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx) + (Collect $ Const (c, UnivT')) NONE + (rtac exI 1 THEN rtac CollectI 1 THEN + QUIET_BREADTH_FIRST (has_fewer_prems 1) + (resolve_tac rep_intrs 1))) + (types_syntax ~~ tyvars ~~ + (take (length newTs) rep_set_names) ~~ new_type_names) ||> + Sign.add_path big_name; + + (*********************** definition of constructors ***********************) + + val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_"; + val rep_names = map (curry op ^ "Rep_") new_type_names; + val rep_names' = map (fn i => big_rep_name ^ (string_of_int i)) + (1 upto (length (flat (tl descr)))); + val all_rep_names = map (Sign.intern_const thy3) rep_names @ + map (Sign.full_bname thy3) rep_names'; + + (* isomorphism declarations *) + + val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn)) + (oldTs ~~ rep_names'); + + (* constructor definitions *) + + fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) = + let + fun constr_arg dt (j, l_args, r_args) = + let val T = typ_of_dtyp descr' sorts dt; + val free_t = mk_Free "x" T j + in (case (strip_dtyp dt, strip_type T) of + ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim + (Const (nth all_rep_names m, U --> Univ_elT) $ + app_bnds free_t (length Us)) Us :: r_args) + | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args)) + end; + + val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); + val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T; + val abs_name = Sign.intern_const thy ("Abs_" ^ tname); + val rep_name = Sign.intern_const thy ("Rep_" ^ tname); + val lhs = list_comb (Const (cname, constrT), l_args); + val rhs = mk_univ_inj r_args n i; + val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs); + val def_name = Long_Name.base_name cname ^ "_def"; + val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const (rep_name, T --> Univ_elT) $ lhs, rhs)); + val ([def_thm], thy') = + thy + |> Sign.add_consts_i [(cname', constrT, mx)] + |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]; + + in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end; + + (* constructor definitions for datatype *) + + fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax) + (thy, defs, eqns, rep_congs, dist_lemmas) = + let + val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong; + val rep_const = cterm_of thy + (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT)); + val cong' = + Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong); + val dist = + Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); + val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs)) + (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1); + in + (Sign.parent_path thy', defs', eqns @ [eqns'], + rep_congs @ [cong'], dist_lemmas @ [dist]) + end; + + val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = + fold dt_constr_defs + (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax) + (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []); + + + (*********** isomorphisms for new types (introduced by typedef) ***********) + + val _ = message config "Proving isomorphism properties ..."; + + val newT_iso_axms = map (fn (_, td) => + (collect_simp (#Abs_inverse td), #Rep_inverse td, + collect_simp (#Rep td))) typedefs; + + val newT_iso_inj_thms = map (fn (_, td) => + (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs; + + (********* isomorphisms between existing types and "unfolded" types *******) + + (*---------------------------------------------------------------------*) + (* isomorphisms are defined using primrec-combinators: *) + (* generate appropriate functions for instantiating primrec-combinator *) + (* *) + (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *) + (* *) + (* also generate characteristic equations for isomorphisms *) + (* *) + (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *) + (*---------------------------------------------------------------------*) + + fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) = + let + val argTs = map (typ_of_dtyp descr' sorts) cargs; + val T = nth recTs k; + val rep_name = nth all_rep_names k; + val rep_const = Const (rep_name, T --> Univ_elT); + val constr = Const (cname, argTs ---> T); + + fun process_arg ks' dt (i2, i2', ts, Ts) = + let + val T' = typ_of_dtyp descr' sorts dt; + val (Us, U) = strip_type T' + in (case strip_dtyp dt of + (_, DtRec j) => if j mem ks' then + (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds + (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us], + Ts @ [Us ---> Univ_elT]) + else + (i2 + 1, i2', ts @ [mk_lim + (Const (nth all_rep_names j, U --> Univ_elT) $ + app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts) + | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts)) + end; + + val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []); + val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1))); + val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1))); + val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i); + + val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []); + val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq + (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i)) + + in (fs @ [f], eqns @ [eqn], i + 1) end; + + (* define isomorphisms for all mutually recursive datatypes in list ds *) + + fun make_iso_defs ds (thy, char_thms) = + let + val ks = map fst ds; + val (_, (tname, _, _)) = hd ds; + val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname); + + fun process_dt (k, (tname, _, constrs)) (fs, eqns, isos) = + let + val (fs', eqns', _) = + fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1); + val iso = (nth recTs k, nth all_rep_names k) + in (fs', eqns', isos @ [iso]) end; + + val (fs, eqns, isos) = fold process_dt ds ([], [], []); + val fTs = map fastype_of fs; + val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"), + Logic.mk_equals (Const (iso_name, T --> Univ_elT), + list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos); + val (def_thms, thy') = + apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy); + + (* prove characteristic equations *) + + val rewrites = def_thms @ (map mk_meta_eq rec_rewrites); + val char_thms' = map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn + (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; + + in (thy', char_thms' @ char_thms) end; + + val (thy5, iso_char_thms) = apfst Theory.checkpoint (fold_rev make_iso_defs + (tl descr) (Sign.add_path big_name thy4, [])); + + (* prove isomorphism properties *) + + fun mk_funs_inv thy thm = + let + val prop = Thm.prop_of thm; + val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $ + (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop; + val used = OldTerm.add_term_tfree_names (a, []); + + fun mk_thm i = + let + val Ts = map (TFree o rpair HOLogic.typeS) + (Name.variant_list used (replicate i "'t")); + val f = Free ("f", Ts ---> U) + in Skip_Proof.prove_global thy [] [] (Logic.mk_implies + (HOLogic.mk_Trueprop (HOLogic.list_all + (map (pair "x") Ts, S $ app_bnds f i)), + HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts, + r $ (a $ app_bnds f i)), f)))) + (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1), + REPEAT (etac allE 1), rtac thm 1, atac 1]) + end + in map (fn r => r RS subst) (thm :: map mk_thm arities) end; + + (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *) + + val fun_congs = map (fn T => make_elim (Drule.instantiate' + [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs; + + fun prove_iso_thms ds (inj_thms, elem_thms) = + let + val (_, (tname, _, _)) = hd ds; + val induct = (#induct o the o Symtab.lookup dt_info) tname; + + fun mk_ind_concl (i, _) = + let + val T = nth recTs i; + val Rep_t = Const (nth all_rep_names i, T --> Univ_elT); + val rep_set_name = nth rep_set_names i + in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $ + HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $ + HOLogic.mk_eq (mk_Free "x" T i, Bound 0)), + Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i)) + end; + + val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds); + + val rewrites = map mk_meta_eq iso_char_thms; + val inj_thms' = map snd newT_iso_inj_thms @ + map (fn r => r RS @{thm injD}) inj_thms; + + val inj_thm = Skip_Proof.prove_global thy5 [] [] + (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY + [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + REPEAT (EVERY + [rtac allI 1, rtac impI 1, + exh_tac (exh_thm_of dt_info) 1, + REPEAT (EVERY + [hyp_subst_tac 1, + rewrite_goals_tac rewrites, + REPEAT (dresolve_tac [In0_inject, In1_inject] 1), + (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1) + ORELSE (EVERY + [REPEAT (eresolve_tac (Scons_inject :: + map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1), + REPEAT (cong_tac 1), rtac refl 1, + REPEAT (atac 1 ORELSE (EVERY + [REPEAT (rtac ext 1), + REPEAT (eresolve_tac (mp :: allE :: + map make_elim (Suml_inject :: Sumr_inject :: + Lim_inject :: inj_thms') @ fun_congs) 1), + atac 1]))])])])]); + + val inj_thms'' = map (fn r => r RS @{thm datatype_injI}) + (split_conj_thm inj_thm); + + val elem_thm = + Skip_Proof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2)) + (fn _ => + EVERY [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + rewrite_goals_tac rewrites, + REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW + ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]); + + in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm)) + end; + + val (iso_inj_thms_unfolded, iso_elem_thms) = + fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms); + val iso_inj_thms = map snd newT_iso_inj_thms @ + map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded; + + (* prove dt_rep_set_i x --> x : range dt_Rep_i *) + + fun mk_iso_t (((set_name, iso_name), i), T) = + let val isoT = T --> Univ_elT + in HOLogic.imp $ + (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $ + (if i < length newTs then HOLogic.true_const + else HOLogic.mk_mem (mk_Free "x" Univ_elT i, + Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $ + Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T))) + end; + + val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t + (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs))); + + (* all the theorems are proved by one single simultaneous induction *) + + val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq})) + iso_inj_thms_unfolded; + + val iso_thms = if length descr = 1 then [] else + drop (length newTs) (split_conj_thm + (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY + [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + REPEAT (rtac TrueI 1), + rewrite_goals_tac (mk_meta_eq choice_eq :: + symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs), + rewrite_goals_tac (map symmetric range_eqs), + REPEAT (EVERY + [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @ + maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1), + TRY (hyp_subst_tac 1), + rtac (sym RS range_eqI) 1, + resolve_tac iso_char_thms 1])]))); + + val Abs_inverse_thms' = + map #1 newT_iso_axms @ + map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp]) + iso_inj_thms_unfolded iso_thms; + + val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms'; + + (******************* freeness theorems for constructors *******************) + + val _ = message config "Proving freeness of constructors ..."; + + (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *) + + fun prove_constr_rep_thm eqn = + let + val inj_thms = map fst newT_iso_inj_thms; + val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms) + in Skip_Proof.prove_global thy5 [] [] eqn (fn _ => EVERY + [resolve_tac inj_thms 1, + rewrite_goals_tac rewrites, + rtac refl 3, + resolve_tac rep_intrs 2, + REPEAT (resolve_tac iso_elem_thms 1)]) + end; + + (*--------------------------------------------------------------*) + (* constr_rep_thms and rep_congs are used to prove distinctness *) + (* of constructors. *) + (*--------------------------------------------------------------*) + + val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns; + + val dist_rewrites = map (fn (rep_thms, dist_lemma) => + dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) + (constr_rep_thms ~~ dist_lemmas); + + fun prove_distinct_thms dist_rewrites' (k, ts) = + let + fun prove [] = [] + | prove (t :: ts) = + let + val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ => + EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1]) + in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end; + in prove ts end; + + val distinct_thms = map2 (prove_distinct_thms) + dist_rewrites (Datatype_Prop.make_distincts descr sorts); + + (* prove injectivity of constructors *) + + fun prove_constr_inj_thm rep_thms t = + let val inj_thms = Scons_inject :: (map make_elim + (iso_inj_thms @ + [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject, + Lim_inject, Suml_inject, Sumr_inject])) + in Skip_Proof.prove_global thy5 [] [] t (fn _ => EVERY + [rtac iffI 1, + REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2, + dresolve_tac rep_congs 1, dtac box_equals 1, + REPEAT (resolve_tac rep_thms 1), + REPEAT (eresolve_tac inj_thms 1), + REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1), + REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1), + atac 1]))]) + end; + + val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts) + ((Datatype_Prop.make_injs descr sorts) ~~ constr_rep_thms); + + val ((constr_inject', distinct_thms'), thy6) = + thy5 + |> Sign.parent_path + |> store_thmss "inject" new_type_names constr_inject + ||>> store_thmss "distinct" new_type_names distinct_thms; + + (*************************** induction theorem ****************************) + + val _ = message config "Proving induction rule for datatypes ..."; + + val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @ + (map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded); + val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded; + + fun mk_indrule_lemma ((i, _), T) (prems, concls) = + let + val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ + mk_Free "x" T i; + + val Abs_t = if i < length newTs then + Const (Sign.intern_const thy6 + ("Abs_" ^ (nth new_type_names i)), Univ_elT --> T) + else Const (@{const_name the_inv_into}, + [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $ + HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT) + + in (prems @ [HOLogic.imp $ + (Const (nth rep_set_names i, UnivT') $ Rep_t) $ + (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], + concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) + end; + + val (indrule_lemma_prems, indrule_lemma_concls) = + fold mk_indrule_lemma (descr' ~~ recTs) ([], []); + + val cert = cterm_of thy6; + + val indrule_lemma = Skip_Proof.prove_global thy6 [] [] + (Logic.mk_implies + (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), + HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY + [REPEAT (etac conjE 1), + REPEAT (EVERY + [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1, + etac mp 1, resolve_tac iso_elem_thms 1])]); + + val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); + val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else + map (Free o apfst fst o dest_Var) Ps; + val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma; + + val dt_induct_prop = Datatype_Prop.make_ind descr sorts; + val dt_induct = Skip_Proof.prove_global thy6 [] + (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) + (fn {prems, ...} => EVERY + [rtac indrule_lemma' 1, + (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + EVERY (map (fn (prem, r) => (EVERY + [REPEAT (eresolve_tac Abs_inverse_thms 1), + simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1, + DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) + (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]); + + val ([dt_induct'], thy7) = + thy6 + |> Sign.add_path big_name + |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] + ||> Sign.parent_path + ||> Theory.checkpoint; + + in + ((constr_inject', distinct_thms', dt_induct'), thy7) + end; + + + +(** definitional introduction of datatypes **) + +fun gen_add_datatype prep_typ config new_type_names dts thy = + let + val _ = Theory.requires thy "Datatype" "datatype definitions"; + + (* this theory is used just for parsing *) + val tmp_thy = thy |> + Theory.copy |> + Sign.add_types (map (fn (tvs, tname, mx, _) => + (tname, length tvs, mx)) dts); + + val (tyvars, _, _, _)::_ = dts; + val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) => + let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname) + in + (case duplicates (op =) tvs of + [] => + if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx)) + else error ("Mutually recursive datatypes must have same type parameters") + | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^ + " : " ^ commas dups)) + end) dts); + val dt_names = map fst new_dts; + + val _ = + (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of + [] => () + | dups => error ("Duplicate datatypes: " ^ commas dups)); + + fun prep_dt_spec (tvs, tname, mx, constrs) tname' (dts', constr_syntax, sorts, i) = + let + fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') = + let + val (cargs', sorts'') = fold_map (prep_typ tmp_thy) cargs sorts'; + val _ = + (case subtract (op =) tvs (fold (curry OldTerm.add_typ_tfree_names) cargs' []) of + [] => () + | vs => error ("Extra type variables on rhs: " ^ commas vs)) + in (constrs @ [(Sign.full_name_path tmp_thy tname' + (Binding.map_name (Syntax.const_name mx') cname), + map (dtyp_of_typ new_dts) cargs')], + constr_syntax' @ [(cname, mx')], sorts'') + end handle ERROR msg => cat_error msg + ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^ + " of datatype " ^ quote (Binding.str_of tname)); + + val (constrs', constr_syntax', sorts') = + fold prep_constr constrs ([], [], sorts) + + in + case duplicates (op =) (map fst constrs') of + [] => + (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname), + map DtTFree tvs, constrs'))], + constr_syntax @ [constr_syntax'], sorts', i + 1) + | dups => error ("Duplicate constructors " ^ commas dups ^ + " in datatype " ^ quote (Binding.str_of tname)) + end; + + val (dts', constr_syntax, sorts', i) = + fold2 prep_dt_spec dts new_type_names ([], [], [], 0); + val sorts = sorts' @ map (rpair (Sign.defaultS tmp_thy)) (subtract (op =) (map fst sorts') tyvars); + val dt_info = Datatype_Data.get_all thy; + val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i; + val _ = check_nonempty descr handle (exn as Datatype_Empty s) => + if #strict config then error ("Nonemptiness check failed for datatype " ^ s) + else raise exn; + + val _ = message config ("Constructing datatype(s) " ^ commas_quote new_type_names); + + in + thy + |> representation_proofs config dt_info new_type_names descr sorts + types_syntax constr_syntax (Datatype_Data.mk_case_names_induct (flat descr)) + |-> (fn (inject, distinct, induct) => Datatype_Data.derive_datatype_props + config dt_names (SOME new_type_names) descr sorts + induct inject distinct) + end; + +val add_datatype = gen_add_datatype Datatype_Data.cert_typ; +val datatype_cmd = snd ooo gen_add_datatype Datatype_Data.read_typ default_config; + +local + +structure P = OuterParse and K = OuterKeyword + +fun prep_datatype_decls args = + let + val names = map + (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args; + val specs = map (fn ((((_, vs), t), mx), cons) => + (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; + in (names, specs) end; + +val parse_datatype_decl = + (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix -- + (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix))); + +val parse_datatype_decls = P.and_list1 parse_datatype_decl >> prep_datatype_decls; + +in + +val _ = + OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl + (parse_datatype_decls >> (fn (names, specs) => Toplevel.theory (datatype_cmd names specs))); end; + +end; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_abs_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,15 +1,9 @@ -(* Title: HOL/Tools/datatype_abs_proofs.ML +(* Title: HOL/Tools/Datatype/datatype_abs_proofs.ML Author: Stefan Berghofer, TU Muenchen -Proofs and defintions independent of concrete representation -of datatypes (i.e. requiring only abstract properties such as -injectivity / distinctness of constructors and induction) - - - case distinction (exhaustion) theorems - - characteristic equations for primrec combinators - - characteristic equations for case combinators - - equations for splitting "P (case ...)" expressions - - "nchotomy" and "case_cong" theorems for TFL +Datatype package: proofs and defintions independent of concrete +representation of datatypes (i.e. requiring only abstract +properties: injectivity / distinctness of constructors and induction). *) signature DATATYPE_ABS_PROOFS = @@ -38,10 +32,10 @@ thm list -> thm list list -> theory -> thm list * theory end; -structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS = +structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS = struct -open DatatypeAux; +open Datatype_Aux; (************************ case distinction theorems ***************************) @@ -78,7 +72,7 @@ end; val casedist_thms = map_index prove_casedist_thm - (newTs ~~ DatatypeProp.make_casedists descr sorts) + (newTs ~~ Datatype_Prop.make_casedists descr sorts) in thy |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms @@ -109,7 +103,7 @@ (1 upto (length descr')); val rec_set_names = map (Sign.full_bname thy0) rec_set_names'; - val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used; + val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr sorts used; val rec_set_Ts = map (fn (T1, T2) => reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); @@ -260,7 +254,7 @@ resolve_tac rec_unique_thms 1, resolve_tac rec_intrs 1, REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)])) - (DatatypeProp.make_primrecs new_type_names descr sorts thy2) + (Datatype_Prop.make_primrecs new_type_names descr sorts thy2) in thy2 @@ -334,7 +328,7 @@ val case_thms = map (map (fn t => Skip_Proof.prove_global thy2 [] [] t (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))) - (DatatypeProp.make_cases new_type_names descr sorts thy2) + (Datatype_Prop.make_cases new_type_names descr sorts thy2) in thy2 |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms) @@ -370,7 +364,7 @@ end; val split_thm_pairs = map prove_split_thms - ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~ + ((Datatype_Prop.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~ dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs); val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs @@ -388,7 +382,7 @@ Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1]) - val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs + val weak_case_congs = map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs new_type_names descr sorts thy) in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end; @@ -413,7 +407,7 @@ end; val nchotomys = - map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms) + map prove_nchotomy (Datatype_Prop.make_nchotomys descr sorts ~~ casedist_thms) in thy |> store_thms "nchotomy" new_type_names nchotomys end; @@ -438,7 +432,7 @@ end) end; - val case_congs = map prove_case_cong (DatatypeProp.make_case_congs + val case_congs = map prove_case_cong (Datatype_Prop.make_case_congs new_type_names descr sorts thy ~~ nchotomys ~~ case_thms) in thy |> store_thms "case_cong" new_type_names case_congs end; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_aux.ML --- a/src/HOL/Tools/Datatype/datatype_aux.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_aux.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,7 +1,7 @@ -(* Title: HOL/Tools/datatype_aux.ML +(* Title: HOL/Tools/Datatype/datatype_aux.ML Author: Stefan Berghofer, TU Muenchen -Auxiliary functions for defining datatypes. +Datatype package: auxiliary data structures and functions. *) signature DATATYPE_COMMON = @@ -79,9 +79,10 @@ val unfold_datatypes : theory -> descr -> (string * sort) list -> info Symtab.table -> descr -> int -> descr list * int + val find_shortest_path : descr -> int -> (string * int) option end; -structure DatatypeAux : DATATYPE_AUX = +structure Datatype_Aux : DATATYPE_AUX = struct (* datatype option flags *) @@ -365,4 +366,24 @@ in (descr' :: descrs, i') end; +(* find shortest path to constructor with no recursive arguments *) + +fun find_nonempty descr is i = + let + val (_, _, constrs) = the (AList.lookup (op =) descr i); + fun arg_nonempty (_, DtRec i) = if member (op =) is i + then NONE + else Option.map (Integer.add 1 o snd) (find_nonempty descr (i::is) i) + | arg_nonempty _ = SOME 0; + fun max xs = Library.foldl + (fn (NONE, _) => NONE + | (SOME i, SOME j) => SOME (Int.max (i, j)) + | (_, NONE) => NONE) (SOME 0, xs); + val xs = sort (int_ord o pairself snd) + (map_filter (fn (s, dts) => Option.map (pair s) + (max (map (arg_nonempty o strip_dtyp) dts))) constrs) + in case xs of [] => NONE | x :: _ => SOME x end; + +fun find_shortest_path descr i = find_nonempty descr [i] i; + end; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_case.ML --- a/src/HOL/Tools/Datatype/datatype_case.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_case.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,30 +1,32 @@ -(* Title: HOL/Tools/datatype_case.ML +(* Title: HOL/Tools/Datatype/datatype_case.ML Author: Konrad Slind, Cambridge University Computer Laboratory Author: Stefan Berghofer, TU Muenchen -Nested case expressions on datatypes. +Datatype package: nested case expressions on datatypes. *) signature DATATYPE_CASE = sig - datatype config = Error | Warning | Quiet; - val make_case: (string * typ -> DatatypeAux.info option) -> + datatype config = Error | Warning | Quiet + type info = Datatype_Aux.info + val make_case: (string * typ -> info option) -> Proof.context -> config -> string list -> term -> (term * term) list -> term * (term * (int * bool)) list - val dest_case: (string -> DatatypeAux.info option) -> bool -> + val dest_case: (string -> info option) -> bool -> string list -> term -> (term * (term * term) list) option - val strip_case: (string -> DatatypeAux.info option) -> bool -> + val strip_case: (string -> info option) -> bool -> term -> (term * (term * term) list) option - val case_tr: bool -> (theory -> string * typ -> DatatypeAux.info option) - -> Proof.context -> term list -> term - val case_tr': (theory -> string -> DatatypeAux.info option) -> + val case_tr: bool -> (theory -> string * typ -> info option) -> + Proof.context -> term list -> term + val case_tr': (theory -> string -> info option) -> string -> Proof.context -> term list -> term end; -structure DatatypeCase : DATATYPE_CASE = +structure Datatype_Case : DATATYPE_CASE = struct datatype config = Error | Warning | Quiet; +type info = Datatype_Aux.info; exception CASE_ERROR of string * int; @@ -36,10 +38,10 @@ fun ty_info tab sT = case tab sT of - SOME ({descr, case_name, index, sorts, ...} : DatatypeAux.info) => + SOME ({descr, case_name, index, sorts, ...} : info) => let val (_, (tname, dts, constrs)) = nth descr index; - val mk_ty = DatatypeAux.typ_of_dtyp descr sorts; + val mk_ty = Datatype_Aux.typ_of_dtyp descr sorts; val T = Type (tname, map mk_ty dts) in SOME {case_name = case_name, @@ -84,7 +86,7 @@ val (_, Ty) = dest_Const c val Ts = binder_types Ty; val names = Name.variant_list used - (DatatypeProp.make_tnames (map Logic.unvarifyT Ts)); + (Datatype_Prop.make_tnames (map Logic.unvarifyT Ts)); val ty = body_type Ty; val ty_theta = ty_match ty colty handle Type.TYPE_MATCH => raise CASE_ERROR ("type mismatch", ~1) diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_codegen.ML --- a/src/HOL/Tools/Datatype/datatype_codegen.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,4 +1,4 @@ -(* Title: HOL/Tools/datatype_codegen.ML +(* Title: HOL/Tools/Datatype/datatype_codegen.ML Author: Stefan Berghofer and Florian Haftmann, TU Muenchen Code generator facilities for inductive datatypes. @@ -6,48 +6,23 @@ signature DATATYPE_CODEGEN = sig - include DATATYPE_COMMON - val find_shortest_path: descr -> int -> (string * int) option val setup: theory -> theory end; -structure DatatypeCodegen : DATATYPE_CODEGEN = +structure Datatype_Codegen : DATATYPE_CODEGEN = struct -open DatatypeAux - -(** find shortest path to constructor with no recursive arguments **) - -fun find_nonempty (descr: descr) is i = - let - val (_, _, constrs) = the (AList.lookup (op =) descr i); - fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i - then NONE - else Option.map (Integer.add 1 o snd) (find_nonempty descr (i::is) i) - | arg_nonempty _ = SOME 0; - fun max xs = Library.foldl - (fn (NONE, _) => NONE - | (SOME i, SOME j) => SOME (Int.max (i, j)) - | (_, NONE) => NONE) (SOME 0, xs); - val xs = sort (int_ord o pairself snd) - (map_filter (fn (s, dts) => Option.map (pair s) - (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs) - in case xs of [] => NONE | x :: _ => SOME x end; - -fun find_shortest_path descr i = find_nonempty descr [i] i; - - (** SML code generator **) open Codegen; (* datatype definition *) -fun add_dt_defs thy defs dep module (descr: descr) sorts gr = +fun add_dt_defs thy defs dep module descr sorts gr = let - val descr' = filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr; + val descr' = filter (can (map Datatype_Aux.dest_DtTFree o #2 o snd)) descr; val rtnames = map (#1 o snd) (filter (fn (_, (_, _, cs)) => - exists (exists DatatypeAux.is_rec_type o snd) cs) descr'); + exists (exists Datatype_Aux.is_rec_type o snd) cs) descr'); val (_, (tname, _, _)) :: _ = descr'; val node_id = tname ^ " (type)"; @@ -56,8 +31,8 @@ fun mk_dtdef prfx [] gr = ([], gr) | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr = let - val tvs = map DatatypeAux.dest_DtTFree dts; - val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs; + val tvs = map Datatype_Aux.dest_DtTFree dts; + val cs' = map (apsnd (map (Datatype_Aux.typ_of_dtyp descr sorts))) cs; val ((_, type_id), gr') = mk_type_id module' tname gr; val (ps, gr'') = gr' |> fold_map (fn (cname, cargs) => @@ -87,8 +62,8 @@ fun mk_term_of_def gr prfx [] = [] | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) = let - val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs; - val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts; + val cs' = map (apsnd (map (Datatype_Aux.typ_of_dtyp descr sorts))) cs; + val dts' = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; val T = Type (tname, dts'); val rest = mk_term_of_def gr "and " xs; val (eqs, _) = fold_map (fn (cname, Ts) => fn prfx => @@ -110,12 +85,12 @@ fun mk_gen_of_def gr prfx [] = [] | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) = let - val tvs = map DatatypeAux.dest_DtTFree dts; - val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts; + val tvs = map Datatype_Aux.dest_DtTFree dts; + val Us = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; val T = Type (tname, Us); val (cs1, cs2) = - List.partition (exists DatatypeAux.is_rec_type o snd) cs; - val SOME (cname, _) = find_shortest_path descr i; + List.partition (exists Datatype_Aux.is_rec_type o snd) cs; + val SOME (cname, _) = Datatype_Aux.find_shortest_path descr i; fun mk_delay p = Pretty.block [str "fn () =>", Pretty.brk 1, p]; @@ -125,14 +100,14 @@ fun mk_constr s b (cname, dts) = let val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s - (DatatypeAux.typ_of_dtyp descr sorts dt)) - [str (if b andalso DatatypeAux.is_rec_type dt then "0" + (Datatype_Aux.typ_of_dtyp descr sorts dt)) + [str (if b andalso Datatype_Aux.is_rec_type dt then "0" else "j")]) dts; - val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts; + val Ts = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; val xs = map str - (DatatypeProp.indexify_names (replicate (length dts) "x")); + (Datatype_Prop.indexify_names (replicate (length dts) "x")); val ts = map str - (DatatypeProp.indexify_names (replicate (length dts) "t")); + (Datatype_Prop.indexify_names (replicate (length dts) "t")); val (_, id) = get_const_id gr cname in mk_let @@ -378,10 +353,10 @@ | _ => NONE) cos; fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) = trueprop $ (equiv $ mk_eq (t1, t2) $ rhs); - val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index); + val injects = map prep_inject (nth (Datatype_Prop.make_injs [descr] vs) index); fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) = [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)]; - val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index)); + val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index)); val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty))); val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms))); @@ -436,7 +411,7 @@ in if null css then thy else thy - |> tap (fn _ => DatatypeAux.message config "Registering datatype for code generator ...") + |> tap (fn _ => Datatype_Aux.message config "Registering datatype for code generator ...") |> fold Code.add_datatype css |> fold_rev Code.add_default_eqn case_rewrites |> fold Code.add_case certs diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_data.ML --- a/src/HOL/Tools/Datatype/datatype_data.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_data.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,7 +1,7 @@ -(* Title: HOL/Tools/datatype.ML +(* Title: HOL/Tools/Datatype/datatype_data.ML Author: Stefan Berghofer, TU Muenchen -Datatype package for Isabelle/HOL. +Datatype package: bookkeeping; interpretation of existing types as datatypes. *) signature DATATYPE_DATA = @@ -25,7 +25,7 @@ val info_of_constr : theory -> string * typ -> info option val info_of_case : theory -> string -> info option val interpretation : (config -> string list -> theory -> theory) -> theory -> theory - val make_case : Proof.context -> DatatypeCase.config -> string list -> term -> + val make_case : Proof.context -> Datatype_Case.config -> string list -> term -> (term * term) list -> term * (term * (int * bool)) list val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list @@ -37,7 +37,7 @@ structure Datatype_Data: DATATYPE_DATA = struct -open DatatypeAux; +open Datatype_Aux; (** theory data **) @@ -104,9 +104,9 @@ val info as { descr, index, sorts = raw_sorts, ... } = the_info thy dtco; val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index; val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v)) - o DatatypeAux.dest_DtTFree) dtys; + o Datatype_Aux.dest_DtTFree) dtys; val cos = map - (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos; + (fn (co, tys) => (co, map (Datatype_Aux.typ_of_dtyp descr sorts) tys)) raw_cos; in (sorts, cos) end; fun the_descr thy (raw_tycos as raw_tyco :: _) = @@ -197,7 +197,7 @@ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end; fun induct_cases descr = - DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr)); + Datatype_Prop.indexify_names (maps (dt_cases descr) (map #2 descr)); fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i)); @@ -214,22 +214,22 @@ (* translation rules for case *) -fun make_case ctxt = DatatypeCase.make_case +fun make_case ctxt = Datatype_Case.make_case (info_of_constr (ProofContext.theory_of ctxt)) ctxt; -fun strip_case ctxt = DatatypeCase.strip_case +fun strip_case ctxt = Datatype_Case.strip_case (info_of_case (ProofContext.theory_of ctxt)); fun add_case_tr' case_names thy = Sign.add_advanced_trfuns ([], [], map (fn case_name => let val case_name' = Sign.const_syntax_name thy case_name - in (case_name', DatatypeCase.case_tr' info_of_case case_name') + in (case_name', Datatype_Case.case_tr' info_of_case case_name') end) case_names, []) thy; val trfun_setup = Sign.add_advanced_trfuns ([], - [("_case_syntax", DatatypeCase.case_tr true info_of_constr)], + [("_case_syntax", Datatype_Case.case_tr true info_of_constr)], [], []); @@ -299,21 +299,21 @@ val new_type_names = map Long_Name.base_name (the_default dt_names alt_names); val _ = message config ("Deriving properties for datatype(s) " ^ commas_quote new_type_names); - val (exhaust, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names + val (exhaust, thy3) = Datatype_Abs_Proofs.prove_casedist_thms config new_type_names descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2; - val (nchotomys, thy4) = DatatypeAbsProofs.prove_nchotomys config new_type_names + val (nchotomys, thy4) = Datatype_Abs_Proofs.prove_nchotomys config new_type_names descr sorts exhaust thy3; - val ((rec_names, rec_rewrites), thy5) = DatatypeAbsProofs.prove_primrec_thms + val ((rec_names, rec_rewrites), thy5) = Datatype_Abs_Proofs.prove_primrec_thms config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4)) inject (distinct, all_distincts thy2 (get_rec_types flat_descr sorts)) induct thy4; - val ((case_rewrites, case_names), thy6) = DatatypeAbsProofs.prove_case_thms + val ((case_rewrites, case_names), thy6) = Datatype_Abs_Proofs.prove_case_thms config new_type_names descr sorts rec_names rec_rewrites thy5; - val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names + val (case_congs, thy7) = Datatype_Abs_Proofs.prove_case_congs new_type_names descr sorts nchotomys case_rewrites thy6; - val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names + val (weak_case_congs, thy8) = Datatype_Abs_Proofs.prove_weak_case_congs new_type_names descr sorts thy7; - val (splits, thy9) = DatatypeAbsProofs.prove_split_thms + val (splits, thy9) = Datatype_Abs_Proofs.prove_split_thms config new_type_names descr sorts inject distinct exhaust case_rewrites thy8; val inducts = Project_Rule.projections (ProofContext.init thy2) induct; @@ -416,9 +416,9 @@ map (DtTFree o fst) vs, (map o apsnd) dtyps_of_typ constr)) val descr = map_index mk_spec cs; - val injs = DatatypeProp.make_injs [descr] vs; - val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs); - val ind = DatatypeProp.make_ind [descr] vs; + val injs = Datatype_Prop.make_injs [descr] vs; + val half_distincts = map snd (Datatype_Prop.make_distincts [descr] vs); + val ind = Datatype_Prop.make_ind [descr] vs; val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts]; fun after_qed' raw_thms = diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_prop.ML --- a/src/HOL/Tools/Datatype/datatype_prop.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_prop.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,38 +1,39 @@ -(* Title: HOL/Tools/datatype_prop.ML +(* Title: HOL/Tools/Datatype/datatype_prop.ML Author: Stefan Berghofer, TU Muenchen -Characteristic properties of datatypes. +Datatype package: characteristic properties of datatypes. *) signature DATATYPE_PROP = sig + include DATATYPE_COMMON val indexify_names: string list -> string list val make_tnames: typ list -> string list - val make_injs : DatatypeAux.descr list -> (string * sort) list -> term list list - val make_distincts : DatatypeAux.descr list -> + val make_injs : descr list -> (string * sort) list -> term list list + val make_distincts : descr list -> (string * sort) list -> (int * term list) list (*no symmetric inequalities*) - val make_ind : DatatypeAux.descr list -> (string * sort) list -> term - val make_casedists : DatatypeAux.descr list -> (string * sort) list -> term list - val make_primrec_Ts : DatatypeAux.descr list -> (string * sort) list -> + val make_ind : descr list -> (string * sort) list -> term + val make_casedists : descr list -> (string * sort) list -> term list + val make_primrec_Ts : descr list -> (string * sort) list -> string list -> typ list * typ list - val make_primrecs : string list -> DatatypeAux.descr list -> + val make_primrecs : string list -> descr list -> (string * sort) list -> theory -> term list - val make_cases : string list -> DatatypeAux.descr list -> + val make_cases : string list -> descr list -> (string * sort) list -> theory -> term list list - val make_splits : string list -> DatatypeAux.descr list -> + val make_splits : string list -> descr list -> (string * sort) list -> theory -> (term * term) list - val make_weak_case_congs : string list -> DatatypeAux.descr list -> + val make_weak_case_congs : string list -> descr list -> (string * sort) list -> theory -> term list - val make_case_congs : string list -> DatatypeAux.descr list -> + val make_case_congs : string list -> descr list -> (string * sort) list -> theory -> term list - val make_nchotomys : DatatypeAux.descr list -> + val make_nchotomys : descr list -> (string * sort) list -> term list end; -structure DatatypeProp : DATATYPE_PROP = +structure Datatype_Prop : DATATYPE_PROP = struct -open DatatypeAux; +open Datatype_Aux; fun indexify_names names = let diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_realizer.ML --- a/src/HOL/Tools/Datatype/datatype_realizer.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1,8 +1,8 @@ -(* Title: HOL/Tools/datatype_realizer.ML +(* Title: HOL/Tools/Datatype/datatype_realizer.ML Author: Stefan Berghofer, TU Muenchen -Porgram extraction from proofs involving datatypes: -Realizers for induction and case analysis +Program extraction from proofs involving datatypes: +realizers for induction and case analysis. *) signature DATATYPE_REALIZER = @@ -11,10 +11,10 @@ val setup: theory -> theory end; -structure DatatypeRealizer : DATATYPE_REALIZER = +structure Datatype_Realizer : DATATYPE_REALIZER = struct -open DatatypeAux; +open Datatype_Aux; fun subsets i j = if i <= j then let val is = subsets (i+1) j @@ -60,7 +60,7 @@ (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j => let val Ts = map (typ_of_dtyp descr sorts) cargs; - val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); + val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts); val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts); val frees = tnames ~~ Ts; @@ -97,7 +97,7 @@ if (j: int) = i then HOLogic.mk_fst t else mk_proj j is (HOLogic.mk_snd t); - val tnames = DatatypeProp.make_tnames recTs; + val tnames = Datatype_Prop.make_tnames recTs; val fTs = map fastype_of rec_fns; val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0))) @@ -132,7 +132,7 @@ (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm) ||> Sign.restore_naming thy; - val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []); + val ivs = rev (Term.add_vars (Logic.varify (Datatype_Prop.make_ind [descr] sorts)) []); val rvs = rev (Thm.fold_terms Term.add_vars thm' []); val ivs1 = map Var (filter_out (fn (_, T) => tname_of (body_type T) mem ["set", "bool"]) ivs); @@ -169,7 +169,7 @@ fun make_casedist_prem T (cname, cargs) = let val Ts = map (typ_of_dtyp descr sorts) cargs; - val frees = Name.variant_list ["P", "y"] (DatatypeProp.make_tnames Ts) ~~ Ts; + val frees = Name.variant_list ["P", "y"] (Datatype_Prop.make_tnames Ts) ~~ Ts; val free_ts = map Free frees; val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT) in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Datatype/datatype_rep_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Mon Nov 30 11:42:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,747 +0,0 @@ -(* Title: HOL/Tools/datatype_rep_proofs.ML - Author: Stefan Berghofer, TU Muenchen - -Definitional introduction of datatypes with proof of characteristic theorems: - - - injectivity of constructors - - distinctness of constructors - - induction theorem -*) - -signature DATATYPE_REP_PROOFS = -sig - val add_datatype : DatatypeAux.config -> string list -> (string list * binding * mixfix * - (binding * typ list * mixfix) list) list -> theory -> string list * theory - val datatype_cmd : string list -> (string list * binding * mixfix * - (binding * string list * mixfix) list) list -> theory -> theory -end; - -structure DatatypeRepProofs : DATATYPE_REP_PROOFS = -struct - -(** auxiliary **) - -open DatatypeAux; - -val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); - -val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; - -fun exh_thm_of (dt_info : info Symtab.table) tname = - #exhaust (the (Symtab.lookup dt_info tname)); - -val node_name = @{type_name "Datatype.node"}; -val In0_name = @{const_name "Datatype.In0"}; -val In1_name = @{const_name "Datatype.In1"}; -val Scons_name = @{const_name "Datatype.Scons"}; -val Leaf_name = @{const_name "Datatype.Leaf"}; -val Numb_name = @{const_name "Datatype.Numb"}; -val Lim_name = @{const_name "Datatype.Lim"}; -val Suml_name = @{const_name "Sum_Type.Suml"}; -val Sumr_name = @{const_name "Sum_Type.Sumr"}; - -val In0_inject = @{thm In0_inject}; -val In1_inject = @{thm In1_inject}; -val Scons_inject = @{thm Scons_inject}; -val Leaf_inject = @{thm Leaf_inject}; -val In0_eq = @{thm In0_eq}; -val In1_eq = @{thm In1_eq}; -val In0_not_In1 = @{thm In0_not_In1}; -val In1_not_In0 = @{thm In1_not_In0}; -val Lim_inject = @{thm Lim_inject}; -val Inl_inject = @{thm Inl_inject}; -val Inr_inject = @{thm Inr_inject}; -val Suml_inject = @{thm Suml_inject}; -val Sumr_inject = @{thm Sumr_inject}; - - - -(** proof of characteristic theorems **) - -fun representation_proofs (config : config) (dt_info : info Symtab.table) - new_type_names descr sorts types_syntax constr_syntax case_names_induct thy = - let - val descr' = flat descr; - val big_name = space_implode "_" new_type_names; - val thy1 = Sign.add_path big_name thy; - val big_rec_name = big_name ^ "_rep_set"; - val rep_set_names' = - (if length descr' = 1 then [big_rec_name] else - (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) - (1 upto (length descr')))); - val rep_set_names = map (Sign.full_bname thy1) rep_set_names'; - - val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr); - val leafTs' = get_nonrec_types descr' sorts; - val branchTs = get_branching_types descr' sorts; - val branchT = if null branchTs then HOLogic.unitT - else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs; - val arities = remove (op =) 0 (get_arities descr'); - val unneeded_vars = - subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars); - val leafTs = leafTs' @ map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars; - val recTs = get_rec_types descr' sorts; - val (newTs, oldTs) = chop (length (hd descr)) recTs; - val sumT = if null leafTs then HOLogic.unitT - else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) leafTs; - val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT])); - val UnivT = HOLogic.mk_setT Univ_elT; - val UnivT' = Univ_elT --> HOLogic.boolT; - val Collect = Const (@{const_name Collect}, UnivT' --> UnivT); - - val In0 = Const (In0_name, Univ_elT --> Univ_elT); - val In1 = Const (In1_name, Univ_elT --> Univ_elT); - val Leaf = Const (Leaf_name, sumT --> Univ_elT); - val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT); - - (* make injections needed for embedding types in leaves *) - - fun mk_inj T' x = - let - fun mk_inj' T n i = - if n = 1 then x else - let val n2 = n div 2; - val Type (_, [T1, T2]) = T - in - if i <= n2 then - Const (@{const_name "Sum_Type.Inl"}, T1 --> T) $ (mk_inj' T1 n2 i) - else - Const (@{const_name "Sum_Type.Inr"}, T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) - end - in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) - end; - - (* make injections for constructors *) - - fun mk_univ_inj ts = Balanced_Tree.access - {left = fn t => In0 $ t, - right = fn t => In1 $ t, - init = - if ts = [] then Const (@{const_name undefined}, Univ_elT) - else foldr1 (HOLogic.mk_binop Scons_name) ts}; - - (* function spaces *) - - fun mk_fun_inj T' x = - let - fun mk_inj T n i = - if n = 1 then x else - let - val n2 = n div 2; - val Type (_, [T1, T2]) = T; - fun mkT U = (U --> Univ_elT) --> T --> Univ_elT - in - if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i - else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2) - end - in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) - end; - - fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t; - - (************** generate introduction rules for representing set **********) - - val _ = message config "Constructing representing sets ..."; - - (* make introduction rule for a single constructor *) - - fun make_intr s n (i, (_, cargs)) = - let - fun mk_prem dt (j, prems, ts) = - (case strip_dtyp dt of - (dts, DtRec k) => - let - val Ts = map (typ_of_dtyp descr' sorts) dts; - val free_t = - app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts) - in (j + 1, list_all (map (pair "x") Ts, - HOLogic.mk_Trueprop - (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems, - mk_lim free_t Ts :: ts) - end - | _ => - let val T = typ_of_dtyp descr' sorts dt - in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts) - end); - - val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []); - val concl = HOLogic.mk_Trueprop - (Free (s, UnivT') $ mk_univ_inj ts n i) - in Logic.list_implies (prems, concl) - end; - - val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) => - map (make_intr rep_set_name (length constrs)) - ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names'); - - val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) = - thy1 - |> Sign.map_naming Name_Space.conceal - |> Inductive.add_inductive_global - {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, - coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false} - (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') [] - (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] - ||> Sign.restore_naming thy1 - ||> Theory.checkpoint; - - (********************************* typedef ********************************) - - val (typedefs, thy3) = thy2 |> - Sign.parent_path |> - fold_map (fn ((((name, mx), tvs), c), name') => - Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx) - (Collect $ Const (c, UnivT')) NONE - (rtac exI 1 THEN rtac CollectI 1 THEN - QUIET_BREADTH_FIRST (has_fewer_prems 1) - (resolve_tac rep_intrs 1))) - (types_syntax ~~ tyvars ~~ - (take (length newTs) rep_set_names) ~~ new_type_names) ||> - Sign.add_path big_name; - - (*********************** definition of constructors ***********************) - - val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_"; - val rep_names = map (curry op ^ "Rep_") new_type_names; - val rep_names' = map (fn i => big_rep_name ^ (string_of_int i)) - (1 upto (length (flat (tl descr)))); - val all_rep_names = map (Sign.intern_const thy3) rep_names @ - map (Sign.full_bname thy3) rep_names'; - - (* isomorphism declarations *) - - val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn)) - (oldTs ~~ rep_names'); - - (* constructor definitions *) - - fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) = - let - fun constr_arg dt (j, l_args, r_args) = - let val T = typ_of_dtyp descr' sorts dt; - val free_t = mk_Free "x" T j - in (case (strip_dtyp dt, strip_type T) of - ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim - (Const (nth all_rep_names m, U --> Univ_elT) $ - app_bnds free_t (length Us)) Us :: r_args) - | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args)) - end; - - val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); - val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T; - val abs_name = Sign.intern_const thy ("Abs_" ^ tname); - val rep_name = Sign.intern_const thy ("Rep_" ^ tname); - val lhs = list_comb (Const (cname, constrT), l_args); - val rhs = mk_univ_inj r_args n i; - val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs); - val def_name = Long_Name.base_name cname ^ "_def"; - val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq - (Const (rep_name, T --> Univ_elT) $ lhs, rhs)); - val ([def_thm], thy') = - thy - |> Sign.add_consts_i [(cname', constrT, mx)] - |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]; - - in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end; - - (* constructor definitions for datatype *) - - fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax) - (thy, defs, eqns, rep_congs, dist_lemmas) = - let - val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong; - val rep_const = cterm_of thy - (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT)); - val cong' = - Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong); - val dist = - Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); - val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs)) - (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1); - in - (Sign.parent_path thy', defs', eqns @ [eqns'], - rep_congs @ [cong'], dist_lemmas @ [dist]) - end; - - val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = - fold dt_constr_defs - (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax) - (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []); - - - (*********** isomorphisms for new types (introduced by typedef) ***********) - - val _ = message config "Proving isomorphism properties ..."; - - val newT_iso_axms = map (fn (_, td) => - (collect_simp (#Abs_inverse td), #Rep_inverse td, - collect_simp (#Rep td))) typedefs; - - val newT_iso_inj_thms = map (fn (_, td) => - (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs; - - (********* isomorphisms between existing types and "unfolded" types *******) - - (*---------------------------------------------------------------------*) - (* isomorphisms are defined using primrec-combinators: *) - (* generate appropriate functions for instantiating primrec-combinator *) - (* *) - (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *) - (* *) - (* also generate characteristic equations for isomorphisms *) - (* *) - (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *) - (*---------------------------------------------------------------------*) - - fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) = - let - val argTs = map (typ_of_dtyp descr' sorts) cargs; - val T = nth recTs k; - val rep_name = nth all_rep_names k; - val rep_const = Const (rep_name, T --> Univ_elT); - val constr = Const (cname, argTs ---> T); - - fun process_arg ks' dt (i2, i2', ts, Ts) = - let - val T' = typ_of_dtyp descr' sorts dt; - val (Us, U) = strip_type T' - in (case strip_dtyp dt of - (_, DtRec j) => if j mem ks' then - (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds - (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us], - Ts @ [Us ---> Univ_elT]) - else - (i2 + 1, i2', ts @ [mk_lim - (Const (nth all_rep_names j, U --> Univ_elT) $ - app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts) - | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts)) - end; - - val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []); - val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1))); - val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1))); - val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i); - - val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []); - val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq - (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i)) - - in (fs @ [f], eqns @ [eqn], i + 1) end; - - (* define isomorphisms for all mutually recursive datatypes in list ds *) - - fun make_iso_defs ds (thy, char_thms) = - let - val ks = map fst ds; - val (_, (tname, _, _)) = hd ds; - val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname); - - fun process_dt (k, (tname, _, constrs)) (fs, eqns, isos) = - let - val (fs', eqns', _) = - fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1); - val iso = (nth recTs k, nth all_rep_names k) - in (fs', eqns', isos @ [iso]) end; - - val (fs, eqns, isos) = fold process_dt ds ([], [], []); - val fTs = map fastype_of fs; - val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"), - Logic.mk_equals (Const (iso_name, T --> Univ_elT), - list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos); - val (def_thms, thy') = - apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy); - - (* prove characteristic equations *) - - val rewrites = def_thms @ (map mk_meta_eq rec_rewrites); - val char_thms' = map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn - (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; - - in (thy', char_thms' @ char_thms) end; - - val (thy5, iso_char_thms) = apfst Theory.checkpoint (fold_rev make_iso_defs - (tl descr) (Sign.add_path big_name thy4, [])); - - (* prove isomorphism properties *) - - fun mk_funs_inv thy thm = - let - val prop = Thm.prop_of thm; - val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $ - (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop; - val used = OldTerm.add_term_tfree_names (a, []); - - fun mk_thm i = - let - val Ts = map (TFree o rpair HOLogic.typeS) - (Name.variant_list used (replicate i "'t")); - val f = Free ("f", Ts ---> U) - in Skip_Proof.prove_global thy [] [] (Logic.mk_implies - (HOLogic.mk_Trueprop (HOLogic.list_all - (map (pair "x") Ts, S $ app_bnds f i)), - HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts, - r $ (a $ app_bnds f i)), f)))) - (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1), - REPEAT (etac allE 1), rtac thm 1, atac 1]) - end - in map (fn r => r RS subst) (thm :: map mk_thm arities) end; - - (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *) - - val fun_congs = map (fn T => make_elim (Drule.instantiate' - [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs; - - fun prove_iso_thms ds (inj_thms, elem_thms) = - let - val (_, (tname, _, _)) = hd ds; - val induct = (#induct o the o Symtab.lookup dt_info) tname; - - fun mk_ind_concl (i, _) = - let - val T = nth recTs i; - val Rep_t = Const (nth all_rep_names i, T --> Univ_elT); - val rep_set_name = nth rep_set_names i - in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $ - HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $ - HOLogic.mk_eq (mk_Free "x" T i, Bound 0)), - Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i)) - end; - - val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds); - - val rewrites = map mk_meta_eq iso_char_thms; - val inj_thms' = map snd newT_iso_inj_thms @ - map (fn r => r RS @{thm injD}) inj_thms; - - val inj_thm = Skip_Proof.prove_global thy5 [] [] - (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY - [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - REPEAT (EVERY - [rtac allI 1, rtac impI 1, - exh_tac (exh_thm_of dt_info) 1, - REPEAT (EVERY - [hyp_subst_tac 1, - rewrite_goals_tac rewrites, - REPEAT (dresolve_tac [In0_inject, In1_inject] 1), - (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1) - ORELSE (EVERY - [REPEAT (eresolve_tac (Scons_inject :: - map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1), - REPEAT (cong_tac 1), rtac refl 1, - REPEAT (atac 1 ORELSE (EVERY - [REPEAT (rtac ext 1), - REPEAT (eresolve_tac (mp :: allE :: - map make_elim (Suml_inject :: Sumr_inject :: - Lim_inject :: inj_thms') @ fun_congs) 1), - atac 1]))])])])]); - - val inj_thms'' = map (fn r => r RS @{thm datatype_injI}) - (split_conj_thm inj_thm); - - val elem_thm = - Skip_Proof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2)) - (fn _ => - EVERY [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - rewrite_goals_tac rewrites, - REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW - ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]); - - in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm)) - end; - - val (iso_inj_thms_unfolded, iso_elem_thms) = - fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms); - val iso_inj_thms = map snd newT_iso_inj_thms @ - map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded; - - (* prove dt_rep_set_i x --> x : range dt_Rep_i *) - - fun mk_iso_t (((set_name, iso_name), i), T) = - let val isoT = T --> Univ_elT - in HOLogic.imp $ - (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $ - (if i < length newTs then HOLogic.true_const - else HOLogic.mk_mem (mk_Free "x" Univ_elT i, - Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $ - Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T))) - end; - - val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t - (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs))); - - (* all the theorems are proved by one single simultaneous induction *) - - val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq})) - iso_inj_thms_unfolded; - - val iso_thms = if length descr = 1 then [] else - drop (length newTs) (split_conj_thm - (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY - [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - REPEAT (rtac TrueI 1), - rewrite_goals_tac (mk_meta_eq choice_eq :: - symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs), - rewrite_goals_tac (map symmetric range_eqs), - REPEAT (EVERY - [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @ - maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1), - TRY (hyp_subst_tac 1), - rtac (sym RS range_eqI) 1, - resolve_tac iso_char_thms 1])]))); - - val Abs_inverse_thms' = - map #1 newT_iso_axms @ - map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp]) - iso_inj_thms_unfolded iso_thms; - - val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms'; - - (******************* freeness theorems for constructors *******************) - - val _ = message config "Proving freeness of constructors ..."; - - (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *) - - fun prove_constr_rep_thm eqn = - let - val inj_thms = map fst newT_iso_inj_thms; - val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms) - in Skip_Proof.prove_global thy5 [] [] eqn (fn _ => EVERY - [resolve_tac inj_thms 1, - rewrite_goals_tac rewrites, - rtac refl 3, - resolve_tac rep_intrs 2, - REPEAT (resolve_tac iso_elem_thms 1)]) - end; - - (*--------------------------------------------------------------*) - (* constr_rep_thms and rep_congs are used to prove distinctness *) - (* of constructors. *) - (*--------------------------------------------------------------*) - - val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns; - - val dist_rewrites = map (fn (rep_thms, dist_lemma) => - dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) - (constr_rep_thms ~~ dist_lemmas); - - fun prove_distinct_thms dist_rewrites' (k, ts) = - let - fun prove [] = [] - | prove (t :: ts) = - let - val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ => - EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1]) - in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end; - in prove ts end; - - val distinct_thms = map2 (prove_distinct_thms) - dist_rewrites (DatatypeProp.make_distincts descr sorts); - - (* prove injectivity of constructors *) - - fun prove_constr_inj_thm rep_thms t = - let val inj_thms = Scons_inject :: (map make_elim - (iso_inj_thms @ - [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject, - Lim_inject, Suml_inject, Sumr_inject])) - in Skip_Proof.prove_global thy5 [] [] t (fn _ => EVERY - [rtac iffI 1, - REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2, - dresolve_tac rep_congs 1, dtac box_equals 1, - REPEAT (resolve_tac rep_thms 1), - REPEAT (eresolve_tac inj_thms 1), - REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1), - REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1), - atac 1]))]) - end; - - val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts) - ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms); - - val ((constr_inject', distinct_thms'), thy6) = - thy5 - |> Sign.parent_path - |> store_thmss "inject" new_type_names constr_inject - ||>> store_thmss "distinct" new_type_names distinct_thms; - - (*************************** induction theorem ****************************) - - val _ = message config "Proving induction rule for datatypes ..."; - - val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @ - (map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded); - val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded; - - fun mk_indrule_lemma ((i, _), T) (prems, concls) = - let - val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ - mk_Free "x" T i; - - val Abs_t = if i < length newTs then - Const (Sign.intern_const thy6 - ("Abs_" ^ (nth new_type_names i)), Univ_elT --> T) - else Const (@{const_name the_inv_into}, - [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $ - HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT) - - in (prems @ [HOLogic.imp $ - (Const (nth rep_set_names i, UnivT') $ Rep_t) $ - (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], - concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) - end; - - val (indrule_lemma_prems, indrule_lemma_concls) = - fold mk_indrule_lemma (descr' ~~ recTs) ([], []); - - val cert = cterm_of thy6; - - val indrule_lemma = Skip_Proof.prove_global thy6 [] [] - (Logic.mk_implies - (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), - HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY - [REPEAT (etac conjE 1), - REPEAT (EVERY - [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1, - etac mp 1, resolve_tac iso_elem_thms 1])]); - - val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); - val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else - map (Free o apfst fst o dest_Var) Ps; - val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma; - - val dt_induct_prop = DatatypeProp.make_ind descr sorts; - val dt_induct = Skip_Proof.prove_global thy6 [] - (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) - (fn {prems, ...} => EVERY - [rtac indrule_lemma' 1, - (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - EVERY (map (fn (prem, r) => (EVERY - [REPEAT (eresolve_tac Abs_inverse_thms 1), - simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1, - DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) - (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]); - - val ([dt_induct'], thy7) = - thy6 - |> Sign.add_path big_name - |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] - ||> Sign.parent_path - ||> Theory.checkpoint; - - in - ((constr_inject', distinct_thms', dt_induct'), thy7) - end; - - - -(** definitional introduction of datatypes **) - -fun gen_add_datatype prep_typ config new_type_names dts thy = - let - val _ = Theory.requires thy "Datatype" "datatype definitions"; - - (* this theory is used just for parsing *) - val tmp_thy = thy |> - Theory.copy |> - Sign.add_types (map (fn (tvs, tname, mx, _) => - (tname, length tvs, mx)) dts); - - val (tyvars, _, _, _)::_ = dts; - val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) => - let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname) - in - (case duplicates (op =) tvs of - [] => - if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx)) - else error ("Mutually recursive datatypes must have same type parameters") - | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^ - " : " ^ commas dups)) - end) dts); - val dt_names = map fst new_dts; - - val _ = - (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of - [] => () - | dups => error ("Duplicate datatypes: " ^ commas dups)); - - fun prep_dt_spec (tvs, tname, mx, constrs) tname' (dts', constr_syntax, sorts, i) = - let - fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') = - let - val (cargs', sorts'') = fold_map (prep_typ tmp_thy) cargs sorts'; - val _ = - (case subtract (op =) tvs (fold (curry OldTerm.add_typ_tfree_names) cargs' []) of - [] => () - | vs => error ("Extra type variables on rhs: " ^ commas vs)) - in (constrs @ [(Sign.full_name_path tmp_thy tname' - (Binding.map_name (Syntax.const_name mx') cname), - map (dtyp_of_typ new_dts) cargs')], - constr_syntax' @ [(cname, mx')], sorts'') - end handle ERROR msg => cat_error msg - ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^ - " of datatype " ^ quote (Binding.str_of tname)); - - val (constrs', constr_syntax', sorts') = - fold prep_constr constrs ([], [], sorts) - - in - case duplicates (op =) (map fst constrs') of - [] => - (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname), - map DtTFree tvs, constrs'))], - constr_syntax @ [constr_syntax'], sorts', i + 1) - | dups => error ("Duplicate constructors " ^ commas dups ^ - " in datatype " ^ quote (Binding.str_of tname)) - end; - - val (dts', constr_syntax, sorts', i) = - fold2 prep_dt_spec dts new_type_names ([], [], [], 0); - val sorts = sorts' @ map (rpair (Sign.defaultS tmp_thy)) (subtract (op =) (map fst sorts') tyvars); - val dt_info = Datatype_Data.get_all thy; - val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i; - val _ = check_nonempty descr handle (exn as Datatype_Empty s) => - if #strict config then error ("Nonemptiness check failed for datatype " ^ s) - else raise exn; - - val _ = message config ("Constructing datatype(s) " ^ commas_quote new_type_names); - - in - thy - |> representation_proofs config dt_info new_type_names descr sorts - types_syntax constr_syntax (Datatype_Data.mk_case_names_induct (flat descr)) - |-> (fn (inject, distinct, induct) => Datatype_Data.derive_datatype_props - config dt_names (SOME new_type_names) descr sorts - induct inject distinct) - end; - -val add_datatype = gen_add_datatype Datatype_Data.cert_typ; -val datatype_cmd = snd ooo gen_add_datatype Datatype_Data.read_typ default_config; - -local - -structure P = OuterParse and K = OuterKeyword - -fun prep_datatype_decls args = - let - val names = map - (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args; - val specs = map (fn ((((_, vs), t), mx), cons) => - (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; - in (names, specs) end; - -val parse_datatype_decl = - (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix -- - (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix))); - -val parse_datatype_decls = P.and_list1 parse_datatype_decl >> prep_datatype_decls; - -in - -val _ = - OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl - (parse_datatype_decls >> (fn (names, specs) => Toplevel.theory (datatype_cmd names specs))); - -end; - -end; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Function/size.ML --- a/src/HOL/Tools/Function/size.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Function/size.ML Mon Nov 30 11:42:49 2009 +0100 @@ -13,7 +13,7 @@ structure Size: SIZE = struct -open DatatypeAux; +open Datatype_Aux; structure SizeData = Theory_Data ( @@ -177,7 +177,7 @@ fun gen_mk_size_eq p size_of size_ofp size_const T (cname, cargs) = let val Ts = map (typ_of_dtyp descr sorts) cargs; - val tnames = Name.variant_list f_names (DatatypeProp.make_tnames Ts); + val tnames = Name.variant_list f_names (Datatype_Prop.make_tnames Ts); val ts = map_filter (fn (sT as (s, T), dt) => Option.map (fn sz => sz $ Free sT) (if p dt then size_of_type size_of extra_size size_ofp T diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Nov 30 11:42:49 2009 +0100 @@ -516,7 +516,7 @@ | NONE => NONE (* FIXME: use antiquotation for "code_numeral" below or detect "rep_datatype", - e.g., by adding a field to "DatatypeAux.info". *) + e.g., by adding a field to "Datatype_Aux.info". *) (* string -> bool *) fun is_basic_datatype s = s mem [@{type_name "*"}, @{type_name bool}, @{type_name unit}, diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Nov 30 11:42:49 2009 +0100 @@ -1282,7 +1282,7 @@ val v' = Free (name', T); in lambda v (fst (Datatype.make_case - (ProofContext.init thy) DatatypeCase.Quiet [] v + (ProofContext.init thy) Datatype_Case.Quiet [] v [(HOLogic.mk_tuple out_ts, if null eqs'' then success_t else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $ diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/inductive_realizer.ML Mon Nov 30 11:42:49 2009 +0100 @@ -243,7 +243,7 @@ thy |> f (map snd dts) |-> (fn dtinfo => pair (map fst dts, SOME dtinfo)) - handle DatatypeAux.Datatype_Empty name' => + handle Datatype_Aux.Datatype_Empty name' => let val name = Long_Name.base_name name'; val dname = Name.variant used "Dummy"; @@ -398,7 +398,7 @@ val (thm', thy') = PureThy.store_thm (Binding.qualified_name (space_implode "_" (Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy; val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp))) - (DatatypeAux.split_conj_thm thm'); + (Datatype_Aux.split_conj_thm thm'); val ([thms'], thy'') = PureThy.add_thmss [((Binding.qualified_name (space_implode "_" (Long_Name.qualify qualifier "inducts" :: vs' @ Ps @ diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/old_primrec.ML --- a/src/HOL/Tools/old_primrec.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/old_primrec.ML Mon Nov 30 11:42:49 2009 +0100 @@ -21,7 +21,7 @@ structure OldPrimrec : OLD_PRIMREC = struct -open DatatypeAux; +open Datatype_Aux; exception RecError of string; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/primrec.ML Mon Nov 30 11:42:49 2009 +0100 @@ -23,7 +23,7 @@ structure Primrec : PRIMREC = struct -open DatatypeAux; +open Datatype_Aux; exception PrimrecError of string * term option; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/quickcheck_generators.ML --- a/src/HOL/Tools/quickcheck_generators.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/quickcheck_generators.ML Mon Nov 30 11:42:49 2009 +0100 @@ -246,10 +246,10 @@ mk_const @{const_name random_fun_lift} [fTs ---> T, fT] $ mk_random_fun_lift fTs t; val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size'); - val size = Option.map snd (DatatypeCodegen.find_shortest_path descr k) + val size = Option.map snd (Datatype_Aux.find_shortest_path descr k) |> the_default 0; in (SOME size, (t, fTs ---> T)) end; - val tss = DatatypeAux.interpret_construction descr vs + val tss = Datatype_Aux.interpret_construction descr vs { atyp = mk_random_call, dtyp = mk_random_aux_call }; fun mk_consexpr simpleT (c, xs) = let @@ -287,9 +287,9 @@ fun mk_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = let - val _ = DatatypeAux.message config "Creating quickcheck generators ..."; + val _ = Datatype_Aux.message config "Creating quickcheck generators ..."; val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq; - fun mk_size_arg k = case DatatypeCodegen.find_shortest_path descr k + fun mk_size_arg k = case Datatype_Aux.find_shortest_path descr k of SOME (_, l) => if l = 0 then size else @{term "max :: code_numeral \ code_numeral \ code_numeral"} $ HOLogic.mk_number @{typ code_numeral} l $ size @@ -328,10 +328,10 @@ val typerep_vs = (map o apsnd) (curry (Sorts.inter_sort algebra) @{sort typerep}) raw_vs; val random_insts = (map (rpair @{sort random}) o flat o maps snd o maps snd) - (DatatypeAux.interpret_construction descr typerep_vs + (Datatype_Aux.interpret_construction descr typerep_vs { atyp = single, dtyp = (K o K o K) [] }); val term_of_insts = (map (rpair @{sort term_of}) o flat o maps snd o maps snd) - (DatatypeAux.interpret_construction descr typerep_vs + (Datatype_Aux.interpret_construction descr typerep_vs { atyp = K [], dtyp = K o K }); val has_inst = exists (fn tyco => can (Sorts.mg_domain algebra tyco) @{sort random}) tycos; diff -r e191b400a8e0 -r f94fb13ecbb3 src/HOL/Tools/refute.ML --- a/src/HOL/Tools/refute.ML Mon Nov 30 11:42:48 2009 +0100 +++ b/src/HOL/Tools/refute.ML Mon Nov 30 11:42:49 2009 +0100 @@ -406,12 +406,12 @@ (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *) (* ------------------------------------------------------------------------- *) - fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) = + fun typ_of_dtyp descr typ_assoc (Datatype_Aux.DtTFree a) = (* replace a 'DtTFree' variable by the associated type *) - the (AList.lookup (op =) typ_assoc (DatatypeAux.DtTFree a)) - | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) = + the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a)) + | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, ds)) = Type (s, map (typ_of_dtyp descr typ_assoc) ds) - | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) = + | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) = let val (s, ds, _) = the (AList.lookup (op =) descr i) in @@ -946,7 +946,7 @@ (* sanity check: every element in 'dtyps' must be a *) (* 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) typs then + case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then raise REFUTE ("ground_types", "datatype argument (for type " ^ Syntax.string_of_typ_global thy T ^ ") is not a variable") else () @@ -958,11 +958,11 @@ val dT = typ_of_dtyp descr typ_assoc d in case d of - DatatypeAux.DtTFree _ => + Datatype_Aux.DtTFree _ => collect_types dT acc - | DatatypeAux.DtType (_, ds) => + | Datatype_Aux.DtType (_, ds) => collect_types dT (fold_rev collect_dtyp ds acc) - | DatatypeAux.DtRec i => + | Datatype_Aux.DtRec i => if dT mem acc then acc (* prevent infinite recursion *) else @@ -971,7 +971,7 @@ (* if the current type is a recursive IDT (i.e. a depth *) (* is required), add it to 'acc' *) val acc_dT = if Library.exists (fn (_, ds) => - Library.exists DatatypeAux.is_rec_type ds) dconstrs then + Library.exists Datatype_Aux.is_rec_type ds) dconstrs then insert (op =) dT acc else acc (* collect argument types *) @@ -985,7 +985,7 @@ in (* argument types 'Ts' could be added here, but they are also *) (* added by 'collect_dtyp' automatically *) - collect_dtyp (DatatypeAux.DtRec index) acc + collect_dtyp (Datatype_Aux.DtRec index) acc end | NONE => (* not an inductive datatype, e.g. defined via "typedef" or *) @@ -1157,7 +1157,7 @@ in (* recursive datatype? *) Library.exists (fn (_, ds) => - Library.exists DatatypeAux.is_rec_type ds) constrs + Library.exists Datatype_Aux.is_rec_type ds) constrs end | NONE => false) | _ => false) types @@ -1952,7 +1952,7 @@ val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_interpreter", "datatype argument (for type " @@ -2076,7 +2076,7 @@ val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_constructor_interpreter", "datatype argument (for type " @@ -2135,7 +2135,7 @@ val typ_assoc = dtyps ~~ Ts' (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_constructor_interpreter", "datatype argument (for type " @@ -2350,7 +2350,7 @@ (* sanity check: every element in 'dtyps' must be a *) (* 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_recursion_interpreter", @@ -2372,10 +2372,10 @@ (case AList.lookup op= acc d of NONE => (case d of - DatatypeAux.DtTFree _ => + Datatype_Aux.DtTFree _ => (* add the association, proceed *) rec_typ_assoc ((d, T)::acc) xs - | DatatypeAux.DtType (s, ds) => + | Datatype_Aux.DtType (s, ds) => let val (s', Ts) = dest_Type T in @@ -2385,7 +2385,7 @@ raise REFUTE ("IDT_recursion_interpreter", "DtType/Type mismatch") end - | DatatypeAux.DtRec i => + | Datatype_Aux.DtRec i => let val (_, ds, _) = the (AList.lookup (op =) descr i) val (_, Ts) = dest_Type T @@ -2401,7 +2401,7 @@ raise REFUTE ("IDT_recursion_interpreter", "different type associations for the same dtyp")) val typ_assoc = filter - (fn (DatatypeAux.DtTFree _, _) => true | (_, _) => false) + (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false) (rec_typ_assoc [] (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT)) (* sanity check: typ_assoc must associate types to the *) @@ -2417,7 +2417,7 @@ val mc_intrs = map (fn (idx, (_, _, cs)) => let val c_return_typ = typ_of_dtyp descr typ_assoc - (DatatypeAux.DtRec idx) + (Datatype_Aux.DtRec idx) in (idx, map (fn (cname, cargs) => (#1 o interpret thy (typs, []) {maxvars=0, @@ -2471,7 +2471,7 @@ val _ = map (fn (idx, intrs) => let val T = typ_of_dtyp descr typ_assoc - (DatatypeAux.DtRec idx) + (Datatype_Aux.DtRec idx) in if length intrs <> size_of_type thy (typs, []) T then raise REFUTE ("IDT_recursion_interpreter", @@ -2552,7 +2552,7 @@ val (_, _, constrs) = the (AList.lookup (op =) descr idx) val (_, dtyps) = List.nth (constrs, c) val rec_dtyps_args = filter - (DatatypeAux.is_rec_type o fst) (dtyps ~~ args) + (Datatype_Aux.is_rec_type o fst) (dtyps ~~ args) (* map those indices to interpretations *) val rec_dtyps_intrs = map (fn (dtyp, arg) => let @@ -2565,18 +2565,18 @@ (* takes the dtyp and interpretation of an element, *) (* and computes the interpretation for the *) (* corresponding recursive argument *) - fun rec_intr (DatatypeAux.DtRec i) (Leaf xs) = + fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) = (* recursive argument is "rec_i params elem" *) compute_array_entry i (find_index (fn x => x = True) xs) - | rec_intr (DatatypeAux.DtRec _) (Node _) = + | rec_intr (Datatype_Aux.DtRec _) (Node _) = raise REFUTE ("IDT_recursion_interpreter", "interpretation for IDT is a node") - | rec_intr (DatatypeAux.DtType ("fun", [dt1, dt2])) + | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2])) (Node xs) = (* recursive argument is something like *) (* "\x::dt1. rec_? params (elem x)" *) Node (map (rec_intr dt2) xs) - | rec_intr (DatatypeAux.DtType ("fun", [_, _])) + | rec_intr (Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) = raise REFUTE ("IDT_recursion_interpreter", "interpretation for function dtyp is a leaf") @@ -3169,7 +3169,7 @@ val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_printer", "datatype argument (for type " ^ Syntax.string_of_typ_global thy (Type (s, Ts)) ^ ") is not a variable")