src/HOL/Tools/datatype_package/datatype_aux.ML
changeset 31605 92d7a5094e23
parent 30364 577edc39b501
child 31668 a616e56a5ec8
equal deleted inserted replaced
31604:eb2f9d709296 31605:92d7a5094e23
       
     1 (*  Title:      HOL/Tools/datatype_aux.ML
       
     2     Author:     Stefan Berghofer, TU Muenchen
       
     3 
       
     4 Auxiliary functions for defining datatypes.
       
     5 *)
       
     6 
       
     7 signature DATATYPE_AUX =
       
     8 sig
       
     9   val quiet_mode : bool ref
       
    10   val message : string -> unit
       
    11   
       
    12   val add_path : bool -> string -> theory -> theory
       
    13   val parent_path : bool -> theory -> theory
       
    14 
       
    15   val store_thmss_atts : string -> string list -> attribute list list -> thm list list
       
    16     -> theory -> thm list list * theory
       
    17   val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory
       
    18   val store_thms_atts : string -> string list -> attribute list list -> thm list
       
    19     -> theory -> thm list * theory
       
    20   val store_thms : string -> string list -> thm list -> theory -> thm list * theory
       
    21 
       
    22   val split_conj_thm : thm -> thm list
       
    23   val mk_conj : term list -> term
       
    24   val mk_disj : term list -> term
       
    25 
       
    26   val app_bnds : term -> int -> term
       
    27 
       
    28   val cong_tac : int -> tactic
       
    29   val indtac : thm -> string list -> int -> tactic
       
    30   val exh_tac : (string -> thm) -> int -> tactic
       
    31 
       
    32   datatype simproc_dist = FewConstrs of thm list
       
    33                         | ManyConstrs of thm * simpset;
       
    34 
       
    35   datatype dtyp =
       
    36       DtTFree of string
       
    37     | DtType of string * (dtyp list)
       
    38     | DtRec of int;
       
    39   type descr
       
    40   type datatype_info
       
    41 
       
    42   exception Datatype
       
    43   exception Datatype_Empty of string
       
    44   val name_of_typ : typ -> string
       
    45   val dtyp_of_typ : (string * string list) list -> typ -> dtyp
       
    46   val mk_Free : string -> typ -> int -> term
       
    47   val is_rec_type : dtyp -> bool
       
    48   val typ_of_dtyp : descr -> (string * sort) list -> dtyp -> typ
       
    49   val dest_DtTFree : dtyp -> string
       
    50   val dest_DtRec : dtyp -> int
       
    51   val strip_dtyp : dtyp -> dtyp list * dtyp
       
    52   val body_index : dtyp -> int
       
    53   val mk_fun_dtyp : dtyp list -> dtyp -> dtyp
       
    54   val get_nonrec_types : descr -> (string * sort) list -> typ list
       
    55   val get_branching_types : descr -> (string * sort) list -> typ list
       
    56   val get_arities : descr -> int list
       
    57   val get_rec_types : descr -> (string * sort) list -> typ list
       
    58   val interpret_construction : descr -> (string * sort) list
       
    59     -> { atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a }
       
    60     -> ((string * Term.typ list) * (string * 'a list) list) list
       
    61   val check_nonempty : descr list -> unit
       
    62   val unfold_datatypes : 
       
    63     theory -> descr -> (string * sort) list -> datatype_info Symtab.table ->
       
    64       descr -> int -> descr list * int
       
    65 end;
       
    66 
       
    67 structure DatatypeAux : DATATYPE_AUX =
       
    68 struct
       
    69 
       
    70 val quiet_mode = ref false;
       
    71 fun message s = if !quiet_mode then () else writeln s;
       
    72 
       
    73 fun add_path flat_names s = if flat_names then I else Sign.add_path s;
       
    74 fun parent_path flat_names = if flat_names then I else Sign.parent_path;
       
    75 
       
    76 
       
    77 (* store theorems in theory *)
       
    78 
       
    79 fun store_thmss_atts label tnames attss thmss =
       
    80   fold_map (fn ((tname, atts), thms) =>
       
    81     Sign.add_path tname
       
    82     #> PureThy.add_thmss [((Binding.name label, thms), atts)]
       
    83     #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
       
    84   ##> Theory.checkpoint;
       
    85 
       
    86 fun store_thmss label tnames = store_thmss_atts label tnames (replicate (length tnames) []);
       
    87 
       
    88 fun store_thms_atts label tnames attss thmss =
       
    89   fold_map (fn ((tname, atts), thms) =>
       
    90     Sign.add_path tname
       
    91     #> PureThy.add_thms [((Binding.name label, thms), atts)]
       
    92     #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
       
    93   ##> Theory.checkpoint;
       
    94 
       
    95 fun store_thms label tnames = store_thms_atts label tnames (replicate (length tnames) []);
       
    96 
       
    97 
       
    98 (* split theorem thm_1 & ... & thm_n into n theorems *)
       
    99 
       
   100 fun split_conj_thm th =
       
   101   ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
       
   102 
       
   103 val mk_conj = foldr1 (HOLogic.mk_binop "op &");
       
   104 val mk_disj = foldr1 (HOLogic.mk_binop "op |");
       
   105 
       
   106 fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
       
   107 
       
   108 
       
   109 fun cong_tac i st = (case Logic.strip_assums_concl
       
   110   (List.nth (prems_of st, i - 1)) of
       
   111     _ $ (_ $ (f $ x) $ (g $ y)) =>
       
   112       let
       
   113         val cong' = Thm.lift_rule (Thm.cprem_of st i) cong;
       
   114         val _ $ (_ $ (f' $ x') $ (g' $ y')) =
       
   115           Logic.strip_assums_concl (prop_of cong');
       
   116         val insts = map (pairself (cterm_of (Thm.theory_of_thm st)) o
       
   117           apsnd (curry list_abs (Logic.strip_params (concl_of cong'))) o
       
   118             apfst head_of) [(f', f), (g', g), (x', x), (y', y)]
       
   119       in compose_tac (false, cterm_instantiate insts cong', 2) i st
       
   120         handle THM _ => no_tac st
       
   121       end
       
   122   | _ => no_tac st);
       
   123 
       
   124 (* instantiate induction rule *)
       
   125 
       
   126 fun indtac indrule indnames i st =
       
   127   let
       
   128     val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
       
   129     val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
       
   130       (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
       
   131     val getP = if can HOLogic.dest_imp (hd ts) then
       
   132       (apfst SOME) o HOLogic.dest_imp else pair NONE;
       
   133     val flt = if null indnames then I else
       
   134       filter (fn Free (s, _) => s mem indnames | _ => false);
       
   135     fun abstr (t1, t2) = (case t1 of
       
   136         NONE => (case flt (OldTerm.term_frees t2) of
       
   137             [Free (s, T)] => SOME (absfree (s, T, t2))
       
   138           | _ => NONE)
       
   139       | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
       
   140     val cert = cterm_of (Thm.theory_of_thm st);
       
   141     val insts = List.mapPartial (fn (t, u) => case abstr (getP u) of
       
   142         NONE => NONE
       
   143       | SOME u' => SOME (t |> getP |> snd |> head_of |> cert, cert u')) (ts ~~ ts');
       
   144     val indrule' = cterm_instantiate insts indrule
       
   145   in
       
   146     rtac indrule' i st
       
   147   end;
       
   148 
       
   149 (* perform exhaustive case analysis on last parameter of subgoal i *)
       
   150 
       
   151 fun exh_tac exh_thm_of i state =
       
   152   let
       
   153     val thy = Thm.theory_of_thm state;
       
   154     val prem = nth (prems_of state) (i - 1);
       
   155     val params = Logic.strip_params prem;
       
   156     val (_, Type (tname, _)) = hd (rev params);
       
   157     val exhaustion = Thm.lift_rule (Thm.cprem_of state i) (exh_thm_of tname);
       
   158     val prem' = hd (prems_of exhaustion);
       
   159     val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
       
   160     val exhaustion' = cterm_instantiate [(cterm_of thy (head_of lhs),
       
   161       cterm_of thy (List.foldr (fn ((_, T), t) => Abs ("z", T, t))
       
   162         (Bound 0) params))] exhaustion
       
   163   in compose_tac (false, exhaustion', nprems_of exhaustion) i state
       
   164   end;
       
   165 
       
   166 (* handling of distinctness theorems *)
       
   167 
       
   168 datatype simproc_dist = FewConstrs of thm list
       
   169                       | ManyConstrs of thm * simpset;
       
   170 
       
   171 (********************** Internal description of datatypes *********************)
       
   172 
       
   173 datatype dtyp =
       
   174     DtTFree of string
       
   175   | DtType of string * (dtyp list)
       
   176   | DtRec of int;
       
   177 
       
   178 (* information about datatypes *)
       
   179 
       
   180 (* index, datatype name, type arguments, constructor name, types of constructor's arguments *)
       
   181 type descr = (int * (string * dtyp list * (string * dtyp list) list)) list;
       
   182 
       
   183 type datatype_info =
       
   184   {index : int,
       
   185    alt_names : string list option,
       
   186    descr : descr,
       
   187    sorts : (string * sort) list,
       
   188    rec_names : string list,
       
   189    rec_rewrites : thm list,
       
   190    case_name : string,
       
   191    case_rewrites : thm list,
       
   192    induction : thm,
       
   193    exhaustion : thm,
       
   194    distinct : simproc_dist,
       
   195    inject : thm list,
       
   196    nchotomy : thm,
       
   197    case_cong : thm,
       
   198    weak_case_cong : thm};
       
   199 
       
   200 fun mk_Free s T i = Free (s ^ (string_of_int i), T);
       
   201 
       
   202 fun subst_DtTFree _ substs (T as (DtTFree name)) =
       
   203       AList.lookup (op =) substs name |> the_default T
       
   204   | subst_DtTFree i substs (DtType (name, ts)) =
       
   205       DtType (name, map (subst_DtTFree i substs) ts)
       
   206   | subst_DtTFree i _ (DtRec j) = DtRec (i + j);
       
   207 
       
   208 exception Datatype;
       
   209 exception Datatype_Empty of string;
       
   210 
       
   211 fun dest_DtTFree (DtTFree a) = a
       
   212   | dest_DtTFree _ = raise Datatype;
       
   213 
       
   214 fun dest_DtRec (DtRec i) = i
       
   215   | dest_DtRec _ = raise Datatype;
       
   216 
       
   217 fun is_rec_type (DtType (_, dts)) = exists is_rec_type dts
       
   218   | is_rec_type (DtRec _) = true
       
   219   | is_rec_type _ = false;
       
   220 
       
   221 fun strip_dtyp (DtType ("fun", [T, U])) = apfst (cons T) (strip_dtyp U)
       
   222   | strip_dtyp T = ([], T);
       
   223 
       
   224 val body_index = dest_DtRec o snd o strip_dtyp;
       
   225 
       
   226 fun mk_fun_dtyp [] U = U
       
   227   | mk_fun_dtyp (T :: Ts) U = DtType ("fun", [T, mk_fun_dtyp Ts U]);
       
   228 
       
   229 fun name_of_typ (Type (s, Ts)) =
       
   230       let val s' = Long_Name.base_name s
       
   231       in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
       
   232         [if Syntax.is_identifier s' then s' else "x"])
       
   233       end
       
   234   | name_of_typ _ = "";
       
   235 
       
   236 fun dtyp_of_typ _ (TFree (n, _)) = DtTFree n
       
   237   | dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
       
   238   | dtyp_of_typ new_dts (Type (tname, Ts)) =
       
   239       (case AList.lookup (op =) new_dts tname of
       
   240          NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
       
   241        | SOME vs => if map (try (fst o dest_TFree)) Ts = map SOME vs then
       
   242              DtRec (find_index (curry op = tname o fst) new_dts)
       
   243            else error ("Illegal occurrence of recursive type " ^ tname));
       
   244 
       
   245 fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, (the o AList.lookup (op =) sorts) a)
       
   246   | typ_of_dtyp descr sorts (DtRec i) =
       
   247       let val (s, ds, _) = (the o AList.lookup (op =) descr) i
       
   248       in Type (s, map (typ_of_dtyp descr sorts) ds) end
       
   249   | typ_of_dtyp descr sorts (DtType (s, ds)) =
       
   250       Type (s, map (typ_of_dtyp descr sorts) ds);
       
   251 
       
   252 (* find all non-recursive types in datatype description *)
       
   253 
       
   254 fun get_nonrec_types descr sorts =
       
   255   map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
       
   256     Library.foldl (fn (Ts', (_, cargs)) =>
       
   257       filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr));
       
   258 
       
   259 (* get all recursive types in datatype description *)
       
   260 
       
   261 fun get_rec_types descr sorts = map (fn (_ , (s, ds, _)) =>
       
   262   Type (s, map (typ_of_dtyp descr sorts) ds)) descr;
       
   263 
       
   264 (* get all branching types *)
       
   265 
       
   266 fun get_branching_types descr sorts =
       
   267   map (typ_of_dtyp descr sorts) (fold (fn (_, (_, _, constrs)) =>
       
   268     fold (fn (_, cargs) => fold (strip_dtyp #> fst #> fold (insert op =)) cargs)
       
   269       constrs) descr []);
       
   270 
       
   271 fun get_arities descr = fold (fn (_, (_, _, constrs)) =>
       
   272   fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp)
       
   273     (List.filter is_rec_type cargs))) constrs) descr [];
       
   274 
       
   275 (* interpret construction of datatype *)
       
   276 
       
   277 fun interpret_construction descr vs { atyp, dtyp } =
       
   278   let
       
   279     val typ_of_dtyp = typ_of_dtyp descr vs;
       
   280     fun interpT dT = case strip_dtyp dT
       
   281      of (dTs, DtRec l) =>
       
   282           let
       
   283             val (tyco, dTs', _) = (the o AList.lookup (op =) descr) l;
       
   284             val Ts = map typ_of_dtyp dTs;
       
   285             val Ts' = map typ_of_dtyp dTs';
       
   286             val is_proper = forall (can dest_TFree) Ts';
       
   287           in dtyp Ts (l, is_proper) (tyco, Ts') end
       
   288       | _ => atyp (typ_of_dtyp dT);
       
   289     fun interpC (c, dTs) = (c, map interpT dTs);
       
   290     fun interpD (_, (tyco, dTs, cs)) = ((tyco, map typ_of_dtyp dTs), map interpC cs);
       
   291   in map interpD descr end;
       
   292 
       
   293 (* nonemptiness check for datatypes *)
       
   294 
       
   295 fun check_nonempty descr =
       
   296   let
       
   297     val descr' = List.concat descr;
       
   298     fun is_nonempty_dt is i =
       
   299       let
       
   300         val (_, _, constrs) = (the o AList.lookup (op =) descr') i;
       
   301         fun arg_nonempty (_, DtRec i) = if i mem is then false
       
   302               else is_nonempty_dt (i::is) i
       
   303           | arg_nonempty _ = true;
       
   304       in exists ((forall (arg_nonempty o strip_dtyp)) o snd) constrs
       
   305       end
       
   306   in assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
       
   307     (fn (_, (s, _, _)) => raise Datatype_Empty s)
       
   308   end;
       
   309 
       
   310 (* unfold a list of mutually recursive datatype specifications *)
       
   311 (* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
       
   312 (* need to be unfolded                                         *)
       
   313 
       
   314 fun unfold_datatypes sign orig_descr sorts (dt_info : datatype_info Symtab.table) descr i =
       
   315   let
       
   316     fun typ_error T msg = error ("Non-admissible type expression\n" ^
       
   317       Syntax.string_of_typ_global sign (typ_of_dtyp (orig_descr @ descr) sorts T) ^ "\n" ^ msg);
       
   318 
       
   319     fun get_dt_descr T i tname dts =
       
   320       (case Symtab.lookup dt_info tname of
       
   321          NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
       
   322            \ nested recursion")
       
   323        | (SOME {index, descr, ...}) =>
       
   324            let val (_, vars, _) = (the o AList.lookup (op =) descr) index;
       
   325                val subst = ((map dest_DtTFree vars) ~~ dts) handle Library.UnequalLengths =>
       
   326                  typ_error T ("Type constructor " ^ tname ^ " used with wrong\
       
   327                   \ number of arguments")
       
   328            in (i + index, map (fn (j, (tn, args, cs)) => (i + j,
       
   329              (tn, map (subst_DtTFree i subst) args,
       
   330               map (apsnd (map (subst_DtTFree i subst))) cs))) descr)
       
   331            end);
       
   332 
       
   333     (* unfold a single constructor argument *)
       
   334 
       
   335     fun unfold_arg ((i, Ts, descrs), T) =
       
   336       if is_rec_type T then
       
   337         let val (Us, U) = strip_dtyp T
       
   338         in if exists is_rec_type Us then
       
   339             typ_error T "Non-strictly positive recursive occurrence of type"
       
   340           else (case U of
       
   341               DtType (tname, dts) =>  
       
   342                 let
       
   343                   val (index, descr) = get_dt_descr T i tname dts;
       
   344                   val (descr', i') = unfold_datatypes sign orig_descr sorts
       
   345                     dt_info descr (i + length descr)
       
   346                 in (i', Ts @ [mk_fun_dtyp Us (DtRec index)], descrs @ descr') end
       
   347             | _ => (i, Ts @ [T], descrs))
       
   348         end
       
   349       else (i, Ts @ [T], descrs);
       
   350 
       
   351     (* unfold a constructor *)
       
   352 
       
   353     fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
       
   354       let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
       
   355       in (i', constrs @ [(cname, cargs')], descrs') end;
       
   356 
       
   357     (* unfold a single datatype *)
       
   358 
       
   359     fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
       
   360       let val (i', constrs', descrs') =
       
   361         Library.foldl unfold_constr ((i, [], descrs), constrs)
       
   362       in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
       
   363       end;
       
   364 
       
   365     val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
       
   366 
       
   367   in (descr' :: descrs, i') end;
       
   368 
       
   369 end;