| author | bulwahn | 
| Mon, 29 Mar 2010 17:30:53 +0200 | |
| changeset 36034 | ee84eac07290 | 
| parent 35845 | e5980f0ad025 | 
| child 36610 | bafd82950e24 | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 1 | (* Title: HOL/Tools/old_primrec.ML | 
| 29266 | 2 | Author: Norbert Voelker, FernUni Hagen | 
| 3 | Author: Stefan Berghofer, TU Muenchen | |
| 25557 | 4 | |
| 5 | Package for defining functions on datatypes by primitive recursion. | |
| 6 | *) | |
| 7 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 8 | signature OLD_PRIMREC = | 
| 25557 | 9 | sig | 
| 10 | val unify_consts: theory -> term list -> term list -> term list * term list | |
| 11 | val add_primrec: string -> ((bstring * string) * Attrib.src list) list | |
| 12 | -> theory -> thm list * theory | |
| 13 | val add_primrec_unchecked: string -> ((bstring * string) * Attrib.src list) list | |
| 14 | -> theory -> thm list * theory | |
| 15 | val add_primrec_i: string -> ((bstring * term) * attribute list) list | |
| 16 | -> theory -> thm list * theory | |
| 17 | val add_primrec_unchecked_i: string -> ((bstring * term) * attribute list) list | |
| 18 | -> theory -> thm list * theory | |
| 19 | end; | |
| 20 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 21 | structure OldPrimrec : OLD_PRIMREC = | 
| 25557 | 22 | struct | 
| 23 | ||
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33963diff
changeset | 24 | open Datatype_Aux; | 
| 25557 | 25 | |
| 26 | exception RecError of string; | |
| 27 | ||
| 28 | fun primrec_err s = error ("Primrec definition error:\n" ^ s);
 | |
| 29 | fun primrec_eq_err thy s eq = | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26478diff
changeset | 30 | primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); | 
| 25557 | 31 | |
| 32 | ||
| 33 | (*the following code ensures that each recursive set always has the | |
| 34 | same type in all introduction rules*) | |
| 35 | fun unify_consts thy cs intr_ts = | |
| 36 | (let | |
| 33338 | 37 | fun varify t (i, ts) = | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35756diff
changeset | 38 | let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify_global [] t)) | 
| 25557 | 39 | in (maxidx_of_term t', t'::ts) end; | 
| 33338 | 40 | val (i, cs') = fold_rev varify cs (~1, []); | 
| 41 | val (i', intr_ts') = fold_rev varify intr_ts (i, []); | |
| 29290 | 42 | val rec_consts = fold Term.add_consts cs' []; | 
| 43 | val intr_consts = fold Term.add_consts intr_ts' []; | |
| 25557 | 44 | fun unify (cname, cT) = | 
| 45 | let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts) | |
| 46 | in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end; | |
| 47 | val (env, _) = fold unify rec_consts (Vartab.empty, i'); | |
| 33832 | 48 | val subst = Type.legacy_freeze o map_types (Envir.norm_type env) | 
| 25557 | 49 | |
| 50 | in (map subst cs', map subst intr_ts') | |
| 51 | end) handle Type.TUNIFY => | |
| 52 | (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts)); | |
| 53 | ||
| 54 | ||
| 55 | (* preprocessing of equations *) | |
| 56 | ||
| 57 | fun process_eqn thy eq rec_fns = | |
| 58 | let | |
| 59 | val (lhs, rhs) = | |
| 29266 | 60 | if null (Term.add_vars eq []) then | 
| 25557 | 61 | HOLogic.dest_eq (HOLogic.dest_Trueprop eq) | 
| 62 | handle TERM _ => raise RecError "not a proper equation" | |
| 63 | else raise RecError "illegal schematic variable(s)"; | |
| 64 | ||
| 65 | val (recfun, args) = strip_comb lhs; | |
| 66 | val fnameT = dest_Const recfun handle TERM _ => | |
| 67 | raise RecError "function is not declared as constant in theory"; | |
| 68 | ||
| 69 | val (ls', rest) = take_prefix is_Free args; | |
| 70 | val (middle, rs') = take_suffix is_Free rest; | |
| 71 | val rpos = length ls'; | |
| 72 | ||
| 73 | val (constr, cargs') = if null middle then raise RecError "constructor missing" | |
| 74 | else strip_comb (hd middle); | |
| 75 | val (cname, T) = dest_Const constr | |
| 76 | handle TERM _ => raise RecError "ill-formed constructor"; | |
| 77 | val (tname, _) = dest_Type (body_type T) handle TYPE _ => | |
| 78 | raise RecError "cannot determine datatype associated with function" | |
| 79 | ||
| 80 | val (ls, cargs, rs) = | |
| 81 | (map dest_Free ls', map dest_Free cargs', map dest_Free rs') | |
| 82 | handle TERM _ => raise RecError "illegal argument in pattern"; | |
| 83 | val lfrees = ls @ rs @ cargs; | |
| 84 | ||
| 85 | fun check_vars _ [] = () | |
| 86 | | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars)) | |
| 87 | in | |
| 88 | if length middle > 1 then | |
| 89 | raise RecError "more than one non-variable in pattern" | |
| 90 | else | |
| 91 | (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); | |
| 92 | check_vars "extra variables on rhs: " | |
| 33040 | 93 | (subtract (op =) lfrees (map dest_Free (OldTerm.term_frees rhs))); | 
| 25557 | 94 | case AList.lookup (op =) rec_fns fnameT of | 
| 95 | NONE => | |
| 96 | (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns | |
| 97 | | SOME (_, rpos', eqns) => | |
| 98 | if AList.defined (op =) eqns cname then | |
| 99 | raise RecError "constructor already occurred as pattern" | |
| 100 | else if rpos <> rpos' then | |
| 101 | raise RecError "position of recursive argument inconsistent" | |
| 102 | else | |
| 103 | AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) | |
| 104 | rec_fns) | |
| 105 | end | |
| 106 | handle RecError s => primrec_eq_err thy s eq; | |
| 107 | ||
| 108 | fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) = | |
| 109 | let | |
| 110 | val (_, (tname, _, constrs)) = List.nth (descr, i); | |
| 111 | ||
| 112 | (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) | |
| 113 | ||
| 114 | fun subst [] t fs = (t, fs) | |
| 115 | | subst subs (Abs (a, T, t)) fs = | |
| 116 | fs | |
| 117 | |> subst subs t | |
| 118 | |-> (fn t' => pair (Abs (a, T, t'))) | |
| 119 | | subst subs (t as (_ $ _)) fs = | |
| 120 | let | |
| 121 | val (f, ts) = strip_comb t; | |
| 122 | in | |
| 123 | if is_Const f andalso dest_Const f mem map fst rec_eqns then | |
| 124 | let | |
| 125 | val fnameT' as (fname', _) = dest_Const f; | |
| 126 | val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT'); | |
| 33957 | 127 | val ls = take rpos ts; | 
| 128 | val rest = drop rpos ts; | |
| 25557 | 129 | val (x', rs) = (hd rest, tl rest) | 
| 130 |                   handle Empty => raise RecError ("not enough arguments\
 | |
| 131 | \ in recursive application\nof function " ^ quote fname' ^ " on rhs"); | |
| 132 | val (x, xs) = strip_comb x' | |
| 133 | in case AList.lookup (op =) subs x | |
| 134 | of NONE => | |
| 135 | fs | |
| 136 | |> fold_map (subst subs) ts | |
| 137 | |-> (fn ts' => pair (list_comb (f, ts'))) | |
| 138 | | SOME (i', y) => | |
| 139 | fs | |
| 140 | |> fold_map (subst subs) (xs @ ls @ rs) | |
| 141 | ||> process_fun thy descr rec_eqns (i', fnameT') | |
| 142 | |-> (fn ts' => pair (list_comb (y, ts'))) | |
| 143 | end | |
| 144 | else | |
| 145 | fs | |
| 146 | |> fold_map (subst subs) (f :: ts) | |
| 147 | |-> (fn (f'::ts') => pair (list_comb (f', ts'))) | |
| 148 | end | |
| 149 | | subst _ t fs = (t, fs); | |
| 150 | ||
| 151 | (* translate rec equations into function arguments suitable for rec comb *) | |
| 152 | ||
| 153 | fun trans eqns (cname, cargs) (fnameTs', fnss', fns) = | |
| 154 | (case AList.lookup (op =) eqns cname of | |
| 155 |           NONE => (warning ("No equation for constructor " ^ quote cname ^
 | |
| 156 | "\nin definition of function " ^ quote fname); | |
| 157 |               (fnameTs', fnss', (Const ("HOL.undefined", dummyT))::fns))
 | |
| 158 | | SOME (ls, cargs', rs, rhs, eq) => | |
| 159 | let | |
| 160 | val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); | |
| 161 | val rargs = map fst recs; | |
| 162 | val subs = map (rpair dummyT o fst) | |
| 29276 | 163 | (rev (Term.rename_wrt_term rhs rargs)); | 
| 25557 | 164 | val (rhs', (fnameTs'', fnss'')) = | 
| 165 | (subst (map (fn ((x, y), z) => | |
| 166 | (Free x, (body_index y, Free z))) | |
| 167 | (recs ~~ subs)) rhs (fnameTs', fnss')) | |
| 168 | handle RecError s => primrec_eq_err thy s eq | |
| 169 | in (fnameTs'', fnss'', | |
| 170 | (list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns) | |
| 171 | end) | |
| 172 | ||
| 173 | in (case AList.lookup (op =) fnameTs i of | |
| 174 | NONE => | |
| 175 | if exists (equal fnameT o snd) fnameTs then | |
| 176 |           raise RecError ("inconsistent functions for datatype " ^ quote tname)
 | |
| 177 | else | |
| 178 | let | |
| 179 | val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT); | |
| 180 | val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs | |
| 181 | ((i, fnameT)::fnameTs, fnss, []) | |
| 182 | in | |
| 183 | (fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') | |
| 184 | end | |
| 185 | | SOME fnameT' => | |
| 186 | if fnameT = fnameT' then (fnameTs, fnss) | |
| 187 |         else raise RecError ("inconsistent functions for datatype " ^ quote tname))
 | |
| 188 | end; | |
| 189 | ||
| 190 | ||
| 191 | (* prepare functions needed for definitions *) | |
| 192 | ||
| 193 | fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) = | |
| 194 | case AList.lookup (op =) fns i of | |
| 195 | NONE => | |
| 196 | let | |
| 197 |          val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
 | |
| 33317 | 198 | replicate (length cargs + length (filter is_rec_type cargs)) | 
| 25557 | 199 | dummyT ---> HOLogic.unitT)) constrs; | 
| 200 |          val _ = warning ("No function definition for datatype " ^ quote tname)
 | |
| 201 | in | |
| 202 | (dummy_fns @ fs, defs) | |
| 203 | end | |
| 204 | | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs); | |
| 205 | ||
| 206 | ||
| 207 | (* make definition *) | |
| 208 | ||
| 209 | fun make_def thy fs (fname, ls, rec_name, tname) = | |
| 210 | let | |
| 211 |     val rhs = fold_rev (fn T => fn t => Abs ("", T, t))
 | |
| 212 | ((map snd ls) @ [dummyT]) | |
| 213 | (list_comb (Const (rec_name, dummyT), | |
| 214 | fs @ map Bound (0 ::(length ls downto 1)))) | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 215 | val def_name = Long_Name.base_name fname ^ "_" ^ Long_Name.base_name tname ^ "_def"; | 
| 25557 | 216 | val def_prop = | 
| 217 | singleton (Syntax.check_terms (ProofContext.init thy)) | |
| 218 | (Logic.mk_equals (Const (fname, dummyT), rhs)); | |
| 219 | in (def_name, def_prop) end; | |
| 220 | ||
| 221 | ||
| 222 | (* find datatypes which contain all datatypes in tnames' *) | |
| 223 | ||
| 31737 | 224 | fun find_dts (dt_info : info Symtab.table) _ [] = [] | 
| 25557 | 225 | | find_dts dt_info tnames' (tname::tnames) = | 
| 226 | (case Symtab.lookup dt_info tname of | |
| 227 | NONE => primrec_err (quote tname ^ " is not a datatype") | |
| 228 | | SOME dt => | |
| 33038 | 229 | if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then | 
| 25557 | 230 | (tname, dt)::(find_dts dt_info tnames' tnames) | 
| 231 | else find_dts dt_info tnames' tnames); | |
| 232 | ||
| 32727 | 233 | fun prepare_induct ({descr, induct, ...}: info) rec_eqns =
 | 
| 25557 | 234 | let | 
| 235 | fun constrs_of (_, (_, _, cs)) = | |
| 236 | map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs; | |
| 32712 
ec5976f4d3d8
registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
 haftmann parents: 
31902diff
changeset | 237 | val params_of = these o AList.lookup (op =) (maps constrs_of rec_eqns); | 
| 25557 | 238 | in | 
| 32712 
ec5976f4d3d8
registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
 haftmann parents: 
31902diff
changeset | 239 | induct | 
| 33368 | 240 | |> Rule_Cases.rename_params (map params_of (maps (map #1 o #3 o #2) descr)) | 
| 241 | |> Rule_Cases.save induct | |
| 25557 | 242 | end; | 
| 243 | ||
| 244 | local | |
| 245 | ||
| 246 | fun gen_primrec_i note def alt_name eqns_atts thy = | |
| 247 | let | |
| 248 | val (eqns, atts) = split_list eqns_atts; | |
| 33963 
977b94b64905
renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33957diff
changeset | 249 | val dt_info = Datatype_Data.get_all thy; | 
| 25557 | 250 | val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ; | 
| 251 | val tnames = distinct (op =) (map (#1 o snd) rec_eqns); | |
| 252 | val dts = find_dts dt_info tnames tnames; | |
| 253 | val main_fns = | |
| 254 |       map (fn (tname, {index, ...}) =>
 | |
| 255 | (index, | |
| 256 | (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns)) | |
| 257 | dts; | |
| 258 |     val {descr, rec_names, rec_rewrites, ...} =
 | |
| 259 | if null dts then | |
| 260 |         primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
 | |
| 261 | else snd (hd dts); | |
| 262 | val (fnameTs, fnss) = | |
| 263 | fold_rev (process_fun thy descr rec_eqns) main_fns ([], []); | |
| 264 | val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []); | |
| 265 | val defs' = map (make_def thy fs) defs; | |
| 266 | val nameTs1 = map snd fnameTs; | |
| 267 | val nameTs2 = map fst rec_eqns; | |
| 33038 | 268 | val _ = if eq_set (op =) (nameTs1, nameTs2) then () | 
| 25557 | 269 |             else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^
 | 
| 270 | "\nare not mutually recursive"); | |
| 271 | val primrec_name = | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 272 | if alt_name = "" then (space_implode "_" (map (Long_Name.base_name o #1) defs)) else alt_name; | 
| 25557 | 273 | val (defs_thms', thy') = | 
| 274 | thy | |
| 275 | |> Sign.add_path primrec_name | |
| 276 | |> fold_map def (map (fn (name, t) => ((name, []), t)) defs'); | |
| 277 | val rewrites = (map mk_meta_eq rec_rewrites) @ map snd defs_thms'; | |
| 278 | val simps = map (fn (_, t) => Goal.prove_global thy' [] [] t | |
| 279 | (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; | |
| 280 | val (simps', thy'') = | |
| 281 | thy' | |
| 282 | |> fold_map note ((map fst eqns ~~ atts) ~~ map single simps); | |
| 283 | val simps'' = maps snd simps'; | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35756diff
changeset | 284 | val lhss = map (Logic.varify_global o fst o Logic.dest_equals o snd) defs'; | 
| 25557 | 285 | in | 
| 286 | thy'' | |
| 31902 | 287 |     |> note (("simps",
 | 
| 33056 
791a4655cae3
renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
 blanchet parents: 
33040diff
changeset | 288 | [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]), simps'') | 
| 25557 | 289 | |> snd | 
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
34952diff
changeset | 290 | |> Spec_Rules.add_global Spec_Rules.Equational (lhss, simps) | 
| 25557 | 291 |     |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns])
 | 
| 292 | |> snd | |
| 293 | |> Sign.parent_path | |
| 294 | |> pair simps'' | |
| 295 | end; | |
| 296 | ||
| 297 | fun gen_primrec note def alt_name eqns thy = | |
| 298 | let | |
| 299 | val ((names, strings), srcss) = apfst split_list (split_list eqns); | |
| 300 | val atts = map (map (Attrib.attribute thy)) srcss; | |
| 301 | val eqn_ts = map (fn s => Syntax.read_prop_global thy s | |
| 302 |       handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings;
 | |
| 303 | val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq))) | |
| 304 | handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts; | |
| 305 | val (_, eqn_ts') = unify_consts thy rec_ts eqn_ts | |
| 306 | in | |
| 307 | gen_primrec_i note def alt_name (names ~~ eqn_ts' ~~ atts) thy | |
| 308 | end; | |
| 309 | ||
| 310 | fun thy_note ((name, atts), thms) = | |
| 29579 | 311 | PureThy.add_thmss [((Binding.name name, thms), atts)] #-> (fn [thms] => pair (name, thms)); | 
| 25557 | 312 | fun thy_def false ((name, atts), t) = | 
| 29579 | 313 | PureThy.add_defs false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm)) | 
| 25557 | 314 | | thy_def true ((name, atts), t) = | 
| 29579 | 315 | PureThy.add_defs_unchecked false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm)); | 
| 25557 | 316 | |
| 317 | in | |
| 318 | ||
| 319 | val add_primrec = gen_primrec thy_note (thy_def false); | |
| 320 | val add_primrec_unchecked = gen_primrec thy_note (thy_def true); | |
| 321 | val add_primrec_i = gen_primrec_i thy_note (thy_def false); | |
| 322 | val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true); | |
| 323 | ||
| 324 | end; | |
| 325 | ||
| 326 | end; |