author | wenzelm |
Tue, 23 May 2023 18:46:15 +0200 | |
changeset 78095 | bc42c074e58f |
parent 74561 | 8e6c973003c8 |
child 80636 | 4041e7c8059d |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Tools/fixrec.ML |
23152 | 2 |
Author: Amber Telfer and Brian Huffman |
3 |
||
4 |
Recursive function definition package for HOLCF. |
|
5 |
*) |
|
6 |
||
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31177
diff
changeset
|
7 |
signature FIXREC = |
23152 | 8 |
sig |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
9 |
val add_fixrec: (binding * typ option * mixfix) list |
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
10 |
-> (bool * (Attrib.binding * term)) list -> local_theory -> local_theory |
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
11 |
val add_fixrec_cmd: (binding * string option * mixfix) list |
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
12 |
-> (bool * (Attrib.binding * string)) list -> local_theory -> local_theory |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
13 |
val add_matchers: (string * string) list -> theory -> theory |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
14 |
val fixrec_simp_tac: Proof.context -> int -> tactic |
40832 | 15 |
end |
23152 | 16 |
|
41296
6aaf80ea9715
switch to transparent ascription, to avoid warning messages
huffman
parents:
40832
diff
changeset
|
17 |
structure Fixrec : FIXREC = |
23152 | 18 |
struct |
19 |
||
40832 | 20 |
open HOLCF_Library |
35527 | 21 |
|
40832 | 22 |
infixr 6 ->> |
23 |
infix -->> |
|
24 |
infix 9 ` |
|
35527 | 25 |
|
40832 | 26 |
val def_cont_fix_eq = @{thm def_cont_fix_eq} |
27 |
val def_cont_fix_ind = @{thm def_cont_fix_ind} |
|
23152 | 28 |
|
40832 | 29 |
fun fixrec_err s = error ("fixrec definition error:\n" ^ s) |
23152 | 30 |
|
30132 | 31 |
(*************************************************************************) |
32 |
(***************************** building types ****************************) |
|
33 |
(*************************************************************************) |
|
34 |
||
33401
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
35 |
local |
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
36 |
|
74305 | 37 |
fun binder_cfun \<^Type>\<open>cfun T U\<close> = T :: binder_cfun U |
38 |
| binder_cfun \<^Type>\<open>fun T U\<close> = T :: binder_cfun U |
|
40832 | 39 |
| binder_cfun _ = [] |
33401
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
40 |
|
74305 | 41 |
fun body_cfun \<^Type>\<open>cfun _ U\<close> = body_cfun U |
42 |
| body_cfun \<^Type>\<open>fun _ U\<close> = body_cfun U |
|
40832 | 43 |
| body_cfun T = T |
33401
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
44 |
|
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
45 |
in |
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
46 |
|
35527 | 47 |
fun matcherT (T, U) = |
40832 | 48 |
body_cfun T ->> (binder_cfun T -->> U) ->> U |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30485
diff
changeset
|
49 |
|
33401
fc43fa403a69
add fixrec support for HOL pair constructor patterns
huffman
parents:
33004
diff
changeset
|
50 |
end |
30132 | 51 |
|
52 |
(*************************************************************************) |
|
53 |
(***************************** building terms ****************************) |
|
54 |
(*************************************************************************) |
|
23152 | 55 |
|
40832 | 56 |
val mk_trp = HOLogic.mk_Trueprop |
23152 | 57 |
|
58 |
(* splits a cterm into the right and lefthand sides of equality *) |
|
40832 | 59 |
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t) |
23152 | 60 |
|
61 |
(* similar to Thm.head_of, but for continuous application *) |
|
74305 | 62 |
fun chead_of \<^Const_>\<open>Rep_cfun _ _ for f _\<close> = chead_of f |
40832 | 63 |
| chead_of u = u |
23152 | 64 |
|
40832 | 65 |
infix 1 === val (op ===) = HOLogic.mk_eq |
30132 | 66 |
|
67 |
fun mk_mplus (t, u) = |
|
74305 | 68 |
let val T = dest_matchT (Term.fastype_of t) |
69 |
in \<^Const>\<open>Fixrec.mplus T\<close> ` t ` u end |
|
30132 | 70 |
|
71 |
fun mk_run t = |
|
37097
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
72 |
let |
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
73 |
val mT = Term.fastype_of t |
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
74 |
val T = dest_matchT mT |
74305 | 75 |
val run = \<^Const>\<open>Fixrec.run T\<close> |
37097
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
76 |
in |
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
77 |
case t of |
74305 | 78 |
\<^Const_>\<open>Rep_cfun _ _\<close> $ \<^Const_>\<open>Fixrec.succeed _\<close> $ u => u |
37097
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
79 |
| _ => run ` t |
40832 | 80 |
end |
30132 | 81 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31023
diff
changeset
|
82 |
|
23152 | 83 |
(*************************************************************************) |
84 |
(************* fixed-point definitions and unfolding theorems ************) |
|
85 |
(*************************************************************************) |
|
86 |
||
33519 | 87 |
structure FixrecUnfoldData = Generic_Data |
88 |
( |
|
40832 | 89 |
type T = thm Symtab.table |
90 |
val empty = Symtab.empty |
|
91 |
fun merge data : T = Symtab.merge (K true) data |
|
92 |
) |
|
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
93 |
|
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
94 |
local |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
95 |
|
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
96 |
fun name_of (Const (n, _)) = n |
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
97 |
| name_of (Free (n, _)) = n |
40832 | 98 |
| name_of t = raise TERM ("Fixrec.add_unfold: lhs not a constant", [t]) |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
99 |
|
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
100 |
val lhs_name = |
59582 | 101 |
name_of o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
102 |
|
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
103 |
in |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
104 |
|
33442 | 105 |
val add_unfold : attribute = |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
106 |
Thm.declaration_attribute |
40832 | 107 |
(fn th => FixrecUnfoldData.map (Symtab.insert (K true) (lhs_name th, th))) |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
108 |
|
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
109 |
end |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
110 |
|
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
111 |
fun add_fixdefs |
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
112 |
(fixes : ((binding * typ) * mixfix) list) |
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
113 |
(spec : (Attrib.binding * term) list) |
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
114 |
(lthy : local_theory) = |
23152 | 115 |
let |
42361 | 116 |
val thy = Proof_Context.theory_of lthy |
40832 | 117 |
val names = map (Binding.name_of o fst o fst) fixes |
118 |
val all_names = space_implode "_" names |
|
119 |
val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec) |
|
120 |
val functional = lambda_tuple lhss (mk_tuple rhss) |
|
121 |
val fixpoint = mk_fix (mk_cabs functional) |
|
36996
63fadc0a33db
more informative error message for fixrec when continuity proof fails
huffman
parents:
36960
diff
changeset
|
122 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31023
diff
changeset
|
123 |
val cont_thm = |
36996
63fadc0a33db
more informative error message for fixrec when continuity proof fails
huffman
parents:
36960
diff
changeset
|
124 |
let |
40832 | 125 |
val prop = mk_trp (mk_cont functional) |
37096 | 126 |
fun err _ = error ( |
40832 | 127 |
"Continuity proof failed please check that cont2cont rules\n" ^ |
37097
476016cbf8b3
for functions with only variable patterns, fixrec definitions no longer use Fixrec.return/Fixrec.run
huffman
parents:
37096
diff
changeset
|
128 |
"or simp rules are configured for all non-HOLCF constants.\n" ^ |
36996
63fadc0a33db
more informative error message for fixrec when continuity proof fails
huffman
parents:
36960
diff
changeset
|
129 |
"The error occurred for the goal statement:\n" ^ |
40832 | 130 |
Syntax.string_of_term lthy prop) |
69597 | 131 |
val rules = Named_Theorems.get lthy \<^named_theorems>\<open>cont2cont\<close> |
58957 | 132 |
val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac lthy (rev rules))) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
133 |
val slow_tac = SOLVED' (simp_tac lthy) |
40832 | 134 |
val tac = fast_tac 1 ORELSE slow_tac 1 ORELSE err |
36996
63fadc0a33db
more informative error message for fixrec when continuity proof fails
huffman
parents:
36960
diff
changeset
|
135 |
in |
63fadc0a33db
more informative error message for fixrec when continuity proof fails
huffman
parents:
36960
diff
changeset
|
136 |
Goal.prove lthy [] [] prop (K tac) |
40832 | 137 |
end |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31023
diff
changeset
|
138 |
|
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
139 |
fun one_def (Free(n,_)) r = |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
140 |
let val b = Long_Name.base_name n |
46909 | 141 |
in ((Binding.name (Thm.def_name b), []), r) end |
40832 | 142 |
| one_def _ _ = fixrec_err "fixdefs: lhs not of correct form" |
23152 | 143 |
fun defs [] _ = [] |
144 |
| defs (l::[]) r = [one_def l r] |
|
40832 | 145 |
| defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r) |
146 |
val fixdefs = defs lhss fixpoint |
|
35772 | 147 |
val (fixdef_thms : (term * (string * thm)) list, lthy) = lthy |
40832 | 148 |
|> fold_map Local_Theory.define (map (apfst fst) fixes ~~ fixdefs) |
149 |
fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2] |
|
150 |
val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms) |
|
74305 | 151 |
val P = Var (("P", 0), map Term.fastype_of lhss ---> \<^Type>\<open>bool\<close>) |
40832 | 152 |
val predicate = lambda_tuple lhss (list_comb (P, lhss)) |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31023
diff
changeset
|
153 |
val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm]) |
60801 | 154 |
|> Thm.instantiate' [] [SOME (Thm.global_cterm_of thy predicate)] |
40832 | 155 |
|> Local_Defs.unfold lthy @{thms split_paired_all split_conv split_strict} |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31023
diff
changeset
|
156 |
val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) |
40832 | 157 |
|> Local_Defs.unfold lthy @{thms split_conv} |
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
158 |
fun unfolds [] _ = [] |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
159 |
| unfolds (n::[]) thm = [(n, thm)] |
23152 | 160 |
| unfolds (n::ns) thm = let |
40832 | 161 |
val thmL = thm RS @{thm Pair_eqD1} |
162 |
val thmR = thm RS @{thm Pair_eqD2} |
|
163 |
in (n, thmL) :: unfolds ns thmR end |
|
164 |
val unfold_thms = unfolds names tuple_unfold_thm |
|
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
165 |
val induct_note : Attrib.binding * Thm.thm list = |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
166 |
let |
40832 | 167 |
val thm_name = Binding.qualify true all_names (Binding.name "induct") |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
168 |
in |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
169 |
((thm_name, []), [tuple_induct_thm]) |
40832 | 170 |
end |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
171 |
fun unfold_note (name, thm) : Attrib.binding * Thm.thm list = |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
172 |
let |
40832 | 173 |
val thm_name = Binding.qualify true name (Binding.name "unfold") |
78095 | 174 |
val src = Attrib.internal \<^here> (K add_unfold) |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
175 |
in |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
176 |
((thm_name, [src]), [thm]) |
40832 | 177 |
end |
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
178 |
val (_, lthy) = lthy |
40832 | 179 |
|> fold_map Local_Theory.note (induct_note :: map unfold_note unfold_thms) |
23152 | 180 |
in |
35772 | 181 |
(lthy, names, fixdef_thms, map snd unfold_thms) |
40832 | 182 |
end |
23152 | 183 |
|
184 |
(*************************************************************************) |
|
185 |
(*********** monadic notation and pattern matching compilation ***********) |
|
186 |
(*************************************************************************) |
|
187 |
||
33522 | 188 |
structure FixrecMatchData = Theory_Data |
189 |
( |
|
40832 | 190 |
type T = string Symtab.table |
191 |
val empty = Symtab.empty |
|
192 |
fun merge data = Symtab.merge (K true) data |
|
193 |
) |
|
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
194 |
|
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
195 |
(* associate match functions with pattern constants *) |
40832 | 196 |
fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms) |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
197 |
|
30157 | 198 |
fun taken_names (t : term) : bstring list = |
199 |
let |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
200 |
fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs |
30157 | 201 |
| taken (Free(a,_) , bs) = insert (op =) a bs |
202 |
| taken (f $ u , bs) = taken (f, taken (u, bs)) |
|
203 |
| taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs) |
|
40832 | 204 |
| taken (_ , bs) = bs |
30157 | 205 |
in |
206 |
taken (t, []) |
|
40832 | 207 |
end |
23152 | 208 |
|
39805 | 209 |
(* builds a monadic term for matching a pattern *) |
210 |
(* returns (rhs, free variable, used varnames) *) |
|
211 |
fun compile_pat match_name pat rhs taken = |
|
212 |
let |
|
213 |
fun comp_pat p rhs taken = |
|
214 |
if is_Free p then (rhs, p, taken) |
|
215 |
else comp_con (fastype_of p) p rhs [] taken |
|
216 |
(* compiles a monadic term for a constructor pattern *) |
|
217 |
and comp_con T p rhs vs taken = |
|
218 |
case p of |
|
74305 | 219 |
\<^Const_>\<open>Rep_cfun _ _ for f x\<close> => |
39805 | 220 |
let val (rhs', v, taken') = comp_pat x rhs taken |
221 |
in comp_con T f rhs' (v::vs) taken' end |
|
222 |
| f $ x => |
|
223 |
let val (rhs', v, taken') = comp_pat x rhs taken |
|
224 |
in comp_con T f rhs' (v::vs) taken' end |
|
225 |
| Const (c, cT) => |
|
226 |
let |
|
43324
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
42361
diff
changeset
|
227 |
val n = singleton (Name.variant_list taken) "v" |
39805 | 228 |
val v = Free(n, T) |
229 |
val m = Const(match_name c, matcherT (cT, fastype_of rhs)) |
|
230 |
val k = big_lambdas vs rhs |
|
231 |
in |
|
232 |
(m`v`k, v, n::taken) |
|
233 |
end |
|
234 |
| _ => raise TERM ("fixrec: invalid pattern ", [p]) |
|
235 |
in |
|
236 |
comp_pat pat rhs taken |
|
40832 | 237 |
end |
23152 | 238 |
|
239 |
(* builds a monadic term for matching a function definition pattern *) |
|
40036 | 240 |
(* returns (constant, (vars, matcher)) *) |
39804
b1cec1fcd95f
rename some fixrec pattern-match compilation functions
huffman
parents:
39557
diff
changeset
|
241 |
fun compile_lhs match_name pat rhs vs taken = |
23152 | 242 |
case pat of |
74305 | 243 |
\<^Const_>\<open>Rep_cfun _ _ for f x\<close> => |
40832 | 244 |
let val (rhs', v, taken') = compile_pat match_name x rhs taken |
39804
b1cec1fcd95f
rename some fixrec pattern-match compilation functions
huffman
parents:
39557
diff
changeset
|
245 |
in compile_lhs match_name f rhs' (v::vs) taken' end |
40036 | 246 |
| Free(_,_) => (pat, (vs, rhs)) |
247 |
| Const(_,_) => (pat, (vs, rhs)) |
|
248 |
| _ => fixrec_err ("invalid function pattern: " |
|
40832 | 249 |
^ ML_Syntax.print_term pat) |
23152 | 250 |
|
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
251 |
fun strip_alls t = |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74305
diff
changeset
|
252 |
(case try Logic.dest_all_global t of |
46895 | 253 |
SOME (_, u) => strip_alls u |
254 |
| NONE => t) |
|
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
255 |
|
39804
b1cec1fcd95f
rename some fixrec pattern-match compilation functions
huffman
parents:
39557
diff
changeset
|
256 |
fun compile_eq match_name eq = |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
257 |
let |
40832 | 258 |
val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq)) |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
259 |
in |
39804
b1cec1fcd95f
rename some fixrec pattern-match compilation functions
huffman
parents:
39557
diff
changeset
|
260 |
compile_lhs match_name lhs (mk_succeed rhs) [] (taken_names eq) |
40832 | 261 |
end |
23152 | 262 |
|
263 |
(* this is the pattern-matching compiler function *) |
|
39804
b1cec1fcd95f
rename some fixrec pattern-match compilation functions
huffman
parents:
39557
diff
changeset
|
264 |
fun compile_eqs match_name eqs = |
23152 | 265 |
let |
40036 | 266 |
val (consts, matchers) = |
40832 | 267 |
ListPair.unzip (map (compile_eq match_name) eqs) |
40036 | 268 |
val const = |
269 |
case distinct (op =) consts of |
|
35903 | 270 |
[n] => n |
45787 | 271 |
| [] => fixrec_err "no defining equations for function" |
40832 | 272 |
| _ => fixrec_err "all equations in block must define the same function" |
40036 | 273 |
val vars = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58957
diff
changeset
|
274 |
case distinct (op = o apply2 length) (map fst matchers) of |
40036 | 275 |
[vars] => vars |
40832 | 276 |
| _ => fixrec_err "all equations in block must have the same arity" |
40036 | 277 |
(* rename so all matchers use same free variables *) |
40832 | 278 |
fun rename (vs, t) = Term.subst_free (filter_out (op =) (vs ~~ vars)) t |
279 |
val rhs = big_lambdas vars (mk_run (foldr1 mk_mplus (map rename matchers))) |
|
23152 | 280 |
in |
40036 | 281 |
mk_trp (const === rhs) |
40832 | 282 |
end |
23152 | 283 |
|
284 |
(*************************************************************************) |
|
285 |
(********************** Proving associated theorems **********************) |
|
286 |
(*************************************************************************) |
|
287 |
||
40832 | 288 |
fun eta_tac i = CONVERSION Thm.eta_conversion i |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
289 |
|
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
290 |
fun fixrec_simp_tac ctxt = |
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
291 |
let |
40832 | 292 |
val tab = FixrecUnfoldData.get (Context.Proof ctxt) |
46895 | 293 |
val concl = HOLogic.dest_Trueprop o Logic.strip_imp_concl o strip_alls |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
294 |
fun tac (t, i) = |
33430 | 295 |
let |
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
296 |
val (c, _) = |
40832 | 297 |
(dest_Const o head_of o chead_of o fst o HOLogic.dest_eq o concl) t |
298 |
val unfold_thm = the (Symtab.lookup tab c) |
|
299 |
val rule = unfold_thm RS @{thm ssubst_lhs} |
|
33430 | 300 |
in |
60754 | 301 |
CHANGED (resolve_tac ctxt [rule] i THEN eta_tac i THEN asm_simp_tac ctxt i) |
33430 | 302 |
end |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
303 |
in |
33505 | 304 |
SUBGOAL (fn ti => the_default no_tac (try tac ti)) |
40832 | 305 |
end |
33425
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
huffman
parents:
33401
diff
changeset
|
306 |
|
23152 | 307 |
(* proves a block of pattern matching equations as theorems, using unfold *) |
32149
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents:
31740
diff
changeset
|
308 |
fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = |
23152 | 309 |
let |
40832 | 310 |
val rule = unfold_thm RS @{thm ssubst_lhs} |
60754 | 311 |
val tac = resolve_tac ctxt [rule] 1 THEN eta_tac 1 THEN asm_simp_tac ctxt 1 |
40832 | 312 |
fun prove_term t = Goal.prove ctxt [] [] t (K tac) |
313 |
fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t) |
|
23152 | 314 |
in |
315 |
map prove_eqn eqns |
|
40832 | 316 |
end |
23152 | 317 |
|
318 |
(*************************************************************************) |
|
319 |
(************************* Main fixrec function **************************) |
|
320 |
(*************************************************************************) |
|
321 |
||
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
322 |
local |
45897
65cef0298158
clarified modules that contribute to datatype package;
wenzelm
parents:
45787
diff
changeset
|
323 |
(* code adapted from HOL/Tools/Datatype/primrec.ML *) |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
324 |
|
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
325 |
fun gen_fixrec |
30485 | 326 |
prep_spec |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
327 |
(raw_fixes : (binding * 'a option * mixfix) list) |
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
328 |
(raw_spec' : (bool * (Attrib.binding * 'b)) list) |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
329 |
(lthy : local_theory) = |
23152 | 330 |
let |
40832 | 331 |
val (skips, raw_spec) = ListPair.unzip raw_spec' |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
332 |
val (fixes : ((binding * typ) * mixfix) list, |
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
333 |
spec : (Attrib.binding * term) list) = |
63182 | 334 |
fst (prep_spec raw_fixes (map (fn s => (s, [], [])) raw_spec) lthy) |
45787 | 335 |
val names = map (Binding.name_of o fst o fst) fixes |
336 |
fun check_head name = |
|
337 |
member (op =) names name orelse |
|
338 |
fixrec_err ("Illegal equation head. Expected " ^ commas_quote names) |
|
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
339 |
val chead_of_spec = |
40832 | 340 |
chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd |
45787 | 341 |
fun name_of (Free (n, _)) = tap check_head n |
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
342 |
| name_of _ = fixrec_err ("unknown term") |
40832 | 343 |
val all_names = map (name_of o chead_of_spec) spec |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
344 |
fun block_of_name n = |
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
345 |
map_filter |
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
346 |
(fn (m,eq) => if m = n then SOME eq else NONE) |
40832 | 347 |
(all_names ~~ (spec ~~ skips)) |
348 |
val blocks = map block_of_name names |
|
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
349 |
|
42361 | 350 |
val matcher_tab = FixrecMatchData.get (Proof_Context.theory_of lthy) |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
351 |
fun match_name c = |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
352 |
case Symtab.lookup matcher_tab c of SOME m => m |
40832 | 353 |
| NONE => fixrec_err ("unknown pattern constructor: " ^ c) |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29585
diff
changeset
|
354 |
|
40832 | 355 |
val matches = map (compile_eqs match_name) (map (map (snd o fst)) blocks) |
63352 | 356 |
val spec' = map (pair Binding.empty_atts) matches |
44080
53d95b52954c
HOLCF: fix warnings about unreferenced identifiers
huffman
parents:
43324
diff
changeset
|
357 |
val (lthy, _, _, unfold_thms) = |
40832 | 358 |
add_fixdefs fixes spec' lthy |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
359 |
|
40832 | 360 |
val blocks' = map (map fst o filter_out snd) blocks |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
361 |
val simps : (Attrib.binding * thm) list list = |
40832 | 362 |
map (make_simps lthy) (unfold_thms ~~ blocks') |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
363 |
fun mk_bind n : Attrib.binding = |
45592 | 364 |
(Binding.qualify true n (Binding.name "simps"), @{attributes [simp]}) |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
365 |
val simps1 : (Attrib.binding * thm list) list = |
40832 | 366 |
map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps) |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
367 |
val simps2 : (Attrib.binding * thm list) list = |
40832 | 368 |
map (apsnd (fn thm => [thm])) (flat simps) |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
369 |
val (_, lthy) = lthy |
40832 | 370 |
|> fold_map Local_Theory.note (simps1 @ simps2) |
23152 | 371 |
in |
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
372 |
lthy |
40832 | 373 |
end |
23152 | 374 |
|
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
375 |
in |
23152 | 376 |
|
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
60801
diff
changeset
|
377 |
val add_fixrec = gen_fixrec Specification.check_multi_specs |
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
60801
diff
changeset
|
378 |
val add_fixrec_cmd = gen_fixrec Specification.read_multi_specs |
30158
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
huffman
parents:
30157
diff
changeset
|
379 |
|
40832 | 380 |
end (* local *) |
23152 | 381 |
|
382 |
||
383 |
(*************************************************************************) |
|
384 |
(******************************** Parsers ********************************) |
|
385 |
(*************************************************************************) |
|
386 |
||
40096
4d1a8fa8cdfd
change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents:
40041
diff
changeset
|
387 |
val opt_thm_name' : (bool * Attrib.binding) parser = |
69597 | 388 |
\<^keyword>\<open>(\<close> -- \<^keyword>\<open>unchecked\<close> -- \<^keyword>\<open>)\<close> >> K (true, Binding.empty_atts) |
40832 | 389 |
|| Parse_Spec.opt_thm_name ":" >> pair false |
40096
4d1a8fa8cdfd
change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents:
40041
diff
changeset
|
390 |
|
4d1a8fa8cdfd
change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents:
40041
diff
changeset
|
391 |
val spec' : (bool * (Attrib.binding * string)) parser = |
40832 | 392 |
opt_thm_name' -- Parse.prop >> (fn ((a, b), c) => (a, (b, c))) |
40096
4d1a8fa8cdfd
change fixrec parser to not accept theorem names with (unchecked) option
huffman
parents:
40041
diff
changeset
|
393 |
|
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
60801
diff
changeset
|
394 |
val multi_specs' : (bool * (Attrib.binding * string)) list parser = |
69597 | 395 |
let val unexpected = Scan.ahead (Parse.name || \<^keyword>\<open>[\<close> || \<^keyword>\<open>(\<close>) |
396 |
in Parse.enum1 "|" (spec' --| Scan.option (unexpected -- Parse.!!! \<^keyword>\<open>|\<close>)) end |
|
40041
1f09b4c7b85e
replace fixrec 'permissive' mode with per-equation 'unchecked' option
huffman
parents:
40036
diff
changeset
|
397 |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
398 |
val _ = |
69597 | 399 |
Outer_Syntax.local_theory \<^command_keyword>\<open>fixrec\<close> "define recursive functions (HOLCF)" |
63285 | 400 |
(Parse.vars -- (Parse.where_ |-- Parse.!!! multi_specs') |
401 |
>> (fn (vars, specs) => add_fixrec_cmd vars specs)) |
|
23152 | 402 |
|
40832 | 403 |
end |