| author | wenzelm | 
| Thu, 27 Oct 2005 13:54:42 +0200 | |
| changeset 17994 | 6a1a49cba5b3 | 
| parent 17985 | d5d576b72371 | 
| child 18358 | 0a733e11021a | 
| permissions | -rw-r--r-- | 
| 16226 | 1  | 
(* Title: HOLCF/fixrec_package.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Amber Telfer and Brian Huffman  | 
|
4  | 
||
5  | 
Recursive function definition package for HOLCF.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature FIXREC_PACKAGE =  | 
|
9  | 
sig  | 
|
| 16552 | 10  | 
val add_fixrec: bool -> ((string * Attrib.src list) * string) list list  | 
11  | 
-> theory -> theory  | 
|
12  | 
val add_fixrec_i: bool -> ((string * theory attribute list) * term) list list  | 
|
13  | 
-> theory -> theory  | 
|
14  | 
val add_fixpat: (string * Attrib.src list) * string list  | 
|
15  | 
-> theory -> theory  | 
|
16  | 
val add_fixpat_i: (string * theory attribute list) * term list  | 
|
17  | 
-> theory -> theory  | 
|
| 16226 | 18  | 
end;  | 
19  | 
||
20  | 
structure FixrecPackage: FIXREC_PACKAGE =  | 
|
21  | 
struct  | 
|
22  | 
||
| 16552 | 23  | 
fun fixrec_err s = error ("fixrec definition error:\n" ^ s);
 | 
| 16628 | 24  | 
fun fixrec_eq_err thy s eq =  | 
25  | 
fixrec_err (s ^ "\nin\n" ^ quote (Sign.string_of_term thy eq));  | 
|
| 16552 | 26  | 
|
| 16226 | 27  | 
(* ->> is taken from holcf_logic.ML *)  | 
28  | 
(* TODO: fix dependencies so we can import HOLCFLogic here *)  | 
|
29  | 
infixr 6 ->>;  | 
|
30  | 
fun S ->> T = Type ("Cfun.->",[S,T]);
 | 
|
31  | 
||
32  | 
(* extern_name is taken from domain/library.ML *)  | 
|
33  | 
fun extern_name con = case Symbol.explode con of  | 
|
34  | 
		   ("o"::"p"::" "::rest) => implode rest
 | 
|
35  | 
| _ => con;  | 
|
36  | 
||
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
37  | 
val mk_trp = HOLogic.mk_Trueprop;  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
38  | 
|
| 16226 | 39  | 
(* splits a cterm into the right and lefthand sides of equality *)  | 
| 16552 | 40  | 
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t);  | 
41  | 
(*  | 
|
| 16226 | 42  | 
fun dest_eqs (Const ("==", _)$lhs$rhs) = (lhs, rhs)
 | 
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
43  | 
  | dest_eqs (Const ("Trueprop", _)$(Const ("op =", _)$lhs$rhs)) = (lhs,rhs)
 | 
| 16552 | 44  | 
| dest_eqs t = fixrec_err (Sign.string_of_term (sign_of (the_context())) t);  | 
45  | 
*)  | 
|
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
46  | 
(* similar to Thm.head_of, but for continuous application *)  | 
| 
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
47  | 
fun chead_of (Const("Cfun.Rep_CFun",_)$f$t) = chead_of f
 | 
| 
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
48  | 
| chead_of u = u;  | 
| 
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
49  | 
|
| 16387 | 50  | 
(* these are helpful functions copied from HOLCF/domain/library.ML *)  | 
51  | 
fun %: s = Free(s,dummyT);  | 
|
52  | 
fun %%: s = Const(s,dummyT);  | 
|
53  | 
infix 0 ==; fun S == T = %%:"==" $ S $ T;  | 
|
54  | 
infix 1 ===; fun S === T = %%:"op =" $ S $ T;  | 
|
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
55  | 
infix 9 ` ; fun f ` x = %%:"Cfun.Rep_CFun" $ f $ x;  | 
| 16387 | 56  | 
|
57  | 
(* infers the type of a term *)  | 
|
| 16402 | 58  | 
(* similar to Theory.inferT_axm, but allows any type, not just propT *)  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
59  | 
fun infer sg t =  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
60  | 
fst (Sign.infer_types (Sign.pp sg) sg (K NONE) (K NONE) [] true ([t],dummyT));  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
61  | 
|
| 16402 | 62  | 
(* Similar to Term.lambda, but also allows abstraction over constants *)  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
63  | 
fun lambda' (v as Free (x, T)) t = Abs (x, T, abstract_over (v, t))  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
64  | 
| lambda' (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
65  | 
| lambda' (v as Const (x, T)) t = Abs (Sign.base_name x, T, abstract_over (v, t))  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
66  | 
  | lambda' v t = raise TERM ("lambda'", [v, t]);
 | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
67  | 
|
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
68  | 
(* builds the expression (LAM v. rhs) *)  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
69  | 
fun big_lambda v rhs = %%:"Cfun.Abs_CFun"$(lambda' v rhs);  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
70  | 
|
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
71  | 
(* builds the expression (LAM v1 v2 .. vn. rhs) *)  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
72  | 
fun big_lambdas [] rhs = rhs  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
73  | 
| big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
74  | 
|
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
75  | 
(* builds the expression (LAM <v1,v2,..,vn>. rhs) *)  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
76  | 
fun lambda_ctuple [] rhs = big_lambda (%:"unit") rhs  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
77  | 
| lambda_ctuple (v::[]) rhs = big_lambda v rhs  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
78  | 
| lambda_ctuple (v::vs) rhs =  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
79  | 
%%:"Cprod.csplit"`(big_lambda v (lambda_ctuple vs rhs));  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
80  | 
|
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
81  | 
(* builds the expression <v1,v2,..,vn> *)  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
82  | 
fun mk_ctuple [] = %%:"UU"  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
83  | 
| mk_ctuple (t::[]) = t  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
84  | 
| mk_ctuple (t::ts) = %%:"Cprod.cpair"`t`(mk_ctuple ts);  | 
| 16387 | 85  | 
|
86  | 
(*************************************************************************)  | 
|
| 16402 | 87  | 
(************* fixed-point definitions and unfolding theorems ************)  | 
| 16387 | 88  | 
(*************************************************************************)  | 
89  | 
||
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
90  | 
fun add_fixdefs eqs thy =  | 
| 16226 | 91  | 
let  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
92  | 
val (lhss,rhss) = ListPair.unzip (map dest_eqs eqs);  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
93  | 
val fixpoint = %%:"Fix.fix"`lambda_ctuple lhss (mk_ctuple rhss);  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
94  | 
|
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
95  | 
fun one_def (l as Const(n,T)) r =  | 
| 16552 | 96  | 
let val b = Sign.base_name n in (b, (b^"_def", l == r)) end  | 
97  | 
| one_def _ _ = fixrec_err "fixdefs: lhs not of correct form";  | 
|
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
98  | 
fun defs [] _ = []  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
99  | 
| defs (l::[]) r = [one_def l r]  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
100  | 
| defs (l::ls) r = one_def l (%%:"Cprod.cfst"`r) :: defs ls (%%:"Cprod.csnd"`r);  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
101  | 
val (names, pre_fixdefs) = ListPair.unzip (defs lhss fixpoint);  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
102  | 
|
| 16628 | 103  | 
val fixdefs = map (inferT_axm thy) pre_fixdefs;  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
104  | 
val (thy', fixdef_thms) =  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
105  | 
PureThy.add_defs_i false (map Thm.no_attributes fixdefs) thy;  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
106  | 
val ctuple_fixdef_thm = foldr1 (fn (x,y) => cpair_equalI OF [x,y]) fixdef_thms;  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
107  | 
|
| 17985 | 108  | 
val ctuple_unfold = infer thy' (mk_trp (mk_ctuple lhss === mk_ctuple rhss));  | 
109  | 
val ctuple_unfold_thm = standard (Goal.prove thy' [] [] ctuple_unfold  | 
|
110  | 
(fn _ => EVERY [rtac (ctuple_fixdef_thm RS fix_eq2 RS trans) 1,  | 
|
111  | 
simp_tac (simpset_of thy') 1]));  | 
|
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
112  | 
val ctuple_induct_thm =  | 
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
113  | 
(space_implode "_" names ^ "_induct", ctuple_fixdef_thm RS def_fix_ind);  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
114  | 
|
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
115  | 
fun unfolds [] thm = []  | 
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
116  | 
| unfolds (n::[]) thm = [(n^"_unfold", thm)]  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
117  | 
| unfolds (n::ns) thm = let  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
118  | 
val thmL = thm RS cpair_eqD1;  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
119  | 
val thmR = thm RS cpair_eqD2;  | 
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
120  | 
in (n^"_unfold", thmL) :: unfolds ns thmR end;  | 
| 
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
121  | 
val unfold_thms = unfolds names ctuple_unfold_thm;  | 
| 
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
122  | 
val thms = ctuple_induct_thm :: unfold_thms;  | 
| 
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
123  | 
val (thy'', _) = PureThy.add_thms (map Thm.no_attributes thms) thy';  | 
| 16226 | 124  | 
in  | 
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
125  | 
(thy'', names, fixdef_thms, map snd unfold_thms)  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
126  | 
end;  | 
| 16226 | 127  | 
|
| 16387 | 128  | 
(*************************************************************************)  | 
129  | 
(*********** monadic notation and pattern matching compilation ***********)  | 
|
130  | 
(*************************************************************************)  | 
|
| 16226 | 131  | 
|
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
132  | 
fun add_names (Const(a,_), bs) = Sign.base_name a ins_string bs  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
133  | 
| add_names (Free(a,_) , bs) = a ins_string bs  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
134  | 
| add_names (f $ u , bs) = add_names (f, add_names(u, bs))  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
135  | 
| add_names (Abs(a,_,t), bs) = add_names (t, a ins_string bs)  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
136  | 
| add_names (_ , bs) = bs;  | 
| 16226 | 137  | 
|
138  | 
fun add_terms ts xs = foldr add_names xs ts;  | 
|
139  | 
||
140  | 
(* builds a monadic term for matching a constructor pattern *)  | 
|
| 16387 | 141  | 
fun pre_build pat rhs vs taken =  | 
142  | 
case pat of  | 
|
143  | 
    Const("Cfun.Rep_CFun",_)$f$(v as Free(n,T)) =>
 | 
|
| 16226 | 144  | 
pre_build f rhs (v::vs) taken  | 
| 16387 | 145  | 
  | Const("Cfun.Rep_CFun",_)$f$x =>
 | 
| 16226 | 146  | 
let val (rhs', v, taken') = pre_build x rhs [] taken;  | 
| 16387 | 147  | 
in pre_build f rhs' (v::vs) taken' end  | 
148  | 
| Const(c,T) =>  | 
|
| 16226 | 149  | 
let  | 
150  | 
val n = variant taken "v";  | 
|
| 16387 | 151  | 
        fun result_type (Type("Cfun.->",[_,T])) (x::xs) = result_type T xs
 | 
| 16226 | 152  | 
| result_type T _ = T;  | 
153  | 
val v = Free(n, result_type T vs);  | 
|
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
154  | 
val m = "match_"^(extern_name(Sign.base_name c));  | 
| 
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
155  | 
val k = lambda_ctuple vs rhs;  | 
| 16226 | 156  | 
in  | 
| 16387 | 157  | 
(%%:"Fixrec.bind"`(%%:m`v)`k, v, n::taken)  | 
| 16552 | 158  | 
end  | 
159  | 
  | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n)
 | 
|
160  | 
| _ => fixrec_err "pre_build: invalid pattern";  | 
|
| 16226 | 161  | 
|
162  | 
(* builds a monadic term for matching a function definition pattern *)  | 
|
163  | 
(* returns (name, arity, matcher) *)  | 
|
| 16387 | 164  | 
fun building pat rhs vs taken =  | 
165  | 
case pat of  | 
|
166  | 
    Const("Cfun.Rep_CFun", _)$f$(v as Free(n,T)) =>
 | 
|
| 16226 | 167  | 
building f rhs (v::vs) taken  | 
| 16387 | 168  | 
  | Const("Cfun.Rep_CFun", _)$f$x =>
 | 
169  | 
let val (rhs', v, taken') = pre_build x rhs [] taken;  | 
|
170  | 
in building f rhs' (v::vs) taken' end  | 
|
| 16552 | 171  | 
| Const(name,_) => (name, length vs, big_lambdas vs rhs)  | 
172  | 
| _ => fixrec_err "function is not declared as constant in theory";  | 
|
| 16226 | 173  | 
|
| 16387 | 174  | 
fun match_eq eq =  | 
| 16552 | 175  | 
let val (lhs,rhs) = dest_eqs eq;  | 
176  | 
in building lhs (%%:"Fixrec.return"`rhs) [] (add_terms [eq] []) end;  | 
|
| 16226 | 177  | 
|
178  | 
(* returns the sum (using +++) of the terms in ms *)  | 
|
179  | 
(* also applies "run" to the result! *)  | 
|
180  | 
fun fatbar arity ms =  | 
|
181  | 
let  | 
|
182  | 
fun unLAM 0 t = t  | 
|
183  | 
| unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t  | 
|
| 16552 | 184  | 
| unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs";  | 
| 16226 | 185  | 
fun reLAM 0 t = t  | 
| 16552 | 186  | 
      | reLAM n t = reLAM (n-1) (%%:"Cfun.Abs_CFun" $ Abs("",dummyT,t));
 | 
| 16226 | 187  | 
fun mplus (x,y) = %%:"Fixrec.mplus"`x`y;  | 
188  | 
val msum = foldr1 mplus (map (unLAM arity) ms);  | 
|
189  | 
in  | 
|
190  | 
reLAM arity (%%:"Fixrec.run"`msum)  | 
|
191  | 
end;  | 
|
192  | 
||
| 16387 | 193  | 
fun unzip3 [] = ([],[],[])  | 
194  | 
| unzip3 ((x,y,z)::ts) =  | 
|
195  | 
let val (xs,ys,zs) = unzip3 ts  | 
|
196  | 
in (x::xs, y::ys, z::zs) end;  | 
|
| 16226 | 197  | 
|
| 16387 | 198  | 
(* this is the pattern-matching compiler function *)  | 
199  | 
fun compile_pats eqs =  | 
|
200  | 
let  | 
|
201  | 
val ((n::names),(a::arities),mats) = unzip3 (map match_eq eqs);  | 
|
202  | 
val cname = if forall (fn x => n=x) names then n  | 
|
| 16552 | 203  | 
else fixrec_err "all equations in block must define the same function";  | 
| 16387 | 204  | 
val arity = if forall (fn x => a=x) arities then a  | 
| 16552 | 205  | 
else fixrec_err "all equations in block must have the same arity";  | 
| 16387 | 206  | 
val rhs = fatbar arity mats;  | 
207  | 
in  | 
|
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
208  | 
mk_trp (%%:cname === rhs)  | 
| 16387 | 209  | 
end;  | 
210  | 
||
211  | 
(*************************************************************************)  | 
|
212  | 
(********************** Proving associated theorems **********************)  | 
|
213  | 
(*************************************************************************)  | 
|
214  | 
||
| 16552 | 215  | 
(* proves a block of pattern matching equations as theorems, using unfold *)  | 
216  | 
fun make_simps thy (unfold_thm, eqns) =  | 
|
| 16226 | 217  | 
let  | 
| 17985 | 218  | 
val tacs = [rtac (unfold_thm RS ssubst_lhs) 1, asm_simp_tac (simpset_of thy) 1];  | 
219  | 
fun prove_term t = standard (Goal.prove thy [] [] t (K (EVERY tacs)));  | 
|
| 16552 | 220  | 
fun prove_eqn ((name, eqn_t), atts) = ((name, prove_term eqn_t), atts);  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
221  | 
in  | 
| 16552 | 222  | 
map prove_eqn eqns  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
223  | 
end;  | 
| 16226 | 224  | 
|
| 16387 | 225  | 
(*************************************************************************)  | 
226  | 
(************************* Main fixrec function **************************)  | 
|
227  | 
(*************************************************************************)  | 
|
| 16226 | 228  | 
|
| 16552 | 229  | 
fun gen_add_fixrec prep_prop prep_attrib strict blocks thy =  | 
| 16226 | 230  | 
let  | 
| 16552 | 231  | 
val eqns = List.concat blocks;  | 
232  | 
val lengths = map length blocks;  | 
|
233  | 
||
234  | 
val ((names, srcss), strings) = apfst split_list (split_list eqns);  | 
|
235  | 
val atts = map (map (prep_attrib thy)) srcss;  | 
|
236  | 
val eqn_ts = map (prep_prop thy) strings;  | 
|
237  | 
val rec_ts = map (fn eq => chead_of (fst (dest_eqs (Logic.strip_imp_concl eq)))  | 
|
| 16628 | 238  | 
handle TERM _ => fixrec_eq_err thy "not a proper equation" eq) eqn_ts;  | 
239  | 
val (_, eqn_ts') = InductivePackage.unify_consts thy rec_ts eqn_ts;  | 
|
| 16552 | 240  | 
|
241  | 
fun unconcat [] _ = []  | 
|
242  | 
| unconcat (n::ns) xs = List.take (xs,n) :: unconcat ns (List.drop (xs,n));  | 
|
243  | 
val pattern_blocks = unconcat lengths (map Logic.strip_imp_concl eqn_ts');  | 
|
| 16628 | 244  | 
val compiled_ts = map (infer thy o compile_pats) pattern_blocks;  | 
| 16552 | 245  | 
val (thy', cnames, fixdef_thms, unfold_thms) = add_fixdefs compiled_ts thy;  | 
| 16226 | 246  | 
in  | 
| 16552 | 247  | 
if strict then let (* only prove simp rules if strict = true *)  | 
248  | 
val eqn_blocks = unconcat lengths ((names ~~ eqn_ts') ~~ atts);  | 
|
249  | 
val simps = List.concat (map (make_simps thy') (unfold_thms ~~ eqn_blocks));  | 
|
250  | 
val (thy'', simp_thms) = PureThy.add_thms simps thy';  | 
|
251  | 
||
252  | 
val simp_names = map (fn name => name^"_simps") cnames;  | 
|
253  | 
val simp_attribute = rpair [Simplifier.simp_add_global];  | 
|
254  | 
val simps' = map simp_attribute (simp_names ~~ unconcat lengths simp_thms);  | 
|
255  | 
in  | 
|
256  | 
(#1 o PureThy.add_thmss simps') thy''  | 
|
257  | 
end  | 
|
258  | 
else thy'  | 
|
| 16226 | 259  | 
end;  | 
260  | 
||
| 16488 | 261  | 
val add_fixrec = gen_add_fixrec Sign.read_prop Attrib.global_attribute;  | 
262  | 
val add_fixrec_i = gen_add_fixrec Sign.cert_prop (K I);  | 
|
263  | 
||
264  | 
||
| 16387 | 265  | 
(*************************************************************************)  | 
266  | 
(******************************** Fixpat *********************************)  | 
|
267  | 
(*************************************************************************)  | 
|
| 16226 | 268  | 
|
| 16552 | 269  | 
fun fix_pat thy t =  | 
| 16226 | 270  | 
let  | 
| 16387 | 271  | 
val T = fastype_of t;  | 
| 
16401
 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 
huffman 
parents: 
16387 
diff
changeset
 | 
272  | 
    val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
 | 
| 
16461
 
e6b431cb8e0c
support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
 
huffman 
parents: 
16402 
diff
changeset
 | 
273  | 
val cname = case chead_of t of Const(c,_) => c | _ =>  | 
| 16552 | 274  | 
fixrec_err "function is not declared as constant in theory";  | 
| 16488 | 275  | 
val unfold_thm = PureThy.get_thm thy (Name (cname^"_unfold"));  | 
| 17985 | 276  | 
val simp = standard (Goal.prove thy [] [] eq  | 
277  | 
(fn _ => EVERY [stac unfold_thm 1, simp_tac (simpset_of thy) 1]));  | 
|
| 16552 | 278  | 
in simp end;  | 
| 16226 | 279  | 
|
| 16552 | 280  | 
fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy =  | 
| 16402 | 281  | 
let  | 
| 16552 | 282  | 
val atts = map (prep_attrib thy) srcs;  | 
283  | 
val ts = map (prep_term thy) strings;  | 
|
284  | 
val simps = map (fix_pat thy) ts;  | 
|
285  | 
in  | 
|
286  | 
(#1 o PureThy.add_thmss [((name, simps), atts)]) thy  | 
|
287  | 
end;  | 
|
| 16226 | 288  | 
|
| 16488 | 289  | 
val add_fixpat = gen_add_fixpat Sign.read_term Attrib.global_attribute;  | 
290  | 
val add_fixpat_i = gen_add_fixpat Sign.cert_term (K I);  | 
|
291  | 
||
292  | 
||
| 16387 | 293  | 
(*************************************************************************)  | 
294  | 
(******************************** Parsers ********************************)  | 
|
295  | 
(*************************************************************************)  | 
|
| 16226 | 296  | 
|
| 17057 | 297  | 
local structure P = OuterParse and K = OuterKeyword in  | 
| 16226 | 298  | 
|
| 16552 | 299  | 
val fixrec_eqn = P.opt_thm_name ":" -- P.prop;  | 
300  | 
||
| 16628 | 301  | 
val fixrec_strict = P.opt_keyword "permissive" >> not;  | 
| 16552 | 302  | 
|
303  | 
val fixrec_decl = fixrec_strict -- P.and_list1 (Scan.repeat1 fixrec_eqn);  | 
|
| 16226 | 304  | 
|
305  | 
(* this builds a parser for a new keyword, fixrec, whose functionality  | 
|
306  | 
is defined by add_fixrec *)  | 
|
307  | 
val fixrecP =  | 
|
| 16402 | 308  | 
OuterSyntax.command "fixrec" "define recursive functions (HOLCF)" K.thy_decl  | 
| 16552 | 309  | 
(fixrec_decl >> (Toplevel.theory o uncurry add_fixrec));  | 
| 16226 | 310  | 
|
311  | 
(* fixpat parser *)  | 
|
| 16552 | 312  | 
val fixpat_decl = P.opt_thm_name ":" -- Scan.repeat1 P.prop;  | 
| 16226 | 313  | 
|
314  | 
val fixpatP =  | 
|
| 16402 | 315  | 
OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl  | 
| 16226 | 316  | 
(fixpat_decl >> (Toplevel.theory o add_fixpat));  | 
317  | 
||
| 16552 | 318  | 
val _ = OuterSyntax.add_parsers [fixrecP, fixpatP];  | 
| 16226 | 319  | 
|
320  | 
end; (* local structure *)  | 
|
321  | 
||
322  | 
end; (* struct *)  |