| author | blanchet | 
| Fri, 30 Aug 2013 11:37:22 +0200 | |
| changeset 53304 | cfef783090eb | 
| parent 51930 | 52fd62618631 | 
| child 54742 | 7a86358a3c0b | 
| permissions | -rw-r--r-- | 
| 6050 | 1  | 
(* Title: ZF/Tools/primrec_package.ML  | 
| 29266 | 2  | 
Author: Norbert Voelker, FernUni Hagen  | 
3  | 
Author: Stefan Berghofer, TU Muenchen  | 
|
4  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
| 6050 | 5  | 
|
| 29266 | 6  | 
Package for defining functions on datatypes by primitive recursion.  | 
| 6050 | 7  | 
*)  | 
8  | 
||
9  | 
signature PRIMREC_PACKAGE =  | 
|
10  | 
sig  | 
|
| 29579 | 11  | 
val add_primrec: ((binding * string) * Attrib.src list) list -> theory -> theory * thm list  | 
12  | 
val add_primrec_i: ((binding * term) * attribute list) list -> theory -> theory * thm list  | 
|
| 6050 | 13  | 
end;  | 
14  | 
||
15  | 
structure PrimrecPackage : PRIMREC_PACKAGE =  | 
|
16  | 
struct  | 
|
17  | 
||
18  | 
exception RecError of string;  | 
|
19  | 
||
| 6141 | 20  | 
(*Remove outer Trueprop and equality sign*)  | 
21  | 
val dest_eqn = FOLogic.dest_eq o FOLogic.dest_Trueprop;  | 
|
| 6050 | 22  | 
|
23  | 
fun primrec_err s = error ("Primrec definition error:\n" ^ s);
 | 
|
24  | 
||
25  | 
fun primrec_eq_err sign s eq =  | 
|
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26478 
diff
changeset
 | 
26  | 
primrec_err (s ^ "\nin equation\n" ^ Syntax.string_of_term_global sign eq);  | 
| 6050 | 27  | 
|
| 12183 | 28  | 
|
| 6050 | 29  | 
(* preprocessing of equations *)  | 
30  | 
||
31  | 
(*rec_fn_opt records equations already noted for this function*)  | 
|
| 12183 | 32  | 
fun process_eqn thy (eq, rec_fn_opt) =  | 
| 6050 | 33  | 
let  | 
| 12183 | 34  | 
val (lhs, rhs) =  | 
| 29266 | 35  | 
if null (Term.add_vars eq []) then  | 
| 12203 | 36  | 
dest_eqn eq handle TERM _ => raise RecError "not a proper equation"  | 
| 12183 | 37  | 
else raise RecError "illegal schematic variable(s)";  | 
| 6050 | 38  | 
|
39  | 
val (recfun, args) = strip_comb lhs;  | 
|
| 12203 | 40  | 
val (fname, ftype) = dest_Const recfun handle TERM _ =>  | 
| 6050 | 41  | 
raise RecError "function is not declared as constant in theory";  | 
42  | 
||
43  | 
val (ls_frees, rest) = take_prefix is_Free args;  | 
|
44  | 
val (middle, rs_frees) = take_suffix is_Free rest;  | 
|
45  | 
||
| 12183 | 46  | 
val (constr, cargs_frees) =  | 
| 6050 | 47  | 
if null middle then raise RecError "constructor missing"  | 
48  | 
else strip_comb (hd middle);  | 
|
49  | 
val (cname, _) = dest_Const constr  | 
|
| 12203 | 50  | 
handle TERM _ => raise RecError "ill-formed constructor";  | 
| 17412 | 51  | 
val con_info = the (Symtab.lookup (ConstructorsData.get thy) cname)  | 
| 
51930
 
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
52  | 
handle Option.Option =>  | 
| 6050 | 53  | 
raise RecError "cannot determine datatype associated with function"  | 
54  | 
||
| 12183 | 55  | 
val (ls, cargs, rs) = (map dest_Free ls_frees,  | 
56  | 
map dest_Free cargs_frees,  | 
|
57  | 
map dest_Free rs_frees)  | 
|
| 12203 | 58  | 
handle TERM _ => raise RecError "illegal argument in pattern";  | 
| 6050 | 59  | 
val lfrees = ls @ rs @ cargs;  | 
60  | 
||
61  | 
(*Constructor, frees to left of pattern, pattern variables,  | 
|
62  | 
frees to right of pattern, rhs of equation, full original equation. *)  | 
|
63  | 
val new_eqn = (cname, (rhs, cargs, eq))  | 
|
64  | 
||
65  | 
in  | 
|
| 18973 | 66  | 
if has_duplicates (op =) lfrees then  | 
| 12183 | 67  | 
raise RecError "repeated variable name in pattern"  | 
| 33038 | 68  | 
else if not (subset (op =) (Term.add_frees rhs [], lfrees)) then  | 
| 6050 | 69  | 
raise RecError "extra variables on rhs"  | 
| 12183 | 70  | 
else if length middle > 1 then  | 
| 6050 | 71  | 
raise RecError "more than one non-variable in pattern"  | 
72  | 
else case rec_fn_opt of  | 
|
| 15531 | 73  | 
NONE => SOME (fname, ftype, ls, rs, con_info, [new_eqn])  | 
74  | 
| SOME (fname', _, ls', rs', con_info': constructor_info, eqns) =>  | 
|
| 17314 | 75  | 
if AList.defined (op =) eqns cname then  | 
| 12183 | 76  | 
raise RecError "constructor already occurred as pattern"  | 
77  | 
else if (ls <> ls') orelse (rs <> rs') then  | 
|
78  | 
raise RecError "non-recursive arguments are inconsistent"  | 
|
79  | 
else if #big_rec_name con_info <> #big_rec_name con_info' then  | 
|
80  | 
             raise RecError ("Mixed datatypes for function " ^ fname)
 | 
|
81  | 
else if fname <> fname' then  | 
|
82  | 
             raise RecError ("inconsistent functions for datatype " ^
 | 
|
83  | 
#big_rec_name con_info)  | 
|
| 15531 | 84  | 
else SOME (fname, ftype, ls, rs, con_info, new_eqn::eqns)  | 
| 6050 | 85  | 
end  | 
| 20342 | 86  | 
handle RecError s => primrec_eq_err thy s eq;  | 
| 6050 | 87  | 
|
88  | 
||
89  | 
(*Instantiates a recursor equation with constructor arguments*)  | 
|
| 12183 | 90  | 
fun inst_recursor ((_ $ constr, rhs), cargs') =  | 
| 6050 | 91  | 
subst_atomic (#2 (strip_comb constr) ~~ map Free cargs') rhs;  | 
92  | 
||
93  | 
||
94  | 
(*Convert a list of recursion equations into a recursor call*)  | 
|
95  | 
fun process_fun thy (fname, ftype, ls, rs, con_info: constructor_info, eqns) =  | 
|
96  | 
let  | 
|
97  | 
val fconst = Const(fname, ftype)  | 
|
98  | 
val fabs = list_comb (fconst, map Free ls @ [Bound 0] @ map Free rs)  | 
|
99  | 
    and {big_rec_name, constructors, rec_rewrites, ...} = con_info
 | 
|
100  | 
||
101  | 
(*Replace X_rec(args,t) by fname(ls,t,rs) *)  | 
|
102  | 
fun use_fabs (_ $ t) = subst_bound (t, fabs)  | 
|
103  | 
| use_fabs t = t  | 
|
104  | 
||
105  | 
val cnames = map (#1 o dest_Const) constructors  | 
|
| 6141 | 106  | 
and recursor_pairs = map (dest_eqn o concl_of) rec_rewrites  | 
| 6050 | 107  | 
|
| 44241 | 108  | 
fun absterm (Free x, body) = absfree x body  | 
109  | 
      | absterm (t, body) = Abs("rec", Ind_Syntax.iT, abstract_over (t, body))
 | 
|
| 6050 | 110  | 
|
111  | 
(*Translate rec equations into function arguments suitable for recursor.  | 
|
112  | 
Missing cases are replaced by 0 and all cases are put into order.*)  | 
|
113  | 
fun add_case ((cname, recursor_pair), cases) =  | 
|
| 12183 | 114  | 
let val (rhs, recursor_rhs, eq) =  | 
| 17314 | 115  | 
case AList.lookup (op =) eqns cname of  | 
| 15531 | 116  | 
                NONE => (warning ("no equation for constructor " ^ cname ^
 | 
| 12183 | 117  | 
"\nin definition of function " ^ fname);  | 
| 41310 | 118  | 
                         (Const (@{const_name zero}, Ind_Syntax.iT),
 | 
119  | 
                          #2 recursor_pair, Const (@{const_name zero}, Ind_Syntax.iT)))
 | 
|
| 15531 | 120  | 
| SOME (rhs, cargs', eq) =>  | 
| 12183 | 121  | 
(rhs, inst_recursor (recursor_pair, cargs'), eq)  | 
122  | 
val allowed_terms = map use_fabs (#2 (strip_comb recursor_rhs))  | 
|
| 30190 | 123  | 
val abs = List.foldr absterm rhs allowed_terms  | 
| 12183 | 124  | 
in  | 
| 6050 | 125  | 
if !Ind_Syntax.trace then  | 
| 12183 | 126  | 
              writeln ("recursor_rhs = " ^
 | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26478 
diff
changeset
 | 
127  | 
Syntax.string_of_term_global thy recursor_rhs ^  | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26478 
diff
changeset
 | 
128  | 
"\nabs = " ^ Syntax.string_of_term_global thy abs)  | 
| 6050 | 129  | 
else();  | 
| 12183 | 130  | 
if Logic.occs (fconst, abs) then  | 
| 20342 | 131  | 
primrec_eq_err thy  | 
| 12183 | 132  | 
                   ("illegal recursive occurrences of " ^ fname)
 | 
133  | 
eq  | 
|
134  | 
else abs :: cases  | 
|
| 6050 | 135  | 
end  | 
136  | 
||
137  | 
val recursor = head_of (#1 (hd recursor_pairs))  | 
|
138  | 
||
139  | 
(** make definition **)  | 
|
140  | 
||
141  | 
(*the recursive argument*)  | 
|
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
41310 
diff
changeset
 | 
142  | 
val rec_arg =  | 
| 
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
41310 
diff
changeset
 | 
143  | 
Free (singleton (Name.variant_list (map #1 (ls@rs))) (Long_Name.base_name big_rec_name),  | 
| 
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
41310 
diff
changeset
 | 
144  | 
Ind_Syntax.iT)  | 
| 6050 | 145  | 
|
146  | 
val def_tm = Logic.mk_equals  | 
|
| 12183 | 147  | 
(subst_bound (rec_arg, fabs),  | 
148  | 
list_comb (recursor,  | 
|
| 30190 | 149  | 
List.foldr add_case [] (cnames ~~ recursor_pairs))  | 
| 12183 | 150  | 
$ rec_arg)  | 
| 6050 | 151  | 
|
152  | 
in  | 
|
| 6065 | 153  | 
if !Ind_Syntax.trace then  | 
| 12183 | 154  | 
            writeln ("primrec def:\n" ^
 | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26478 
diff
changeset
 | 
155  | 
Syntax.string_of_term_global thy def_tm)  | 
| 6065 | 156  | 
else();  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
157  | 
(Long_Name.base_name fname ^ "_" ^ Long_Name.base_name big_rec_name ^ "_def",  | 
| 6050 | 158  | 
def_tm)  | 
159  | 
end;  | 
|
160  | 
||
161  | 
||
162  | 
(* prepare functions needed for definitions *)  | 
|
163  | 
||
| 12183 | 164  | 
fun add_primrec_i args thy =  | 
| 6050 | 165  | 
let  | 
| 12183 | 166  | 
val ((eqn_names, eqn_terms), eqn_atts) = apfst split_list (split_list args);  | 
| 15531 | 167  | 
val SOME (fname, ftype, ls, rs, con_info, eqns) =  | 
| 30190 | 168  | 
List.foldr (process_eqn thy) NONE eqn_terms;  | 
| 12183 | 169  | 
val def = process_fun thy (fname, ftype, ls, rs, con_info, eqns);  | 
170  | 
||
| 18358 | 171  | 
val ([def_thm], thy1) = thy  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
172  | 
|> Sign.add_path (Long_Name.base_name fname)  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38514 
diff
changeset
 | 
173  | 
|> Global_Theory.add_defs false [Thm.no_attributes (apfst Binding.name def)];  | 
| 12183 | 174  | 
|
| 8438 | 175  | 
val rewrites = def_thm :: map mk_meta_eq (#rec_rewrites con_info)  | 
| 12183 | 176  | 
val eqn_thms =  | 
| 17985 | 177  | 
eqn_terms |> map (fn t =>  | 
| 20046 | 178  | 
Goal.prove_global thy1 [] [] (Ind_Syntax.traceIt "next primrec equation = " thy1 t)  | 
| 35409 | 179  | 
          (fn _ => EVERY [rewrite_goals_tac rewrites, rtac @{thm refl} 1]));
 | 
| 12183 | 180  | 
|
| 18377 | 181  | 
val (eqn_thms', thy2) =  | 
182  | 
thy1  | 
|
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38514 
diff
changeset
 | 
183  | 
|> Global_Theory.add_thms ((eqn_names ~~ eqn_thms) ~~ eqn_atts);  | 
| 18377 | 184  | 
val (_, thy3) =  | 
185  | 
thy2  | 
|
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38514 
diff
changeset
 | 
186  | 
|> Global_Theory.add_thmss [((Binding.name "simps", eqn_thms'), [Simplifier.simp_add])]  | 
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24707 
diff
changeset
 | 
187  | 
||> Sign.parent_path;  | 
| 12183 | 188  | 
in (thy3, eqn_thms') end;  | 
189  | 
||
190  | 
fun add_primrec args thy =  | 
|
191  | 
add_primrec_i (map (fn ((name, s), srcs) =>  | 
|
| 47815 | 192  | 
((name, Syntax.read_prop_global thy s), map (Attrib.attribute_cmd_global thy) srcs))  | 
| 12183 | 193  | 
args) thy;  | 
194  | 
||
195  | 
||
196  | 
(* outer syntax *)  | 
|
| 6050 | 197  | 
|
| 24867 | 198  | 
val _ =  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
44241 
diff
changeset
 | 
199  | 
  Outer_Syntax.command @{command_spec "primrec"} "define primitive recursive functions on datatypes"
 | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
200  | 
(Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)  | 
| 
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
201  | 
>> (Toplevel.theory o (#1 oo (add_primrec o map Parse.triple_swap))));  | 
| 6050 | 202  | 
|
203  | 
end;  | 
|
| 12183 | 204  |