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