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