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