author | wenzelm |
Wed, 02 Nov 2005 14:47:00 +0100 | |
changeset 18063 | c4bffc47c11b |
parent 17985 | d5d576b72371 |
child 18358 | 0a733e11021a |
permissions | -rw-r--r-- |
6050 | 1 |
(* Title: ZF/Tools/primrec_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Stefan Berghofer and Norbert Voelker |
|
4 |
Copyright 1998 TU Muenchen |
|
5 |
ZF version by Lawrence C Paulson (Cambridge) |
|
6 |
||
7 |
Package for defining functions on datatypes by primitive recursion |
|
8 |
*) |
|
9 |
||
10 |
signature PRIMREC_PACKAGE = |
|
11 |
sig |
|
12183 | 12 |
val quiet_mode: bool ref |
15703 | 13 |
val add_primrec: ((bstring * string) * Attrib.src list) list -> theory -> theory * thm list |
12183 | 14 |
val add_primrec_i: ((bstring * term) * theory attribute list) list -> theory -> theory * thm list |
6050 | 15 |
end; |
16 |
||
17 |
structure PrimrecPackage : PRIMREC_PACKAGE = |
|
18 |
struct |
|
19 |
||
12183 | 20 |
(* messages *) |
21 |
||
22 |
val quiet_mode = ref false; |
|
23 |
fun message s = if ! quiet_mode then () else writeln s; |
|
24 |
||
25 |
||
6050 | 26 |
exception RecError of string; |
27 |
||
6141 | 28 |
(*Remove outer Trueprop and equality sign*) |
29 |
val dest_eqn = FOLogic.dest_eq o FOLogic.dest_Trueprop; |
|
6050 | 30 |
|
31 |
fun primrec_err s = error ("Primrec definition error:\n" ^ s); |
|
32 |
||
33 |
fun primrec_eq_err sign s eq = |
|
34 |
primrec_err (s ^ "\nin equation\n" ^ Sign.string_of_term sign eq); |
|
35 |
||
12183 | 36 |
|
6050 | 37 |
(* preprocessing of equations *) |
38 |
||
39 |
(*rec_fn_opt records equations already noted for this function*) |
|
12183 | 40 |
fun process_eqn thy (eq, rec_fn_opt) = |
6050 | 41 |
let |
12183 | 42 |
val (lhs, rhs) = |
43 |
if null (term_vars eq) then |
|
12203 | 44 |
dest_eqn eq handle TERM _ => raise RecError "not a proper equation" |
12183 | 45 |
else raise RecError "illegal schematic variable(s)"; |
6050 | 46 |
|
47 |
val (recfun, args) = strip_comb lhs; |
|
12203 | 48 |
val (fname, ftype) = dest_Const recfun handle TERM _ => |
6050 | 49 |
raise RecError "function is not declared as constant in theory"; |
50 |
||
51 |
val (ls_frees, rest) = take_prefix is_Free args; |
|
52 |
val (middle, rs_frees) = take_suffix is_Free rest; |
|
53 |
||
12183 | 54 |
val (constr, cargs_frees) = |
6050 | 55 |
if null middle then raise RecError "constructor missing" |
56 |
else strip_comb (hd middle); |
|
57 |
val (cname, _) = dest_Const constr |
|
12203 | 58 |
handle TERM _ => raise RecError "ill-formed constructor"; |
17412 | 59 |
val con_info = the (Symtab.lookup (ConstructorsData.get thy) cname) |
15531 | 60 |
handle Option => |
6050 | 61 |
raise RecError "cannot determine datatype associated with function" |
62 |
||
12183 | 63 |
val (ls, cargs, rs) = (map dest_Free ls_frees, |
64 |
map dest_Free cargs_frees, |
|
65 |
map dest_Free rs_frees) |
|
12203 | 66 |
handle TERM _ => raise RecError "illegal argument in pattern"; |
6050 | 67 |
val lfrees = ls @ rs @ cargs; |
68 |
||
69 |
(*Constructor, frees to left of pattern, pattern variables, |
|
70 |
frees to right of pattern, rhs of equation, full original equation. *) |
|
71 |
val new_eqn = (cname, (rhs, cargs, eq)) |
|
72 |
||
73 |
in |
|
12183 | 74 |
if not (null (duplicates lfrees)) then |
75 |
raise RecError "repeated variable name in pattern" |
|
6050 | 76 |
else if not ((map dest_Free (term_frees rhs)) subset lfrees) then |
77 |
raise RecError "extra variables on rhs" |
|
12183 | 78 |
else if length middle > 1 then |
6050 | 79 |
raise RecError "more than one non-variable in pattern" |
80 |
else case rec_fn_opt of |
|
15531 | 81 |
NONE => SOME (fname, ftype, ls, rs, con_info, [new_eqn]) |
82 |
| SOME (fname', _, ls', rs', con_info': constructor_info, eqns) => |
|
17314 | 83 |
if AList.defined (op =) eqns cname then |
12183 | 84 |
raise RecError "constructor already occurred as pattern" |
85 |
else if (ls <> ls') orelse (rs <> rs') then |
|
86 |
raise RecError "non-recursive arguments are inconsistent" |
|
87 |
else if #big_rec_name con_info <> #big_rec_name con_info' then |
|
88 |
raise RecError ("Mixed datatypes for function " ^ fname) |
|
89 |
else if fname <> fname' then |
|
90 |
raise RecError ("inconsistent functions for datatype " ^ |
|
91 |
#big_rec_name con_info) |
|
15531 | 92 |
else SOME (fname, ftype, ls, rs, con_info, new_eqn::eqns) |
6050 | 93 |
end |
94 |
handle RecError s => primrec_eq_err (sign_of thy) s eq; |
|
95 |
||
96 |
||
97 |
(*Instantiates a recursor equation with constructor arguments*) |
|
12183 | 98 |
fun inst_recursor ((_ $ constr, rhs), cargs') = |
6050 | 99 |
subst_atomic (#2 (strip_comb constr) ~~ map Free cargs') rhs; |
100 |
||
101 |
||
102 |
(*Convert a list of recursion equations into a recursor call*) |
|
103 |
fun process_fun thy (fname, ftype, ls, rs, con_info: constructor_info, eqns) = |
|
104 |
let |
|
105 |
val fconst = Const(fname, ftype) |
|
106 |
val fabs = list_comb (fconst, map Free ls @ [Bound 0] @ map Free rs) |
|
107 |
and {big_rec_name, constructors, rec_rewrites, ...} = con_info |
|
108 |
||
109 |
(*Replace X_rec(args,t) by fname(ls,t,rs) *) |
|
110 |
fun use_fabs (_ $ t) = subst_bound (t, fabs) |
|
111 |
| use_fabs t = t |
|
112 |
||
113 |
val cnames = map (#1 o dest_Const) constructors |
|
6141 | 114 |
and recursor_pairs = map (dest_eqn o concl_of) rec_rewrites |
6050 | 115 |
|
116 |
fun absterm (Free(a,T), body) = absfree (a,T,body) |
|
9179
44eabc57ed46
no longer depends upon a prior "open Ind_Syntax" from elsewhere
paulson
parents:
8819
diff
changeset
|
117 |
| absterm (t,body) = Abs("rec", Ind_Syntax.iT, abstract_over (t, body)) |
6050 | 118 |
|
119 |
(*Translate rec equations into function arguments suitable for recursor. |
|
120 |
Missing cases are replaced by 0 and all cases are put into order.*) |
|
121 |
fun add_case ((cname, recursor_pair), cases) = |
|
12183 | 122 |
let val (rhs, recursor_rhs, eq) = |
17314 | 123 |
case AList.lookup (op =) eqns cname of |
15531 | 124 |
NONE => (warning ("no equation for constructor " ^ cname ^ |
12183 | 125 |
"\nin definition of function " ^ fname); |
126 |
(Const ("0", Ind_Syntax.iT), |
|
127 |
#2 recursor_pair, Const ("0", Ind_Syntax.iT))) |
|
15531 | 128 |
| SOME (rhs, cargs', eq) => |
12183 | 129 |
(rhs, inst_recursor (recursor_pair, cargs'), eq) |
130 |
val allowed_terms = map use_fabs (#2 (strip_comb recursor_rhs)) |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
131 |
val abs = foldr absterm rhs allowed_terms |
12183 | 132 |
in |
6050 | 133 |
if !Ind_Syntax.trace then |
12183 | 134 |
writeln ("recursor_rhs = " ^ |
135 |
Sign.string_of_term (sign_of thy) recursor_rhs ^ |
|
136 |
"\nabs = " ^ Sign.string_of_term (sign_of thy) abs) |
|
6050 | 137 |
else(); |
12183 | 138 |
if Logic.occs (fconst, abs) then |
139 |
primrec_eq_err (sign_of thy) |
|
140 |
("illegal recursive occurrences of " ^ fname) |
|
141 |
eq |
|
142 |
else abs :: cases |
|
6050 | 143 |
end |
144 |
||
145 |
val recursor = head_of (#1 (hd recursor_pairs)) |
|
146 |
||
147 |
(** make definition **) |
|
148 |
||
149 |
(*the recursive argument*) |
|
150 |
val rec_arg = Free (variant (map #1 (ls@rs)) (Sign.base_name big_rec_name), |
|
12183 | 151 |
Ind_Syntax.iT) |
6050 | 152 |
|
153 |
val def_tm = Logic.mk_equals |
|
12183 | 154 |
(subst_bound (rec_arg, fabs), |
155 |
list_comb (recursor, |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
156 |
foldr add_case [] (cnames ~~ recursor_pairs)) |
12183 | 157 |
$ rec_arg) |
6050 | 158 |
|
159 |
in |
|
6065 | 160 |
if !Ind_Syntax.trace then |
12183 | 161 |
writeln ("primrec def:\n" ^ |
162 |
Sign.string_of_term (sign_of thy) def_tm) |
|
6065 | 163 |
else(); |
6050 | 164 |
(Sign.base_name fname ^ "_" ^ Sign.base_name big_rec_name ^ "_def", |
165 |
def_tm) |
|
166 |
end; |
|
167 |
||
168 |
||
169 |
(* prepare functions needed for definitions *) |
|
170 |
||
12183 | 171 |
fun add_primrec_i args thy = |
6050 | 172 |
let |
12183 | 173 |
val ((eqn_names, eqn_terms), eqn_atts) = apfst split_list (split_list args); |
15531 | 174 |
val SOME (fname, ftype, ls, rs, con_info, eqns) = |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
175 |
foldr (process_eqn thy) NONE eqn_terms; |
12183 | 176 |
val def = process_fun thy (fname, ftype, ls, rs, con_info, eqns); |
177 |
||
178 |
val (thy1, [def_thm]) = thy |
|
12188 | 179 |
|> Theory.add_path (Sign.base_name fname) |
12183 | 180 |
|> (PureThy.add_defs_i false o map Thm.no_attributes) [def]; |
181 |
||
8438 | 182 |
val rewrites = def_thm :: map mk_meta_eq (#rec_rewrites con_info) |
12183 | 183 |
val eqn_thms = |
17985 | 184 |
(message ("Proving equations for primrec function " ^ fname); |
185 |
eqn_terms |> map (fn t => |
|
186 |
standard (Goal.prove thy1 [] [] (Ind_Syntax.traceIt "next primrec equation = " thy1 t) |
|
187 |
(fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])))); |
|
12183 | 188 |
|
189 |
val (thy2, eqn_thms') = thy1 |> PureThy.add_thms ((eqn_names ~~ eqn_thms) ~~ eqn_atts); |
|
190 |
val thy3 = thy2 |
|
191 |
|> (#1 o PureThy.add_thmss [(("simps", eqn_thms'), [Simplifier.simp_add_global])]) |
|
6050 | 192 |
|> Theory.parent_path; |
12183 | 193 |
in (thy3, eqn_thms') end; |
194 |
||
195 |
fun add_primrec args thy = |
|
196 |
add_primrec_i (map (fn ((name, s), srcs) => |
|
197 |
((name, Sign.simple_read_term (sign_of thy) propT s), map (Attrib.global_attribute thy) srcs)) |
|
198 |
args) thy; |
|
199 |
||
200 |
||
201 |
(* outer syntax *) |
|
6050 | 202 |
|
17057 | 203 |
local structure P = OuterParse and K = OuterKeyword in |
12183 | 204 |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12203
diff
changeset
|
205 |
val primrec_decl = Scan.repeat1 (P.opt_thm_name ":" -- P.prop); |
12183 | 206 |
|
207 |
val primrecP = |
|
208 |
OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl |
|
209 |
(primrec_decl >> (Toplevel.theory o (#1 oo (add_primrec o map P.triple_swap)))); |
|
210 |
||
211 |
val _ = OuterSyntax.add_parsers [primrecP]; |
|
6050 | 212 |
|
213 |
end; |
|
12183 | 214 |
|
215 |
end; |
|
15705 | 216 |
|
217 |