tuned signature;
authorwenzelm
Fri, 16 Dec 2011 10:38:38 +0100
changeset 45896 100fb1f33e3e
parent 45895 36f3f0010b7d
child 45897 65cef0298158
tuned signature;
src/HOL/Nominal/nominal_datatype.ML
src/HOL/Tools/ATP/atp_util.ML
src/HOL/Tools/Datatype/datatype_abs_proofs.ML
src/HOL/Tools/Datatype/datatype_prop.ML
src/HOL/Tools/Datatype/datatype_realizer.ML
src/HOL/Tools/Function/size.ML
src/HOL/Tools/Nitpick/nitpick_util.ML
src/HOL/Tools/refute.ML
--- a/src/HOL/Nominal/nominal_datatype.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -42,8 +42,8 @@
 (* theory data *)
 
 type descr =
-  (int * (string * Datatype_Aux.dtyp list *
-      (string * (Datatype_Aux.dtyp list * Datatype_Aux.dtyp) list) list)) list;
+  (int * (string * Datatype.dtyp list *
+      (string * (Datatype.dtyp list * Datatype.dtyp) list) list)) list;
 
 type nominal_datatype_info =
   {index : int,
@@ -200,7 +200,7 @@
     val (full_new_type_names',thy1) = Datatype.add_datatype config dts'' thy;
 
     val {descr, induct, ...} = Datatype.the_info thy1 (hd full_new_type_names');
-    fun nth_dtyp i = Datatype_Aux.typ_of_dtyp descr (Datatype_Aux.DtRec i);
+    fun nth_dtyp i = Datatype_Aux.typ_of_dtyp descr (Datatype.DtRec i);
 
     val big_name = space_implode "_" new_type_names;
 
@@ -465,13 +465,13 @@
           | (i, _) => SOME (Datatype_Aux.name_of_typ (nth_dtyp i))) descr)) ^ "_set";
     val _ = warning ("big_rep_name: " ^ big_rep_name);
 
-    fun strip_option (dtf as Datatype_Aux.DtType ("fun", [dt, Datatype_Aux.DtRec i])) =
+    fun strip_option (dtf as Datatype.DtType ("fun", [dt, Datatype.DtRec i])) =
           (case AList.lookup op = descr i of
              SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
                apfst (cons dt) (strip_option dt')
            | _ => ([], dtf))
-      | strip_option (Datatype_Aux.DtType ("fun",
-            [dt, Datatype_Aux.DtType ("Nominal.noption", [dt'])])) =
+      | strip_option (Datatype.DtType ("fun",
+            [dt, Datatype.DtType ("Nominal.noption", [dt'])])) =
           apfst (cons dt) (strip_option dt')
       | strip_option dt = ([], dt);
 
@@ -497,7 +497,7 @@
               end
           in (j + 1, j' + length Ts,
             case dt'' of
-                Datatype_Aux.DtRec k => list_all (map (pair "x") Us,
+                Datatype.DtRec k => list_all (map (pair "x") Us,
                   HOLogic.mk_Trueprop (Free (nth rep_set_names k,
                     T --> HOLogic.boolT) $ free')) :: prems
               | _ => prems,
@@ -676,8 +676,8 @@
       (fn ((i, ("Nominal.noption", _, _)), p) => p
         | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
 
-    fun reindex (Datatype_Aux.DtType (s, dts)) = Datatype_Aux.DtType (s, map reindex dts)
-      | reindex (Datatype_Aux.DtRec i) = Datatype_Aux.DtRec (the (AList.lookup op = ty_idxs i))
+    fun reindex (Datatype.DtType (s, dts)) = Datatype.DtType (s, map reindex dts)
+      | reindex (Datatype.DtRec i) = Datatype.DtRec (the (AList.lookup op = ty_idxs i))
       | reindex dt = dt;
 
     fun strip_suffix i s = implode (List.take (raw_explode s, size s - i));  (* FIXME Symbol.explode (?) *)
@@ -713,7 +713,7 @@
       map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
         (constrs ~~ idxss)))) (descr'' ~~ ndescr);
 
-    fun nth_dtyp' i = Datatype_Aux.typ_of_dtyp descr'' (Datatype_Aux.DtRec i);
+    fun nth_dtyp' i = Datatype_Aux.typ_of_dtyp descr'' (Datatype.DtRec i);
 
     val rep_names = map (fn s =>
       Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
@@ -740,7 +740,7 @@
              xs @ x :: l_args,
              fold_rev mk_abs_fun xs
                (case dt of
-                  Datatype_Aux.DtRec k => if k < length new_type_names then
+                  Datatype.DtRec k => if k < length new_type_names then
                       Const (nth rep_names k, Datatype_Aux.typ_of_dtyp descr'' dt -->
                         Datatype_Aux.typ_of_dtyp descr dt) $ x
                     else error "nested recursion not (yet) supported"
@@ -1435,7 +1435,7 @@
         val binders = maps fst frees';
         val atomTs = distinct op = (maps (map snd o fst) frees');
         val recs = map_filter
-          (fn ((_, Datatype_Aux.DtRec i), p) => SOME (i, p) | _ => NONE)
+          (fn ((_, Datatype.DtRec i), p) => SOME (i, p) | _ => NONE)
           (partition_cargs idxs cargs ~~ frees');
         val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
           map (fn (i, _) => nth rec_result_Ts i) recs;
--- a/src/HOL/Tools/ATP/atp_util.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/ATP/atp_util.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -23,9 +23,7 @@
   val varify_type : Proof.context -> typ -> typ
   val instantiate_type : theory -> typ -> typ -> typ -> typ
   val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
-  val typ_of_dtyp :
-    Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
-    -> typ
+  val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
   val is_type_surely_finite : Proof.context -> typ -> bool
   val is_type_surely_infinite : Proof.context -> bool -> typ list -> typ -> bool
   val s_not : term -> term
@@ -148,11 +146,11 @@
     instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
   end
 
-fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
-    the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
-  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
+fun typ_of_dtyp _ typ_assoc (Datatype.DtTFree a) =
+    the (AList.lookup (op =) typ_assoc (Datatype.DtTFree a))
+  | typ_of_dtyp descr typ_assoc (Datatype.DtType (s, Us)) =
     Type (s, map (typ_of_dtyp descr typ_assoc) Us)
-  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
+  | typ_of_dtyp descr typ_assoc (Datatype.DtRec i) =
     let val (s, ds, _) = the (AList.lookup (op =) descr i) in
       Type (s, map (typ_of_dtyp descr typ_assoc) ds)
     end
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -8,7 +8,8 @@
 
 signature DATATYPE_ABS_PROOFS =
 sig
-  include DATATYPE_COMMON
+  type config = Datatype_Aux.config
+  type descr = Datatype_Aux.descr
   val prove_casedist_thms : config -> string list -> descr list -> thm ->
     attribute list -> theory -> thm list * theory
   val prove_primrec_thms : config -> string list -> descr list ->
@@ -29,10 +30,13 @@
 structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS =
 struct
 
+type config = Datatype_Aux.config;
+type descr = Datatype_Aux.descr;
+
+
 (************************ case distinction theorems ***************************)
 
-fun prove_casedist_thms (config : Datatype_Aux.config)
-    new_type_names descr induct case_names_exhausts thy =
+fun prove_casedist_thms (config : config) new_type_names descr induct case_names_exhausts thy =
   let
     val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
 
@@ -79,7 +83,7 @@
 
 (*************************** primrec combinators ******************************)
 
-fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr
+fun prove_primrec_thms (config : config) new_type_names descr
     injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
   let
     val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
@@ -275,8 +279,7 @@
 
 (***************************** case combinators *******************************)
 
-fun prove_case_thms (config : Datatype_Aux.config)
-    new_type_names descr reccomb_names primrec_thms thy =
+fun prove_case_thms (config : config) new_type_names descr reccomb_names primrec_thms thy =
   let
     val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
 
@@ -350,7 +353,7 @@
 
 (******************************* case splitting *******************************)
 
-fun prove_split_thms (config : Datatype_Aux.config)
+fun prove_split_thms (config : config)
     new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
   let
     val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
@@ -399,7 +402,7 @@
 
 (************************* additional theorems for TFL ************************)
 
-fun prove_nchotomys (config : Datatype_Aux.config) new_type_names descr casedist_thms thy =
+fun prove_nchotomys (config : config) new_type_names descr casedist_thms thy =
   let
     val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
 
@@ -449,7 +452,4 @@
 
   in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
 
-
-open Datatype_Aux;
-
 end;
--- a/src/HOL/Tools/Datatype/datatype_prop.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -6,7 +6,7 @@
 
 signature DATATYPE_PROP =
 sig
-  include DATATYPE_COMMON
+  type descr = Datatype_Aux.descr
   val indexify_names: string list -> string list
   val make_tnames: typ list -> string list
   val make_injs : descr list -> term list list
@@ -26,6 +26,9 @@
 structure Datatype_Prop : DATATYPE_PROP =
 struct
 
+type descr = Datatype_Aux.descr;
+
+
 fun indexify_names names =
   let
     fun index (x :: xs) tab =
@@ -429,7 +432,4 @@
       (hd descr ~~ newTs)
   end;
 
-
-open Datatype_Aux;
-
 end;
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -25,7 +25,7 @@
 fun tname_of (Type (s, _)) = s
   | tname_of _ = "";
 
-fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Datatype_Aux.info) is thy =
+fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Datatype.info) is thy =
   let
     val recTs = Datatype_Aux.get_rec_types descr;
     val pnames =
@@ -157,7 +157,7 @@
   in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
 
 
-fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Datatype_Aux.info) thy =
+fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Datatype.info) thy =
   let
     val cert = cterm_of thy;
     val rT = TFree ("'P", HOLogic.typeS);
--- a/src/HOL/Tools/Function/size.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/Function/size.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -40,7 +40,7 @@
       NONE => Abs ("x", T, HOLogic.zero)
     | SOME t => t);
 
-fun is_poly thy (Datatype_Aux.DtType (name, dts)) =
+fun is_poly thy (Datatype.DtType (name, dts)) =
       (case Datatype.get_info thy name of
          NONE => false
        | SOME _ => exists (is_poly thy) dts)
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -61,9 +61,7 @@
   val int_T : typ
   val simple_string_of_typ : typ -> string
   val is_real_constr : theory -> string * typ -> bool
-  val typ_of_dtyp :
-    Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
-    -> typ
+  val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
   val varify_type : Proof.context -> typ -> typ
   val instantiate_type : theory -> typ -> typ -> typ -> typ
   val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
--- a/src/HOL/Tools/refute.ML	Thu Dec 15 21:46:52 2011 +0100
+++ b/src/HOL/Tools/refute.ML	Fri Dec 16 10:38:38 2011 +0100
@@ -862,7 +862,7 @@
                 (* sanity check: every element in 'dtyps' must be a *)
                 (* 'DtTFree'                                        *)
                 val _ = if Library.exists (fn d =>
-                  case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then
+                  case d of Datatype.DtTFree _ => false | _ => true) typs then
                   raise REFUTE ("ground_types", "datatype argument (for type "
                     ^ Syntax.string_of_typ ctxt T ^ ") is not a variable")
                 else ()
@@ -874,11 +874,11 @@
                     val dT = typ_of_dtyp descr typ_assoc d
                   in
                     case d of
-                      Datatype_Aux.DtTFree _ =>
+                      Datatype.DtTFree _ =>
                       collect_types dT acc
-                    | Datatype_Aux.DtType (_, ds) =>
+                    | Datatype.DtType (_, ds) =>
                       collect_types dT (fold_rev collect_dtyp ds acc)
-                    | Datatype_Aux.DtRec i =>
+                    | Datatype.DtRec i =>
                       if member (op =) acc dT then
                         acc  (* prevent infinite recursion *)
                       else
@@ -901,7 +901,7 @@
               in
                 (* argument types 'Ts' could be added here, but they are also *)
                 (* added by 'collect_dtyp' automatically                      *)
-                collect_dtyp (Datatype_Aux.DtRec index) acc
+                collect_dtyp (Datatype.DtRec index) acc
               end
           | NONE =>
             (* not an inductive datatype, e.g. defined via "typedef" or *)
@@ -1853,7 +1853,7 @@
                     (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
                     val _ =
                       if Library.exists (fn d =>
-                        case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+                        case d of Datatype.DtTFree _ => false | _ => true) dtyps
                       then
                         raise REFUTE ("IDT_interpreter",
                           "datatype argument (for type "
@@ -1975,7 +1975,7 @@
                   (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
                   val _ =
                     if Library.exists (fn d =>
-                      case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+                      case d of Datatype.DtTFree _ => false | _ => true) dtyps
                     then
                       raise REFUTE ("IDT_constructor_interpreter",
                         "datatype argument (for type "
@@ -2035,7 +2035,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 Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+                          case d of Datatype.DtTFree _ => false | _ => true) dtyps
                         then
                           raise REFUTE ("IDT_constructor_interpreter",
                             "datatype argument (for type "
@@ -2250,7 +2250,7 @@
                       (*               'DtTFree'                          *)
                       val _ =
                         if Library.exists (fn d =>
-                          case d of Datatype_Aux.DtTFree _ => false
+                          case d of Datatype.DtTFree _ => false
                                   | _ => true) dtyps
                         then
                           raise REFUTE ("IDT_recursion_interpreter",
@@ -2271,10 +2271,10 @@
                             (case AList.lookup op= acc d of
                               NONE =>
                                 (case d of
-                                  Datatype_Aux.DtTFree _ =>
+                                  Datatype.DtTFree _ =>
                                   (* add the association, proceed *)
                                   rec_typ_assoc ((d, T)::acc) xs
-                                | Datatype_Aux.DtType (s, ds) =>
+                                | Datatype.DtType (s, ds) =>
                                     let
                                       val (s', Ts) = dest_Type T
                                     in
@@ -2284,7 +2284,7 @@
                                         raise REFUTE ("IDT_recursion_interpreter",
                                           "DtType/Type mismatch")
                                     end
-                                | Datatype_Aux.DtRec i =>
+                                | Datatype.DtRec i =>
                                     let
                                       val (_, ds, _) = the (AList.lookup (op =) descr i)
                                       val (_, Ts)    = dest_Type T
@@ -2300,7 +2300,7 @@
                                   raise REFUTE ("IDT_recursion_interpreter",
                                     "different type associations for the same dtyp"))
                       val typ_assoc = filter
-                        (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false)
+                        (fn (Datatype.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   *)
@@ -2317,7 +2317,7 @@
                       val mc_intrs = map (fn (idx, (_, _, cs)) =>
                         let
                           val c_return_typ = typ_of_dtyp descr typ_assoc
-                            (Datatype_Aux.DtRec idx)
+                            (Datatype.DtRec idx)
                         in
                           (idx, map (fn (cname, cargs) =>
                             (#1 o interpret ctxt (typs, []) {maxvars=0,
@@ -2371,7 +2371,7 @@
                       val _ = map (fn (idx, intrs) =>
                         let
                           val T = typ_of_dtyp descr typ_assoc
-                            (Datatype_Aux.DtRec idx)
+                            (Datatype.DtRec idx)
                         in
                           if length intrs <> size_of_type ctxt (typs, []) T then
                             raise REFUTE ("IDT_recursion_interpreter",
@@ -2465,17 +2465,17 @@
                               (* takes the dtyp and interpretation of an element, *)
                               (* and computes the interpretation for the          *)
                               (* corresponding recursive argument                 *)
-                              fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) =
+                              fun rec_intr (Datatype.DtRec i) (Leaf xs) =
                                     (* recursive argument is "rec_i params elem" *)
                                     compute_array_entry i (find_index (fn x => x = True) xs)
-                                | rec_intr (Datatype_Aux.DtRec _) (Node _) =
+                                | rec_intr (Datatype.DtRec _) (Node _) =
                                     raise REFUTE ("IDT_recursion_interpreter",
                                       "interpretation for IDT is a node")
-                                | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2])) (Node xs) =
+                                | rec_intr (Datatype.DtType ("fun", [dt1, dt2])) (Node xs) =
                                     (* recursive argument is something like     *)
                                     (* "\<lambda>x::dt1. rec_? params (elem x)" *)
                                     Node (map (rec_intr dt2) xs)
-                                | rec_intr (Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) =
+                                | rec_intr (Datatype.DtType ("fun", [_, _])) (Leaf _) =
                                     raise REFUTE ("IDT_recursion_interpreter",
                                       "interpretation for function dtyp is a leaf")
                                 | rec_intr _ _ =
@@ -3024,7 +3024,7 @@
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ =
                 if Library.exists (fn d =>
-                  case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+                  case d of Datatype.DtTFree _ => false | _ => true) dtyps
                 then
                   raise REFUTE ("IDT_printer", "datatype argument (for type " ^
                     Syntax.string_of_typ ctxt (Type (s, Ts)) ^ ") is not a variable")