| author | blanchet | 
| Tue, 24 May 2011 10:01:03 +0200 | |
| changeset 42968 | 74415622d293 | 
| parent 42949 | 618adb3584e5 | 
| child 43080 | 73a1d6a7ef1d | 
| permissions | -rw-r--r-- | 
| 40107 | 1  | 
(* Title: HOL/Tools/Function/partial_function.ML  | 
2  | 
Author: Alexander Krauss, TU Muenchen  | 
|
3  | 
||
4  | 
Partial function definitions based on least fixed points in ccpos.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature PARTIAL_FUNCTION =  | 
|
8  | 
sig  | 
|
9  | 
val setup: theory -> theory  | 
|
| 
42949
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
42495 
diff
changeset
 | 
10  | 
val init: string -> term -> term -> thm -> declaration  | 
| 40107 | 11  | 
|
12  | 
val add_partial_function: string -> (binding * typ option * mixfix) list ->  | 
|
13  | 
Attrib.binding * term -> local_theory -> local_theory  | 
|
14  | 
||
15  | 
val add_partial_function_cmd: string -> (binding * string option * mixfix) list ->  | 
|
16  | 
Attrib.binding * string -> local_theory -> local_theory  | 
|
17  | 
end;  | 
|
18  | 
||
19  | 
||
20  | 
structure Partial_Function: PARTIAL_FUNCTION =  | 
|
21  | 
struct  | 
|
22  | 
||
23  | 
(*** Context Data ***)  | 
|
24  | 
||
25  | 
structure Modes = Generic_Data  | 
|
26  | 
(  | 
|
27  | 
type T = ((term * term) * thm) Symtab.table;  | 
|
28  | 
val empty = Symtab.empty;  | 
|
29  | 
val extend = I;  | 
|
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41117 
diff
changeset
 | 
30  | 
fun merge data = Symtab.merge (K true) data;  | 
| 40107 | 31  | 
)  | 
32  | 
||
| 
42949
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
42495 
diff
changeset
 | 
33  | 
fun init mode fixp mono fixp_eq phi =  | 
| 40107 | 34  | 
let  | 
35  | 
val term = Morphism.term phi;  | 
|
36  | 
val data' = ((term fixp, term mono), Morphism.thm phi fixp_eq);  | 
|
37  | 
in  | 
|
| 
42949
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
42495 
diff
changeset
 | 
38  | 
Modes.map (Symtab.update (mode, data'))  | 
| 40107 | 39  | 
end  | 
40  | 
||
41  | 
val known_modes = Symtab.keys o Modes.get o Context.Proof;  | 
|
42  | 
val lookup_mode = Symtab.lookup o Modes.get o Context.Proof;  | 
|
43  | 
||
44  | 
||
45  | 
structure Mono_Rules = Named_Thms  | 
|
46  | 
(  | 
|
47  | 
val name = "partial_function_mono";  | 
|
48  | 
val description = "monotonicity rules for partial function definitions";  | 
|
49  | 
);  | 
|
50  | 
||
51  | 
||
52  | 
(*** Automated monotonicity proofs ***)  | 
|
53  | 
||
54  | 
fun strip_cases ctac = ctac #> Seq.map snd;  | 
|
55  | 
||
56  | 
(*rewrite conclusion with k-th assumtion*)  | 
|
57  | 
fun rewrite_with_asm_tac ctxt k =  | 
|
58  | 
  Subgoal.FOCUS (fn {context=ctxt', prems, ...} =>
 | 
|
59  | 
Local_Defs.unfold_tac ctxt' [nth prems k]) ctxt;  | 
|
60  | 
||
61  | 
fun dest_case thy t =  | 
|
62  | 
case strip_comb t of  | 
|
63  | 
(Const (case_comb, _), args) =>  | 
|
64  | 
(case Datatype.info_of_case thy case_comb of  | 
|
65  | 
NONE => NONE  | 
|
66  | 
       | SOME {case_rewrites, ...} =>
 | 
|
67  | 
let  | 
|
68  | 
val lhs = prop_of (hd case_rewrites)  | 
|
69  | 
|> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst;  | 
|
70  | 
val arity = length (snd (strip_comb lhs));  | 
|
71  | 
val conv = funpow (length args - arity) Conv.fun_conv  | 
|
72  | 
(Conv.rewrs_conv (map mk_meta_eq case_rewrites));  | 
|
73  | 
in  | 
|
74  | 
SOME (nth args (arity - 1), conv)  | 
|
75  | 
end)  | 
|
76  | 
| _ => NONE;  | 
|
77  | 
||
78  | 
(*split on case expressions*)  | 
|
79  | 
val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context=ctxt, ...} =>
 | 
|
80  | 
SUBGOAL (fn (t, i) => case t of  | 
|
81  | 
_ $ (_ $ Abs (_, _, body)) =>  | 
|
| 42361 | 82  | 
(case dest_case (Proof_Context.theory_of ctxt) body of  | 
| 40107 | 83  | 
NONE => no_tac  | 
84  | 
| SOME (arg, conv) =>  | 
|
85  | 
let open Conv in  | 
|
| 
42083
 
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
 
wenzelm 
parents: 
41472 
diff
changeset
 | 
86  | 
if Term.is_open arg then no_tac  | 
| 40107 | 87  | 
else ((DETERM o strip_cases o Induct.cases_tac ctxt false [[SOME arg]] NONE [])  | 
88  | 
THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0)  | 
|
89  | 
                THEN_ALL_NEW etac @{thm thin_rl}
 | 
|
90  | 
THEN_ALL_NEW (CONVERSION  | 
|
91  | 
(params_conv ~1 (fn ctxt' =>  | 
|
92  | 
arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i  | 
|
93  | 
end)  | 
|
94  | 
| _ => no_tac) 1);  | 
|
95  | 
||
96  | 
(*monotonicity proof: apply rules + split case expressions*)  | 
|
97  | 
fun mono_tac ctxt =  | 
|
98  | 
  K (Local_Defs.unfold_tac ctxt [@{thm curry_def}])
 | 
|
99  | 
THEN' (TRY o REPEAT_ALL_NEW  | 
|
100  | 
(resolve_tac (Mono_Rules.get ctxt)  | 
|
101  | 
ORELSE' split_cases_tac ctxt));  | 
|
102  | 
||
103  | 
||
104  | 
(*** Auxiliary functions ***)  | 
|
105  | 
||
106  | 
(*positional instantiation with computed type substitution.  | 
|
107  | 
internal version of attribute "[of s t u]".*)  | 
|
108  | 
fun cterm_instantiate' cts thm =  | 
|
109  | 
let  | 
|
110  | 
val thy = Thm.theory_of_thm thm;  | 
|
111  | 
val vs = rev (Term.add_vars (prop_of thm) [])  | 
|
112  | 
|> map (Thm.cterm_of thy o Var);  | 
|
113  | 
in  | 
|
114  | 
cterm_instantiate (zip_options vs cts) thm  | 
|
115  | 
end;  | 
|
116  | 
||
117  | 
(*Returns t $ u, but instantiates the type of t to make the  | 
|
118  | 
application type correct*)  | 
|
119  | 
fun apply_inst ctxt t u =  | 
|
120  | 
let  | 
|
| 42361 | 121  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 40107 | 122  | 
val T = domain_type (fastype_of t);  | 
123  | 
val T' = fastype_of u;  | 
|
| 42388 | 124  | 
val subst = Sign.typ_match thy (T, T') Vartab.empty  | 
| 40107 | 125  | 
      handle Type.TYPE_MATCH => raise TYPE ("apply_inst", [T, T'], [t, u])
 | 
126  | 
in  | 
|
127  | 
map_types (Envir.norm_type subst) t $ u  | 
|
128  | 
end;  | 
|
129  | 
||
130  | 
fun head_conv cv ct =  | 
|
131  | 
if can Thm.dest_comb ct then Conv.fun_conv (head_conv cv) ct else cv ct;  | 
|
132  | 
||
133  | 
||
134  | 
(*** currying transformation ***)  | 
|
135  | 
||
136  | 
fun curry_const (A, B, C) =  | 
|
137  | 
  Const (@{const_name Product_Type.curry},
 | 
|
138  | 
[HOLogic.mk_prodT (A, B) --> C, A, B] ---> C);  | 
|
139  | 
||
140  | 
fun mk_curry f =  | 
|
141  | 
case fastype_of f of  | 
|
142  | 
    Type ("fun", [Type (_, [S, T]), U]) =>
 | 
|
143  | 
curry_const (S, T, U) $ f  | 
|
144  | 
  | T => raise TYPE ("mk_curry", [T], [f]);
 | 
|
145  | 
||
146  | 
(* iterated versions. Nonstandard left-nested tuples arise naturally  | 
|
147  | 
from "split o split o split"*)  | 
|
148  | 
fun curry_n arity = funpow (arity - 1) mk_curry;  | 
|
149  | 
fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split;  | 
|
150  | 
||
151  | 
val curry_uncurry_ss = HOL_basic_ss addsimps  | 
|
152  | 
  [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}]
 | 
|
153  | 
||
154  | 
||
155  | 
(*** partial_function definition ***)  | 
|
156  | 
||
157  | 
fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy =  | 
|
158  | 
let  | 
|
159  | 
val ((fixp, mono), fixp_eq) = the (lookup_mode lthy mode)  | 
|
160  | 
handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".",  | 
|
161  | 
"Known modes are " ^ commas_quote (known_modes lthy) ^ "."]);  | 
|
162  | 
||
163  | 
val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;  | 
|
| 
42495
 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 
wenzelm 
parents: 
42388 
diff
changeset
 | 
164  | 
val ((_, plain_eqn), _) = Variable.focus eqn lthy;  | 
| 40107 | 165  | 
|
166  | 
val ((f_binding, fT), mixfix) = the_single fixes;  | 
|
167  | 
val fname = Binding.name_of f_binding;  | 
|
168  | 
||
| 42361 | 169  | 
val cert = cterm_of (Proof_Context.theory_of lthy);  | 
| 40107 | 170  | 
val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn);  | 
171  | 
val (head, args) = strip_comb lhs;  | 
|
172  | 
val F = fold_rev lambda (head :: args) rhs;  | 
|
173  | 
||
174  | 
val arity = length args;  | 
|
175  | 
val (aTs, bTs) = chop arity (binder_types fT);  | 
|
176  | 
||
177  | 
val tupleT = foldl1 HOLogic.mk_prodT aTs;  | 
|
178  | 
val fT_uc = tupleT :: bTs ---> body_type fT;  | 
|
179  | 
val f_uc = Var ((fname, 0), fT_uc);  | 
|
180  | 
    val x_uc = Var (("x", 0), tupleT);
 | 
|
181  | 
val uncurry = lambda head (uncurry_n arity head);  | 
|
182  | 
val curry = lambda f_uc (curry_n arity f_uc);  | 
|
183  | 
||
184  | 
val F_uc =  | 
|
185  | 
lambda f_uc (uncurry_n arity (F $ curry_n arity f_uc));  | 
|
186  | 
||
187  | 
val mono_goal = apply_inst lthy mono (lambda f_uc (F_uc $ f_uc $ x_uc))  | 
|
188  | 
|> HOLogic.mk_Trueprop  | 
|
189  | 
|> Logic.all x_uc;  | 
|
190  | 
||
191  | 
val mono_thm = Goal.prove_internal [] (cert mono_goal)  | 
|
192  | 
(K (mono_tac lthy 1))  | 
|
193  | 
|> Thm.forall_elim (cert x_uc);  | 
|
194  | 
||
195  | 
val f_def_rhs = curry_n arity (apply_inst lthy fixp F_uc);  | 
|
196  | 
val f_def_binding = Binding.conceal (Binding.name (Thm.def_name fname));  | 
|
197  | 
val ((f, (_, f_def)), lthy') = Local_Theory.define  | 
|
198  | 
((f_binding, mixfix), ((f_def_binding, []), f_def_rhs)) lthy;  | 
|
199  | 
||
200  | 
val eqn = HOLogic.mk_eq (list_comb (f, args),  | 
|
201  | 
Term.betapplys (F, f :: args))  | 
|
202  | 
|> HOLogic.mk_Trueprop;  | 
|
203  | 
||
204  | 
val unfold =  | 
|
205  | 
(cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq  | 
|
206  | 
OF [mono_thm, f_def])  | 
|
207  | 
|> Tactic.rule_by_tactic lthy (Simplifier.simp_tac curry_uncurry_ss 1);  | 
|
208  | 
||
209  | 
val rec_rule = let open Conv in  | 
|
210  | 
Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>  | 
|
211  | 
CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1  | 
|
212  | 
        THEN rtac @{thm refl} 1) end;
 | 
|
213  | 
in  | 
|
214  | 
lthy'  | 
|
215  | 
|> Local_Theory.note (eq_abinding, [rec_rule])  | 
|
216  | 
|-> (fn (_, rec') =>  | 
|
| 40180 | 217  | 
Spec_Rules.add Spec_Rules.Equational ([f], rec')  | 
218  | 
#> Local_Theory.note ((Binding.qualify true fname (Binding.name "simps"), []), rec') #> snd)  | 
|
| 40107 | 219  | 
end;  | 
220  | 
||
221  | 
val add_partial_function = gen_add_partial_function Specification.check_spec;  | 
|
222  | 
val add_partial_function_cmd = gen_add_partial_function Specification.read_spec;  | 
|
223  | 
||
224  | 
val mode = Parse.$$$ "(" |-- Parse.xname --| Parse.$$$ ")";
 | 
|
225  | 
||
226  | 
val _ = Outer_Syntax.local_theory  | 
|
| 40186 | 227  | 
"partial_function" "define partial function" Keyword.thy_decl  | 
| 40107 | 228  | 
((mode -- (Parse.fixes -- (Parse.where_ |-- Parse_Spec.spec)))  | 
229  | 
>> (fn (mode, (fixes, spec)) => add_partial_function_cmd mode fixes spec));  | 
|
230  | 
||
231  | 
||
232  | 
val setup = Mono_Rules.setup;  | 
|
233  | 
||
234  | 
end  |