| author | huffman | 
| Tue, 02 Jun 2009 15:13:22 -0700 | |
| changeset 31390 | 1d0478b16613 | 
| parent 31177 | c39994cb152a | 
| child 31738 | 7b9b9ba532ca | 
| permissions | -rw-r--r-- | 
| 23152 | 1 | (* Title: HOLCF/Tools/fixrec_package.ML | 
| 2 | Author: Amber Telfer and Brian Huffman | |
| 3 | ||
| 4 | Recursive function definition package for HOLCF. | |
| 5 | *) | |
| 6 | ||
| 7 | signature FIXREC_PACKAGE = | |
| 8 | sig | |
| 30485 | 9 | val add_fixrec: bool -> (binding * typ option * mixfix) list | 
| 10 | -> (Attrib.binding * term) list -> local_theory -> local_theory | |
| 11 | val add_fixrec_cmd: bool -> (binding * string option * mixfix) list | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 12 | -> (Attrib.binding * string) list -> local_theory -> local_theory | 
| 30485 | 13 | val add_fixpat: Thm.binding * term list -> theory -> theory | 
| 14 | val add_fixpat_cmd: Attrib.binding * string list -> theory -> theory | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 15 | val add_matchers: (string * string) list -> theory -> theory | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 16 | val setup: theory -> theory | 
| 23152 | 17 | end; | 
| 18 | ||
| 31023 | 19 | structure FixrecPackage :> FIXREC_PACKAGE = | 
| 23152 | 20 | struct | 
| 21 | ||
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 22 | val def_cont_fix_eq = @{thm def_cont_fix_eq};
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 23 | val def_cont_fix_ind = @{thm def_cont_fix_ind};
 | 
| 23152 | 24 | |
| 25 | ||
| 26 | fun fixrec_err s = error ("fixrec definition error:\n" ^ s);
 | |
| 27 | fun fixrec_eq_err thy s eq = | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26343diff
changeset | 28 | fixrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); | 
| 23152 | 29 | |
| 30132 | 30 | (*************************************************************************) | 
| 31 | (***************************** building types ****************************) | |
| 32 | (*************************************************************************) | |
| 33 | ||
| 23152 | 34 | (* ->> is taken from holcf_logic.ML *) | 
| 30132 | 35 | fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
 | 
| 36 | ||
| 37 | infixr 6 ->>; val (op ->>) = cfunT; | |
| 38 | ||
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 39 | fun cfunsT (Ts, U) = foldr cfunT U Ts; | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 40 | |
| 30132 | 41 | fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U)
 | 
| 42 |   | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
 | |
| 43 | ||
| 44 | fun binder_cfun (Type(@{type_name "->"},[T, U])) = T :: binder_cfun U
 | |
| 45 | | binder_cfun _ = []; | |
| 46 | ||
| 47 | fun body_cfun (Type(@{type_name "->"},[T, U])) = body_cfun U
 | |
| 48 | | body_cfun T = T; | |
| 23152 | 49 | |
| 30132 | 50 | fun strip_cfun T : typ list * typ = | 
| 51 | (binder_cfun T, body_cfun T); | |
| 52 | ||
| 53 | fun maybeT T = Type(@{type_name "maybe"}, [T]);
 | |
| 54 | ||
| 55 | fun dest_maybeT (Type(@{type_name "maybe"}, [T])) = T
 | |
| 56 |   | dest_maybeT T = raise TYPE ("dest_maybeT", [T], []);
 | |
| 57 | ||
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 58 | fun tupleT [] = HOLogic.unitT | 
| 30132 | 59 | | tupleT [T] = T | 
| 60 | | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts); | |
| 61 | ||
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 62 | fun matchT (T, U) = | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 63 | body_cfun T ->> cfunsT (binder_cfun T, U) ->> U; | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 64 | |
| 30132 | 65 | |
| 66 | (*************************************************************************) | |
| 67 | (***************************** building terms ****************************) | |
| 68 | (*************************************************************************) | |
| 23152 | 69 | |
| 70 | val mk_trp = HOLogic.mk_Trueprop; | |
| 71 | ||
| 72 | (* splits a cterm into the right and lefthand sides of equality *) | |
| 73 | fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t); | |
| 74 | ||
| 75 | (* similar to Thm.head_of, but for continuous application *) | |
| 26045 | 76 | fun chead_of (Const(@{const_name Rep_CFun},_)$f$t) = chead_of f
 | 
| 23152 | 77 | | chead_of u = u; | 
| 78 | ||
| 30132 | 79 | fun capply_const (S, T) = | 
| 80 |   Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T));
 | |
| 81 | ||
| 82 | fun cabs_const (S, T) = | |
| 83 |   Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T));
 | |
| 84 | ||
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 85 | fun mk_cabs t = | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 86 | let val T = Term.fastype_of t | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 87 | in cabs_const (Term.domain_type T, Term.range_type T) $ t end | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 88 | |
| 30132 | 89 | fun mk_capply (t, u) = | 
| 90 | let val (S, T) = | |
| 91 | case Term.fastype_of t of | |
| 92 |         Type(@{type_name "->"}, [S, T]) => (S, T)
 | |
| 93 |       | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]);
 | |
| 94 | in capply_const (S, T) $ t $ u end; | |
| 95 | ||
| 96 | infix 0 ==; val (op ==) = Logic.mk_equals; | |
| 97 | infix 1 ===; val (op ===) = HOLogic.mk_eq; | |
| 98 | infix 9 ` ; val (op `) = mk_capply; | |
| 99 | ||
| 23152 | 100 | (* builds the expression (LAM v. rhs) *) | 
| 30132 | 101 | fun big_lambda v rhs = | 
| 102 | cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs; | |
| 23152 | 103 | |
| 104 | (* builds the expression (LAM v1 v2 .. vn. rhs) *) | |
| 105 | fun big_lambdas [] rhs = rhs | |
| 106 | | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs); | |
| 107 | ||
| 30132 | 108 | fun mk_return t = | 
| 109 | let val T = Term.fastype_of t | |
| 110 |   in Const(@{const_name Fixrec.return}, T ->> maybeT T) ` t end;
 | |
| 111 | ||
| 112 | fun mk_bind (t, u) = | |
| 113 | let val (T, mU) = dest_cfunT (Term.fastype_of u); | |
| 114 | val bindT = maybeT T ->> (T ->> mU) ->> mU; | |
| 115 |   in Const(@{const_name Fixrec.bind}, bindT) ` t ` u end;
 | |
| 116 | ||
| 117 | fun mk_mplus (t, u) = | |
| 118 | let val mT = Term.fastype_of t | |
| 119 |   in Const(@{const_name Fixrec.mplus}, mT ->> mT ->> mT) ` t ` u end;
 | |
| 120 | ||
| 121 | fun mk_run t = | |
| 122 | let val mT = Term.fastype_of t | |
| 123 | val T = dest_maybeT mT | |
| 124 |   in Const(@{const_name Fixrec.run}, mT ->> T) ` t end;
 | |
| 125 | ||
| 126 | fun mk_fix t = | |
| 127 | let val (T, _) = dest_cfunT (Term.fastype_of t) | |
| 128 |   in Const(@{const_name fix}, (T ->> T) ->> T) ` t end;
 | |
| 23152 | 129 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 130 | fun mk_cont t = | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 131 | let val T = Term.fastype_of t | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 132 |   in Const(@{const_name cont}, T --> HOLogic.boolT) $ t end;
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 133 | |
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 134 | val mk_fst = HOLogic.mk_fst | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 135 | val mk_snd = HOLogic.mk_snd | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 136 | |
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 137 | (* builds the expression (v1,v2,..,vn) *) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 138 | fun mk_tuple [] = HOLogic.unit | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 139 | | mk_tuple (t::[]) = t | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 140 | | mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts); | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 141 | |
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 142 | (* builds the expression (%(v1,v2,..,vn). rhs) *) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 143 | fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 144 | | lambda_tuple (v::[]) rhs = Term.lambda v rhs | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 145 | | lambda_tuple (v::vs) rhs = | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 146 | HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs)); | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 147 | |
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 148 | |
| 23152 | 149 | (*************************************************************************) | 
| 150 | (************* fixed-point definitions and unfolding theorems ************) | |
| 151 | (*************************************************************************) | |
| 152 | ||
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 153 | fun add_fixdefs | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 154 | (fixes : ((binding * typ) * mixfix) list) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 155 | (spec : (Attrib.binding * term) list) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 156 | (lthy : local_theory) = | 
| 23152 | 157 | let | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 158 | val thy = ProofContext.theory_of lthy; | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
30211diff
changeset | 159 | val names = map (Binding.name_of o fst o fst) fixes; | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 160 | val all_names = space_implode "_" names; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 161 | val (lhss,rhss) = ListPair.unzip (map (dest_eqs o snd) spec); | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 162 | val functional = lambda_tuple lhss (mk_tuple rhss); | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 163 | val fixpoint = mk_fix (mk_cabs functional); | 
| 23152 | 164 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 165 | val cont_thm = | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 166 | Goal.prove lthy [] [] (mk_trp (mk_cont functional)) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 167 | (K (simp_tac (local_simpset_of lthy) 1)); | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 168 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 169 | fun one_def (l as Free(n,_)) r = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 170 | let val b = Long_Name.base_name n | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 171 | in ((Binding.name (b^"_def"), []), r) end | 
| 23152 | 172 | | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form"; | 
| 173 | fun defs [] _ = [] | |
| 174 | | defs (l::[]) r = [one_def l r] | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 175 | | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r); | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 176 | val fixdefs = defs lhss fixpoint; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 177 | val define_all = fold_map (LocalTheory.define Thm.definitionK); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 178 | val (fixdef_thms : (term * (string * thm)) list, lthy') = lthy | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 179 | |> define_all (map (apfst fst) fixes ~~ fixdefs); | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 180 |     fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2];
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 181 | val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms); | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 182 |     val P = Var (("P", 0), map Term.fastype_of lhss ---> HOLogic.boolT);
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 183 | val predicate = lambda_tuple lhss (list_comb (P, lhss)); | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 184 | 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 | 185 | |> Drule.instantiate' [] [SOME (Thm.cterm_of thy predicate)] | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 186 |       |> LocalDefs.unfold lthy @{thms split_paired_all split_conv split_strict};
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 187 | val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 188 |       |> LocalDefs.unfold lthy' @{thms split_conv};
 | 
| 23152 | 189 | fun unfolds [] thm = [] | 
| 190 | | unfolds (n::[]) thm = [(n^"_unfold", thm)] | |
| 191 | | unfolds (n::ns) thm = let | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 192 |           val thmL = thm RS @{thm Pair_eqD1};
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 193 |           val thmR = thm RS @{thm Pair_eqD2};
 | 
| 23152 | 194 | in (n^"_unfold", thmL) :: unfolds ns thmR end; | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 195 | val unfold_thms = unfolds names tuple_unfold_thm; | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 196 | fun mk_note (n, thm) = ((Binding.name n, []), [thm]); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 197 | val (thmss, lthy'') = lthy' | 
| 31177 | 198 | |> fold_map (LocalTheory.note Thm.generatedK o mk_note) | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 199 | ((all_names ^ "_induct", tuple_induct_thm) :: unfold_thms); | 
| 23152 | 200 | in | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 201 | (lthy'', names, fixdef_thms, map snd unfold_thms) | 
| 23152 | 202 | end; | 
| 203 | ||
| 204 | (*************************************************************************) | |
| 205 | (*********** monadic notation and pattern matching compilation ***********) | |
| 206 | (*************************************************************************) | |
| 207 | ||
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 208 | structure FixrecMatchData = TheoryDataFun ( | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 209 | type T = string Symtab.table; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 210 | val empty = Symtab.empty; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 211 | val copy = I; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 212 | val extend = I; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 213 | fun merge _ tabs : T = Symtab.merge (K true) tabs; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 214 | ); | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 215 | |
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 216 | (* associate match functions with pattern constants *) | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 217 | fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms); | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 218 | |
| 30157 | 219 | fun taken_names (t : term) : bstring list = | 
| 220 | let | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 221 | fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs | 
| 30157 | 222 | | taken (Free(a,_) , bs) = insert (op =) a bs | 
| 223 | | taken (f $ u , bs) = taken (f, taken (u, bs)) | |
| 224 | | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs) | |
| 225 | | taken (_ , bs) = bs; | |
| 226 | in | |
| 227 | taken (t, []) | |
| 228 | end; | |
| 23152 | 229 | |
| 230 | (* builds a monadic term for matching a constructor pattern *) | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 231 | fun pre_build match_name pat rhs vs taken = | 
| 23152 | 232 | case pat of | 
| 26045 | 233 |     Const(@{const_name Rep_CFun},_)$f$(v as Free(n,T)) =>
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 234 | pre_build match_name f rhs (v::vs) taken | 
| 26045 | 235 |   | Const(@{const_name Rep_CFun},_)$f$x =>
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 236 | let val (rhs', v, taken') = pre_build match_name x rhs [] taken; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 237 | in pre_build match_name f rhs' (v::vs) taken' end | 
| 23152 | 238 | | Const(c,T) => | 
| 239 | let | |
| 240 | val n = Name.variant taken "v"; | |
| 26045 | 241 |         fun result_type (Type(@{type_name "->"},[_,T])) (x::xs) = result_type T xs
 | 
| 23152 | 242 | | result_type T _ = T; | 
| 243 | val v = Free(n, result_type T vs); | |
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 244 | val m = Const(match_name c, matchT (T, fastype_of rhs)); | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 245 | val k = big_lambdas vs rhs; | 
| 23152 | 246 | in | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 247 | (m`v`k, v, n::taken) | 
| 23152 | 248 | end | 
| 249 |   | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n)
 | |
| 250 | | _ => fixrec_err "pre_build: invalid pattern"; | |
| 251 | ||
| 252 | (* builds a monadic term for matching a function definition pattern *) | |
| 253 | (* returns (name, arity, matcher) *) | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 254 | fun building match_name pat rhs vs taken = | 
| 23152 | 255 | case pat of | 
| 26045 | 256 |     Const(@{const_name Rep_CFun}, _)$f$(v as Free(n,T)) =>
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 257 | building match_name f rhs (v::vs) taken | 
| 26045 | 258 |   | Const(@{const_name Rep_CFun}, _)$f$x =>
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 259 | let val (rhs', v, taken') = pre_build match_name x rhs [] taken; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 260 | in building match_name f rhs' (v::vs) taken' end | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 261 | | Free(_,_) => ((pat, length vs), big_lambdas vs rhs) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 262 | | Const(_,_) => ((pat, length vs), big_lambdas vs rhs) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 263 |   | _ => fixrec_err ("function is not declared as constant in theory: "
 | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 264 | ^ ML_Syntax.print_term pat); | 
| 23152 | 265 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 266 | fun strip_alls t = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 267 | if Logic.is_all t then strip_alls (snd (Logic.dest_all t)) else t; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 268 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 269 | fun match_eq match_name eq = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 270 | let | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 271 | 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 | 272 | in | 
| 30157 | 273 | building match_name lhs (mk_return rhs) [] (taken_names eq) | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 274 | end; | 
| 23152 | 275 | |
| 276 | (* returns the sum (using +++) of the terms in ms *) | |
| 277 | (* also applies "run" to the result! *) | |
| 278 | fun fatbar arity ms = | |
| 279 | let | |
| 30132 | 280 | fun LAM_Ts 0 t = ([], Term.fastype_of t) | 
| 281 | | LAM_Ts n (_ $ Abs(_,T,t)) = | |
| 282 | let val (Ts, U) = LAM_Ts (n-1) t in (T::Ts, U) end | |
| 283 | | LAM_Ts _ _ = fixrec_err "fatbar: internal error, not enough LAMs"; | |
| 23152 | 284 | fun unLAM 0 t = t | 
| 285 | | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t | |
| 286 | | unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs"; | |
| 30132 | 287 | fun reLAM ([], U) t = t | 
| 288 |       | reLAM (T::Ts, U) t = reLAM (Ts, T ->> U) (cabs_const(T,U)$Abs("",T,t));
 | |
| 289 | val msum = foldr1 mk_mplus (map (unLAM arity) ms); | |
| 290 | val (Ts, U) = LAM_Ts arity (hd ms) | |
| 23152 | 291 | in | 
| 30132 | 292 | reLAM (rev Ts, dest_maybeT U) (mk_run msum) | 
| 23152 | 293 | end; | 
| 294 | ||
| 295 | (* this is the pattern-matching compiler function *) | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 296 | fun compile_pats match_name eqs = | 
| 23152 | 297 | let | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 298 | val (((n::names),(a::arities)),mats) = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 299 | apfst ListPair.unzip (ListPair.unzip (map (match_eq match_name) eqs)); | 
| 23152 | 300 | val cname = if forall (fn x => n=x) names then n | 
| 301 | else fixrec_err "all equations in block must define the same function"; | |
| 302 | val arity = if forall (fn x => a=x) arities then a | |
| 303 | else fixrec_err "all equations in block must have the same arity"; | |
| 304 | val rhs = fatbar arity mats; | |
| 305 | in | |
| 30132 | 306 | mk_trp (cname === rhs) | 
| 23152 | 307 | end; | 
| 308 | ||
| 309 | (*************************************************************************) | |
| 310 | (********************** Proving associated theorems **********************) | |
| 311 | (*************************************************************************) | |
| 312 | ||
| 313 | (* proves a block of pattern matching equations as theorems, using unfold *) | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 314 | fun make_simps lthy (unfold_thm, eqns : (Attrib.binding * term) list) = | 
| 23152 | 315 | let | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 316 | val tacs = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 317 |       [rtac (unfold_thm RS @{thm ssubst_lhs}) 1,
 | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 318 | asm_simp_tac (local_simpset_of lthy) 1]; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 319 | fun prove_term t = Goal.prove lthy [] [] t (K (EVERY tacs)); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 320 | fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t); | 
| 23152 | 321 | in | 
| 322 | map prove_eqn eqns | |
| 323 | end; | |
| 324 | ||
| 325 | (*************************************************************************) | |
| 326 | (************************* Main fixrec function **************************) | |
| 327 | (*************************************************************************) | |
| 328 | ||
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 329 | local | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 330 | (* code adapted from HOL/Tools/primrec_package.ML *) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 331 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 332 | fun gen_fixrec | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 333 | (set_group : bool) | 
| 30485 | 334 | prep_spec | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 335 | (strict : bool) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 336 | raw_fixes | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 337 | raw_spec | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 338 | (lthy : local_theory) = | 
| 23152 | 339 | let | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 340 | val (fixes : ((binding * typ) * mixfix) list, | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 341 | spec : (Attrib.binding * term) list) = | 
| 30485 | 342 | fst (prep_spec raw_fixes raw_spec lthy); | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 343 | val chead_of_spec = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 344 | chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 345 | fun name_of (Free (n, _)) = n | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 346 |       | name_of t = fixrec_err ("unknown term");
 | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 347 | val all_names = map (name_of o chead_of_spec) spec; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 348 | val names = distinct (op =) all_names; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 349 | fun block_of_name n = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 350 | map_filter | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 351 | (fn (m,eq) => if m = n then SOME eq else NONE) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 352 | (all_names ~~ spec); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 353 | val blocks = map block_of_name names; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 354 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 355 | val matcher_tab = FixrecMatchData.get (ProofContext.theory_of lthy); | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 356 | fun match_name c = | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 357 | case Symtab.lookup matcher_tab c of SOME m => m | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 358 |         | NONE => fixrec_err ("unknown pattern constructor: " ^ c);
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 359 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 360 | val matches = map (compile_pats match_name) (map (map snd) blocks); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 361 | val spec' = map (pair Attrib.empty_binding) matches; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 362 | val (lthy', cnames, fixdef_thms, unfold_thms) = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 363 | add_fixdefs fixes spec' lthy; | 
| 23152 | 364 | in | 
| 365 | if strict then let (* only prove simp rules if strict = true *) | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 366 | val simps : (Attrib.binding * thm) list list = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 367 | map (make_simps lthy') (unfold_thms ~~ blocks); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 368 | fun mk_bind n : Attrib.binding = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 369 | (Binding.name (n ^ "_simps"), | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 370 | [Attrib.internal (K Simplifier.simp_add)]); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 371 | val simps1 : (Attrib.binding * thm list) list = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 372 | map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 373 | val simps2 : (Attrib.binding * thm list) list = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 374 | map (apsnd (fn thm => [thm])) (List.concat simps); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 375 | val (_, lthy'') = lthy' | 
| 31177 | 376 | |> fold_map (LocalTheory.note Thm.generatedK) (simps1 @ simps2); | 
| 23152 | 377 | in | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 378 | lthy'' | 
| 23152 | 379 | end | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 380 | else lthy' | 
| 23152 | 381 | end; | 
| 382 | ||
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 383 | in | 
| 23152 | 384 | |
| 30485 | 385 | val add_fixrec = gen_fixrec false Specification.check_spec; | 
| 386 | val add_fixrec_cmd = gen_fixrec true Specification.read_spec; | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 387 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 388 | end; (* local *) | 
| 23152 | 389 | |
| 390 | (*************************************************************************) | |
| 391 | (******************************** Fixpat *********************************) | |
| 392 | (*************************************************************************) | |
| 393 | ||
| 394 | fun fix_pat thy t = | |
| 395 | let | |
| 396 | val T = fastype_of t; | |
| 397 |     val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
 | |
| 398 | val cname = case chead_of t of Const(c,_) => c | _ => | |
| 399 | fixrec_err "function is not declared as constant in theory"; | |
| 26343 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 400 | val unfold_thm = PureThy.get_thm thy (cname^"_unfold"); | 
| 23152 | 401 | val simp = Goal.prove_global thy [] [] eq | 
| 402 | (fn _ => EVERY [stac unfold_thm 1, simp_tac (simpset_of thy) 1]); | |
| 403 | in simp end; | |
| 404 | ||
| 405 | fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy = | |
| 406 | let | |
| 407 | val atts = map (prep_attrib thy) srcs; | |
| 408 | val ts = map (prep_term thy) strings; | |
| 409 | val simps = map (fix_pat thy) ts; | |
| 410 | in | |
| 29585 | 411 | (snd o PureThy.add_thmss [((name, simps), atts)]) thy | 
| 23152 | 412 | end; | 
| 413 | ||
| 30485 | 414 | val add_fixpat = gen_add_fixpat Sign.cert_term (K I); | 
| 415 | val add_fixpat_cmd = gen_add_fixpat Syntax.read_term_global Attrib.attribute; | |
| 23152 | 416 | |
| 417 | ||
| 418 | (*************************************************************************) | |
| 419 | (******************************** Parsers ********************************) | |
| 420 | (*************************************************************************) | |
| 421 | ||
| 422 | local structure P = OuterParse and K = OuterKeyword in | |
| 423 | ||
| 30485 | 424 | val _ = OuterSyntax.local_theory "fixrec" "define recursive functions (HOLCF)" K.thy_decl | 
| 425 | ((P.opt_keyword "permissive" >> not) -- P.fixes -- SpecParse.where_alt_specs | |
| 426 | >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs)); | |
| 23152 | 427 | |
| 30485 | 428 | val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl | 
| 429 | (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd)); | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 430 | |
| 30485 | 431 | end; | 
| 23152 | 432 | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 433 | val setup = FixrecMatchData.init; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 434 | |
| 24867 | 435 | end; |