| author | wenzelm | 
| Mon, 26 Apr 2010 11:20:18 +0200 | |
| changeset 36345 | 3cbce59ed78d | 
| parent 35903 | 0b43ff2d2e91 | 
| child 36628 | 1a251f69e96b | 
| permissions | -rw-r--r-- | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 1 | (* Title: 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 | 
| 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 | 
| 33442 | 16 | val fixrec_simp_add: attribute | 
| 17 | val fixrec_simp_del: attribute | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 18 | val fixrec_simp_tac: Proof.context -> int -> tactic | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 19 | val setup: theory -> theory | 
| 23152 | 20 | end; | 
| 21 | ||
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 22 | structure Fixrec :> FIXREC = | 
| 23152 | 23 | struct | 
| 24 | ||
| 35527 | 25 | open HOLCF_Library; | 
| 26 | ||
| 27 | infixr 6 ->>; | |
| 28 | infix -->>; | |
| 29 | infix 9 `; | |
| 30 | ||
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 31 | 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 | 32 | val def_cont_fix_ind = @{thm def_cont_fix_ind};
 | 
| 23152 | 33 | |
| 34 | fun fixrec_err s = error ("fixrec definition error:\n" ^ s);
 | |
| 35 | fun fixrec_eq_err thy s eq = | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26343diff
changeset | 36 | fixrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); | 
| 23152 | 37 | |
| 30132 | 38 | (*************************************************************************) | 
| 39 | (***************************** building types ****************************) | |
| 40 | (*************************************************************************) | |
| 41 | ||
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 42 | local | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 43 | |
| 35525 | 44 | 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 | 45 |   | binder_cfun (Type(@{type_name "fun"},[T, U])) = T :: binder_cfun U
 | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 46 | | binder_cfun _ = []; | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 47 | |
| 35525 | 48 | fun body_cfun (Type(@{type_name cfun},[T, U])) = body_cfun U
 | 
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 49 |   | body_cfun (Type(@{type_name "fun"},[T, U])) = body_cfun U
 | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 50 | | body_cfun T = T; | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 51 | |
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 52 | fun strip_cfun T : typ list * typ = | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 53 | (binder_cfun T, body_cfun T); | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 54 | |
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 55 | in | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 56 | |
| 35527 | 57 | fun matcherT (T, U) = | 
| 58 | body_cfun T ->> (binder_cfun T -->> U) ->> U; | |
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 59 | |
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 60 | end | 
| 30132 | 61 | |
| 62 | (*************************************************************************) | |
| 63 | (***************************** building terms ****************************) | |
| 64 | (*************************************************************************) | |
| 23152 | 65 | |
| 66 | val mk_trp = HOLogic.mk_Trueprop; | |
| 67 | ||
| 68 | (* splits a cterm into the right and lefthand sides of equality *) | |
| 69 | fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t); | |
| 70 | ||
| 71 | (* similar to Thm.head_of, but for continuous application *) | |
| 26045 | 72 | fun chead_of (Const(@{const_name Rep_CFun},_)$f$t) = chead_of f
 | 
| 23152 | 73 | | chead_of u = u; | 
| 74 | ||
| 30132 | 75 | infix 0 ==; val (op ==) = Logic.mk_equals; | 
| 76 | infix 1 ===; val (op ===) = HOLogic.mk_eq; | |
| 77 | ||
| 78 | fun mk_mplus (t, u) = | |
| 79 | let val mT = Term.fastype_of t | |
| 80 |   in Const(@{const_name Fixrec.mplus}, mT ->> mT ->> mT) ` t ` u end;
 | |
| 81 | ||
| 82 | fun mk_run t = | |
| 83 | let val mT = Term.fastype_of t | |
| 35527 | 84 | val T = dest_matchT mT | 
| 30132 | 85 |   in Const(@{const_name Fixrec.run}, mT ->> T) ` t end;
 | 
| 86 | ||
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 87 | |
| 23152 | 88 | (*************************************************************************) | 
| 89 | (************* fixed-point definitions and unfolding theorems ************) | |
| 90 | (*************************************************************************) | |
| 91 | ||
| 33519 | 92 | structure FixrecUnfoldData = Generic_Data | 
| 93 | ( | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 94 | type T = thm Symtab.table; | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 95 | val empty = Symtab.empty; | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 96 | val extend = I; | 
| 33519 | 97 | fun merge data : T = Symtab.merge (K true) data; | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 98 | ); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 99 | |
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 100 | local | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 101 | |
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 102 | fun name_of (Const (n, T)) = n | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 103 | | name_of (Free (n, T)) = n | 
| 35770 
a57ab2c01369
fixes to allow using fixrec_simp inside a locale, with test in ex/Fixrec_ex.thy
 huffman parents: 
35769diff
changeset | 104 |   | 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 | 105 | |
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 106 | val lhs_name = | 
| 35770 
a57ab2c01369
fixes to allow using fixrec_simp inside a locale, with test in ex/Fixrec_ex.thy
 huffman parents: 
35769diff
changeset | 107 | 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 | 108 | |
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 109 | in | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 110 | |
| 33442 | 111 | val add_unfold : attribute = | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 112 | Thm.declaration_attribute | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 113 | (fn th => FixrecUnfoldData.map (Symtab.insert (K true) (lhs_name th, th))); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 114 | |
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 115 | end | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 116 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 117 | fun add_fixdefs | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 118 | (fixes : ((binding * typ) * mixfix) list) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 119 | (spec : (Attrib.binding * term) list) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 120 | (lthy : local_theory) = | 
| 23152 | 121 | let | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 122 | 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 | 123 | 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 | 124 | val all_names = space_implode "_" names; | 
| 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 | 125 | 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 | 126 | 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 | 127 | val fixpoint = mk_fix (mk_cabs functional); | 
| 23152 | 128 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 129 | val cont_thm = | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 130 | Goal.prove lthy [] [] (mk_trp (mk_cont functional)) | 
| 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 | 131 | (K (simp_tac (simpset_of lthy) 1)); | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 132 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 133 | 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 | 134 | let val b = Long_Name.base_name n | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 135 | in ((Binding.name (b^"_def"), []), r) end | 
| 23152 | 136 | | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form"; | 
| 137 | fun defs [] _ = [] | |
| 138 | | 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 | 139 | | 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 | 140 | val fixdefs = defs lhss fixpoint; | 
| 35772 | 141 | val (fixdef_thms : (term * (string * thm)) list, lthy) = lthy | 
| 33766 
c679f05600cd
adapted Local_Theory.define -- eliminated odd thm kind;
 wenzelm parents: 
33726diff
changeset | 142 | |> fold_map Local_Theory.define (map (apfst fst) fixes ~~ fixdefs); | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 143 |     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 | 144 | 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 | 145 |     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 | 146 | 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 | 147 | 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 | 148 | |> Drule.instantiate' [] [SOME (Thm.cterm_of thy predicate)] | 
| 35624 | 149 |       |> 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 | 150 | val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) | 
| 35772 | 151 |       |> Local_Defs.unfold lthy @{thms split_conv};
 | 
| 23152 | 152 | fun unfolds [] thm = [] | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 153 | | unfolds (n::[]) thm = [(n, thm)] | 
| 23152 | 154 | | unfolds (n::ns) thm = let | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 155 |           val thmL = thm RS @{thm Pair_eqD1};
 | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 156 |           val thmR = thm RS @{thm Pair_eqD2};
 | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 157 | in (n, thmL) :: unfolds ns thmR end; | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31023diff
changeset | 158 | val unfold_thms = unfolds names tuple_unfold_thm; | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 159 | val induct_note : Attrib.binding * Thm.thm list = | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 160 | let | 
| 35769 | 161 | 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 | 162 | in | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 163 | ((thm_name, []), [tuple_induct_thm]) | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 164 | end; | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 165 | 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 | 166 | let | 
| 35769 | 167 | val thm_name = Binding.qualify true name (Binding.name "unfold"); | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 168 | val src = Attrib.internal (K add_unfold); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 169 | in | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 170 | ((thm_name, [src]), [thm]) | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 171 | end; | 
| 35772 | 172 | val (thmss, lthy) = lthy | 
| 33671 | 173 | |> fold_map Local_Theory.note (induct_note :: map unfold_note unfold_thms); | 
| 23152 | 174 | in | 
| 35772 | 175 | (lthy, names, fixdef_thms, map snd unfold_thms) | 
| 23152 | 176 | end; | 
| 177 | ||
| 178 | (*************************************************************************) | |
| 179 | (*********** monadic notation and pattern matching compilation ***********) | |
| 180 | (*************************************************************************) | |
| 181 | ||
| 33522 | 182 | structure FixrecMatchData = Theory_Data | 
| 183 | ( | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 184 | type T = string Symtab.table; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 185 | val empty = Symtab.empty; | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 186 | val extend = I; | 
| 33522 | 187 | fun merge data = Symtab.merge (K true) data; | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 188 | ); | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 189 | |
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 190 | (* associate match functions with pattern constants *) | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 191 | fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms); | 
| 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 192 | |
| 30157 | 193 | fun taken_names (t : term) : bstring list = | 
| 194 | let | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 195 | fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs | 
| 30157 | 196 | | taken (Free(a,_) , bs) = insert (op =) a bs | 
| 197 | | taken (f $ u , bs) = taken (f, taken (u, bs)) | |
| 198 | | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs) | |
| 199 | | taken (_ , bs) = bs; | |
| 200 | in | |
| 201 | taken (t, []) | |
| 202 | end; | |
| 23152 | 203 | |
| 204 | (* builds a monadic term for matching a constructor pattern *) | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 205 | fun pre_build match_name pat rhs vs taken = | 
| 23152 | 206 | case pat of | 
| 26045 | 207 |     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 | 208 | pre_build match_name f rhs (v::vs) taken | 
| 26045 | 209 |   | Const(@{const_name Rep_CFun},_)$f$x =>
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 210 | 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 | 211 | in pre_build match_name f rhs' (v::vs) taken' end | 
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 212 | | f$(v as Free(n,T)) => | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 213 | pre_build match_name f rhs (v::vs) taken | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 214 | | f$x => | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 215 | let val (rhs', v, taken') = pre_build match_name x rhs [] taken; | 
| 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 216 | in pre_build match_name f rhs' (v::vs) taken' end | 
| 23152 | 217 | | Const(c,T) => | 
| 218 | let | |
| 219 | val n = Name.variant taken "v"; | |
| 35525 | 220 |         fun result_type (Type(@{type_name cfun},[_,T])) (x::xs) = result_type T xs
 | 
| 33401 
fc43fa403a69
add fixrec support for HOL pair constructor patterns
 huffman parents: 
33004diff
changeset | 221 |           | result_type (Type (@{type_name "fun"},[_,T])) (x::xs) = result_type T xs
 | 
| 23152 | 222 | | result_type T _ = T; | 
| 223 | val v = Free(n, result_type T vs); | |
| 35527 | 224 | val m = Const(match_name c, matcherT (T, fastype_of rhs)); | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 225 | val k = big_lambdas vs rhs; | 
| 23152 | 226 | in | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30485diff
changeset | 227 | (m`v`k, v, n::taken) | 
| 23152 | 228 | end | 
| 229 |   | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n)
 | |
| 230 | | _ => fixrec_err "pre_build: invalid pattern"; | |
| 231 | ||
| 232 | (* builds a monadic term for matching a function definition pattern *) | |
| 233 | (* returns (name, arity, matcher) *) | |
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 234 | fun building match_name pat rhs vs taken = | 
| 23152 | 235 | case pat of | 
| 26045 | 236 |     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 | 237 | building match_name f rhs (v::vs) taken | 
| 26045 | 238 |   | Const(@{const_name Rep_CFun}, _)$f$x =>
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 239 | 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 | 240 | 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 | 241 | | Free(_,_) => ((pat, length vs), big_lambdas vs rhs) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 242 | | Const(_,_) => ((pat, length vs), big_lambdas vs rhs) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 243 |   | _ => 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 | 244 | ^ ML_Syntax.print_term pat); | 
| 23152 | 245 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 246 | fun strip_alls t = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 247 | 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 | 248 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 249 | fun match_eq match_name eq = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 250 | let | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 251 | 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 | 252 | in | 
| 30157 | 253 | 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 | 254 | end; | 
| 23152 | 255 | |
| 256 | (* returns the sum (using +++) of the terms in ms *) | |
| 257 | (* also applies "run" to the result! *) | |
| 258 | fun fatbar arity ms = | |
| 259 | let | |
| 30132 | 260 | fun LAM_Ts 0 t = ([], Term.fastype_of t) | 
| 261 | | LAM_Ts n (_ $ Abs(_,T,t)) = | |
| 262 | let val (Ts, U) = LAM_Ts (n-1) t in (T::Ts, U) end | |
| 263 | | LAM_Ts _ _ = fixrec_err "fatbar: internal error, not enough LAMs"; | |
| 23152 | 264 | fun unLAM 0 t = t | 
| 265 | | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t | |
| 266 | | unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs"; | |
| 30132 | 267 | fun reLAM ([], U) t = t | 
| 268 |       | reLAM (T::Ts, U) t = reLAM (Ts, T ->> U) (cabs_const(T,U)$Abs("",T,t));
 | |
| 269 | val msum = foldr1 mk_mplus (map (unLAM arity) ms); | |
| 270 | val (Ts, U) = LAM_Ts arity (hd ms) | |
| 23152 | 271 | in | 
| 35527 | 272 | reLAM (rev Ts, dest_matchT U) (mk_run msum) | 
| 23152 | 273 | end; | 
| 274 | ||
| 275 | (* this is the pattern-matching compiler function *) | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 276 | fun compile_pats match_name eqs = | 
| 23152 | 277 | let | 
| 35903 | 278 | val ((names, arities), mats) = | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 279 | apfst ListPair.unzip (ListPair.unzip (map (match_eq match_name) eqs)); | 
| 35903 | 280 | val cname = | 
| 281 | case distinct (op =) names of | |
| 282 | [n] => n | |
| 283 | | _ => fixrec_err "all equations in block must define the same function"; | |
| 284 | val arity = | |
| 285 | case distinct (op =) arities of | |
| 286 | [a] => a | |
| 287 | | _ => fixrec_err "all equations in block must have the same arity"; | |
| 23152 | 288 | val rhs = fatbar arity mats; | 
| 289 | in | |
| 30132 | 290 | mk_trp (cname === rhs) | 
| 23152 | 291 | end; | 
| 292 | ||
| 293 | (*************************************************************************) | |
| 294 | (********************** Proving associated theorems **********************) | |
| 295 | (*************************************************************************) | |
| 296 | ||
| 33519 | 297 | structure FixrecSimpData = Generic_Data | 
| 298 | ( | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 299 | type T = simpset; | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 300 | val empty = | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 301 | HOL_basic_ss | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 302 | addsimps simp_thms | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 303 |       addsimps [@{thm beta_cfun}]
 | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 304 |       addsimprocs [@{simproc cont_proc}];
 | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 305 | val extend = I; | 
| 33519 | 306 | val merge = merge_ss; | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 307 | ); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 308 | |
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 309 | fun fixrec_simp_tac ctxt = | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 310 | let | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 311 | val tab = FixrecUnfoldData.get (Context.Proof ctxt); | 
| 35780 
98fd7910f70a
use Simplifier.context to avoid 'no proof context in simpset' errors from fixrec_simp after theory merge
 huffman parents: 
35779diff
changeset | 312 | val ss = Simplifier.context ctxt (FixrecSimpData.get (Context.Proof ctxt)); | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 313 | fun concl t = | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 314 | if Logic.is_all t then concl (snd (Logic.dest_all t)) | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 315 | else HOLogic.dest_Trueprop (Logic.strip_imp_concl t); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 316 | fun tac (t, i) = | 
| 33430 | 317 | let | 
| 35903 | 318 | val (c, T) = | 
| 319 | (dest_Const o head_of o chead_of o fst o HOLogic.dest_eq o concl) t; | |
| 33430 | 320 | val unfold_thm = the (Symtab.lookup tab c); | 
| 321 |         val rule = unfold_thm RS @{thm ssubst_lhs};
 | |
| 322 | in | |
| 323 | CHANGED (rtac rule i THEN asm_simp_tac ss i) | |
| 324 | end | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 325 | in | 
| 33505 | 326 | SUBGOAL (fn ti => the_default no_tac (try tac ti)) | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 327 | end; | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 328 | |
| 33442 | 329 | val fixrec_simp_add : attribute = | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 330 | Thm.declaration_attribute | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 331 | (fn th => FixrecSimpData.map (fn ss => ss addsimps [th])); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 332 | |
| 33442 | 333 | val fixrec_simp_del : attribute = | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 334 | Thm.declaration_attribute | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 335 | (fn th => FixrecSimpData.map (fn ss => ss delsimps [th])); | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 336 | |
| 23152 | 337 | (* 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 | 338 | fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = | 
| 23152 | 339 | let | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 340 | val tacs = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 341 |       [rtac (unfold_thm RS @{thm ssubst_lhs}) 1,
 | 
| 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 | 342 | asm_simp_tac (simpset_of ctxt) 1]; | 
| 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
31740diff
changeset | 343 | fun prove_term t = Goal.prove ctxt [] [] t (K (EVERY tacs)); | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 344 | fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t); | 
| 23152 | 345 | in | 
| 346 | map prove_eqn eqns | |
| 347 | end; | |
| 348 | ||
| 349 | (*************************************************************************) | |
| 350 | (************************* Main fixrec function **************************) | |
| 351 | (*************************************************************************) | |
| 352 | ||
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 353 | local | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 354 | (* code adapted from HOL/Tools/primrec.ML *) | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 355 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 356 | fun gen_fixrec | 
| 30485 | 357 | prep_spec | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 358 | (strict : bool) | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 359 | raw_fixes | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 360 | raw_spec | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 361 | (lthy : local_theory) = | 
| 23152 | 362 | let | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 363 | val (fixes : ((binding * typ) * mixfix) list, | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 364 | spec : (Attrib.binding * term) list) = | 
| 30485 | 365 | fst (prep_spec raw_fixes raw_spec lthy); | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 366 | val chead_of_spec = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 367 | 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 | 368 | fun name_of (Free (n, _)) = n | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 369 |       | name_of t = fixrec_err ("unknown term");
 | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 370 | 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 | 371 | val names = distinct (op =) all_names; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 372 | fun block_of_name n = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 373 | map_filter | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 374 | (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 | 375 | (all_names ~~ spec); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 376 | val blocks = map block_of_name names; | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 377 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 378 | val matcher_tab = FixrecMatchData.get (ProofContext.theory_of lthy); | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 379 | fun match_name c = | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 380 | 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 | 381 |         | NONE => fixrec_err ("unknown pattern constructor: " ^ c);
 | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 382 | |
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 383 | 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 | 384 | val spec' = map (pair Attrib.empty_binding) matches; | 
| 35772 | 385 | val (lthy, cnames, fixdef_thms, unfold_thms) = | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 386 | add_fixdefs fixes spec' lthy; | 
| 23152 | 387 | in | 
| 388 | 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 | 389 | val simps : (Attrib.binding * thm) list list = | 
| 35772 | 390 | map (make_simps lthy) (unfold_thms ~~ blocks); | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 391 | fun mk_bind n : Attrib.binding = | 
| 35769 | 392 | (Binding.qualify true n (Binding.name "simps"), | 
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 393 | [Attrib.internal (K Simplifier.simp_add)]); | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 394 | val simps1 : (Attrib.binding * thm list) list = | 
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 395 | 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 | 396 | val simps2 : (Attrib.binding * thm list) list = | 
| 32952 | 397 | map (apsnd (fn thm => [thm])) (flat simps); | 
| 35772 | 398 | val (_, lthy) = lthy | 
| 33671 | 399 | |> fold_map Local_Theory.note (simps1 @ simps2); | 
| 23152 | 400 | in | 
| 35772 | 401 | lthy | 
| 23152 | 402 | end | 
| 35772 | 403 | else lthy | 
| 23152 | 404 | end; | 
| 405 | ||
| 30158 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 406 | in | 
| 23152 | 407 | |
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 408 | val add_fixrec = gen_fixrec Specification.check_spec; | 
| 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 409 | 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 | 410 | |
| 
83c50c62cf23
fixrec package uses new-style syntax and local-theory interface
 huffman parents: 
30157diff
changeset | 411 | end; (* local *) | 
| 23152 | 412 | |
| 413 | (*************************************************************************) | |
| 414 | (******************************** Fixpat *********************************) | |
| 415 | (*************************************************************************) | |
| 416 | ||
| 417 | fun fix_pat thy t = | |
| 418 | let | |
| 419 | val T = fastype_of t; | |
| 420 |     val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
 | |
| 421 | val cname = case chead_of t of Const(c,_) => c | _ => | |
| 422 | fixrec_err "function is not declared as constant in theory"; | |
| 35769 | 423 | val unfold_thm = PureThy.get_thm thy (cname^".unfold"); | 
| 23152 | 424 | val simp = Goal.prove_global thy [] [] eq | 
| 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 | 425 | (fn _ => EVERY [stac unfold_thm 1, simp_tac (global_simpset_of thy) 1]); | 
| 23152 | 426 | in simp end; | 
| 427 | ||
| 428 | fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy = | |
| 429 | let | |
| 35779 | 430 | val _ = legacy_feature "\"fixpat\"; prefer \"fixrec_simp\" method instead"; | 
| 23152 | 431 | val atts = map (prep_attrib thy) srcs; | 
| 432 | val ts = map (prep_term thy) strings; | |
| 433 | val simps = map (fix_pat thy) ts; | |
| 434 | in | |
| 29585 | 435 | (snd o PureThy.add_thmss [((name, simps), atts)]) thy | 
| 23152 | 436 | end; | 
| 437 | ||
| 30485 | 438 | val add_fixpat = gen_add_fixpat Sign.cert_term (K I); | 
| 439 | val add_fixpat_cmd = gen_add_fixpat Syntax.read_term_global Attrib.attribute; | |
| 23152 | 440 | |
| 441 | ||
| 442 | (*************************************************************************) | |
| 443 | (******************************** Parsers ********************************) | |
| 444 | (*************************************************************************) | |
| 445 | ||
| 446 | local structure P = OuterParse and K = OuterKeyword in | |
| 447 | ||
| 30485 | 448 | val _ = OuterSyntax.local_theory "fixrec" "define recursive functions (HOLCF)" K.thy_decl | 
| 449 | ((P.opt_keyword "permissive" >> not) -- P.fixes -- SpecParse.where_alt_specs | |
| 450 | >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs)); | |
| 23152 | 451 | |
| 30485 | 452 | val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl | 
| 453 | (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd)); | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 454 | |
| 30485 | 455 | end; | 
| 23152 | 456 | |
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 457 | val setup = | 
| 33522 | 458 |   Attrib.setup @{binding fixrec_simp}
 | 
| 33427 | 459 | (Attrib.add_del fixrec_simp_add fixrec_simp_del) | 
| 33425 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 460 | "declaration of fixrec simp rule" | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 461 |   #> Method.setup @{binding fixrec_simp}
 | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 462 | (Scan.succeed (SIMPLE_METHOD' o fixrec_simp_tac)) | 
| 
7e4f3c66190d
add fixrec_simp attribute and method (eventually to replace fixpat)
 huffman parents: 
33401diff
changeset | 463 | "pattern prover for fixrec constants"; | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29585diff
changeset | 464 | |
| 24867 | 465 | end; |