| author | traytel | 
| Tue, 03 Mar 2015 19:08:04 +0100 | |
| changeset 59580 | cbc38731d42f | 
| parent 59498 | 50b60f501b05 | 
| child 59582 | 0fbed69ff081 | 
| 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  | 
|
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
9  | 
val init: string -> term -> term -> thm -> thm -> thm option -> declaration  | 
| 40107 | 10  | 
|
| 52727 | 11  | 
val mono_tac: Proof.context -> int -> tactic  | 
12  | 
||
| 40107 | 13  | 
val add_partial_function: string -> (binding * typ option * mixfix) list ->  | 
14  | 
Attrib.binding * term -> local_theory -> local_theory  | 
|
15  | 
||
16  | 
val add_partial_function_cmd: string -> (binding * string option * mixfix) list ->  | 
|
17  | 
Attrib.binding * string -> local_theory -> local_theory  | 
|
18  | 
end;  | 
|
19  | 
||
20  | 
||
21  | 
structure Partial_Function: PARTIAL_FUNCTION =  | 
|
22  | 
struct  | 
|
23  | 
||
24  | 
(*** Context Data ***)  | 
|
25  | 
||
| 43080 | 26  | 
datatype setup_data = Setup_Data of  | 
27  | 
 {fixp: term,
 | 
|
28  | 
mono: term,  | 
|
29  | 
fixp_eq: thm,  | 
|
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
30  | 
fixp_induct: thm,  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
31  | 
fixp_induct_user: thm option};  | 
| 43080 | 32  | 
|
| 40107 | 33  | 
structure Modes = Generic_Data  | 
34  | 
(  | 
|
| 43080 | 35  | 
type T = setup_data Symtab.table;  | 
| 40107 | 36  | 
val empty = Symtab.empty;  | 
37  | 
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
 | 
38  | 
fun merge data = Symtab.merge (K true) data;  | 
| 40107 | 39  | 
)  | 
40  | 
||
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
41  | 
fun init mode fixp mono fixp_eq fixp_induct fixp_induct_user phi =  | 
| 40107 | 42  | 
let  | 
43  | 
val term = Morphism.term phi;  | 
|
| 43080 | 44  | 
val thm = Morphism.thm phi;  | 
45  | 
val data' = Setup_Data  | 
|
46  | 
      {fixp=term fixp, mono=term mono, fixp_eq=thm fixp_eq,
 | 
|
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
47  | 
fixp_induct=thm fixp_induct, fixp_induct_user=Option.map thm fixp_induct_user};  | 
| 40107 | 48  | 
in  | 
| 
42949
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
42495 
diff
changeset
 | 
49  | 
Modes.map (Symtab.update (mode, data'))  | 
| 40107 | 50  | 
end  | 
51  | 
||
52  | 
val known_modes = Symtab.keys o Modes.get o Context.Proof;  | 
|
53  | 
val lookup_mode = Symtab.lookup o Modes.get o Context.Proof;  | 
|
54  | 
||
55  | 
||
56  | 
(*** Automated monotonicity proofs ***)  | 
|
57  | 
||
58  | 
fun strip_cases ctac = ctac #> Seq.map snd;  | 
|
59  | 
||
60  | 
(*rewrite conclusion with k-th assumtion*)  | 
|
61  | 
fun rewrite_with_asm_tac ctxt k =  | 
|
| 45403 | 62  | 
  Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
 | 
| 40107 | 63  | 
Local_Defs.unfold_tac ctxt' [nth prems k]) ctxt;  | 
64  | 
||
| 
54405
 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
52728 
diff
changeset
 | 
65  | 
fun dest_case ctxt t =  | 
| 40107 | 66  | 
case strip_comb t of  | 
67  | 
(Const (case_comb, _), args) =>  | 
|
| 
54405
 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
52728 
diff
changeset
 | 
68  | 
(case Ctr_Sugar.ctr_sugar_of_case ctxt case_comb of  | 
| 40107 | 69  | 
NONE => NONE  | 
| 
54405
 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
52728 
diff
changeset
 | 
70  | 
       | SOME {case_thms, ...} =>
 | 
| 40107 | 71  | 
let  | 
| 
54405
 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
52728 
diff
changeset
 | 
72  | 
val lhs = prop_of (hd case_thms)  | 
| 40107 | 73  | 
|> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst;  | 
74  | 
val arity = length (snd (strip_comb lhs));  | 
|
75  | 
val conv = funpow (length args - arity) Conv.fun_conv  | 
|
| 
54405
 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
52728 
diff
changeset
 | 
76  | 
(Conv.rewrs_conv (map mk_meta_eq case_thms));  | 
| 40107 | 77  | 
in  | 
78  | 
SOME (nth args (arity - 1), conv)  | 
|
79  | 
end)  | 
|
80  | 
| _ => NONE;  | 
|
81  | 
||
82  | 
(*split on case expressions*)  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
83  | 
val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} =>
 | 
| 40107 | 84  | 
SUBGOAL (fn (t, i) => case t of  | 
85  | 
_ $ (_ $ Abs (_, _, body)) =>  | 
|
| 
54405
 
88f6d5b1422f
ported 'partial_function' to 'Ctr_Sugar' abstraction
 
blanchet 
parents: 
52728 
diff
changeset
 | 
86  | 
(case dest_case ctxt body of  | 
| 40107 | 87  | 
NONE => no_tac  | 
88  | 
| SOME (arg, conv) =>  | 
|
89  | 
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
 | 
90  | 
if Term.is_open arg then no_tac  | 
| 40107 | 91  | 
else ((DETERM o strip_cases o Induct.cases_tac ctxt false [[SOME arg]] NONE [])  | 
92  | 
THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0)  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
93  | 
                THEN_ALL_NEW eresolve_tac ctxt @{thms thin_rl}
 | 
| 40107 | 94  | 
THEN_ALL_NEW (CONVERSION  | 
95  | 
(params_conv ~1 (fn ctxt' =>  | 
|
96  | 
arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i  | 
|
97  | 
end)  | 
|
98  | 
| _ => no_tac) 1);  | 
|
99  | 
||
100  | 
(*monotonicity proof: apply rules + split case expressions*)  | 
|
101  | 
fun mono_tac ctxt =  | 
|
102  | 
  K (Local_Defs.unfold_tac ctxt [@{thm curry_def}])
 | 
|
103  | 
THEN' (TRY o REPEAT_ALL_NEW  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
104  | 
   (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems partial_function_mono}))
 | 
| 40107 | 105  | 
ORELSE' split_cases_tac ctxt));  | 
106  | 
||
107  | 
||
108  | 
(*** Auxiliary functions ***)  | 
|
109  | 
||
110  | 
(*positional instantiation with computed type substitution.  | 
|
111  | 
internal version of attribute "[of s t u]".*)  | 
|
112  | 
fun cterm_instantiate' cts thm =  | 
|
113  | 
let  | 
|
114  | 
val thy = Thm.theory_of_thm thm;  | 
|
115  | 
val vs = rev (Term.add_vars (prop_of thm) [])  | 
|
116  | 
|> map (Thm.cterm_of thy o Var);  | 
|
117  | 
in  | 
|
118  | 
cterm_instantiate (zip_options vs cts) thm  | 
|
119  | 
end;  | 
|
120  | 
||
121  | 
(*Returns t $ u, but instantiates the type of t to make the  | 
|
122  | 
application type correct*)  | 
|
123  | 
fun apply_inst ctxt t u =  | 
|
124  | 
let  | 
|
| 42361 | 125  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 40107 | 126  | 
val T = domain_type (fastype_of t);  | 
127  | 
val T' = fastype_of u;  | 
|
| 42388 | 128  | 
val subst = Sign.typ_match thy (T, T') Vartab.empty  | 
| 40107 | 129  | 
      handle Type.TYPE_MATCH => raise TYPE ("apply_inst", [T, T'], [t, u])
 | 
130  | 
in  | 
|
131  | 
map_types (Envir.norm_type subst) t $ u  | 
|
132  | 
end;  | 
|
133  | 
||
134  | 
fun head_conv cv ct =  | 
|
135  | 
if can Thm.dest_comb ct then Conv.fun_conv (head_conv cv) ct else cv ct;  | 
|
136  | 
||
137  | 
||
138  | 
(*** currying transformation ***)  | 
|
139  | 
||
140  | 
fun curry_const (A, B, C) =  | 
|
141  | 
  Const (@{const_name Product_Type.curry},
 | 
|
142  | 
[HOLogic.mk_prodT (A, B) --> C, A, B] ---> C);  | 
|
143  | 
||
144  | 
fun mk_curry f =  | 
|
145  | 
case fastype_of f of  | 
|
146  | 
    Type ("fun", [Type (_, [S, T]), U]) =>
 | 
|
147  | 
curry_const (S, T, U) $ f  | 
|
148  | 
  | T => raise TYPE ("mk_curry", [T], [f]);
 | 
|
149  | 
||
150  | 
(* iterated versions. Nonstandard left-nested tuples arise naturally  | 
|
151  | 
from "split o split o split"*)  | 
|
152  | 
fun curry_n arity = funpow (arity - 1) mk_curry;  | 
|
153  | 
fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split;  | 
|
154  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
155  | 
val curry_uncurry_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
156  | 
  simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
157  | 
    addsimps [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}])
 | 
| 40107 | 158  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
159  | 
val split_conv_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
160  | 
  simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
161  | 
    addsimps [@{thm Product_Type.split_conv}]);
 | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
162  | 
|
| 
54630
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
54405 
diff
changeset
 | 
163  | 
val curry_K_ss =  | 
| 
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
54405 
diff
changeset
 | 
164  | 
  simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
54405 
diff
changeset
 | 
165  | 
    addsimps [@{thm Product_Type.curry_K}]);
 | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
166  | 
|
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
167  | 
(* instantiate generic fixpoint induction and eliminate the canonical assumptions;  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
168  | 
curry induction predicate *)  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
169  | 
fun specialize_fixp_induct ctxt args fT fT_uc F curry uncurry mono_thm f_def rule =  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
170  | 
let  | 
| 59580 | 171  | 
val cert = Proof_Context.cterm_of ctxt  | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
172  | 
val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
173  | 
    val P_inst = Abs ("f", fT_uc, Free (P, fT --> HOLogic.boolT) $ (curry $ Bound 0))
 | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
174  | 
in  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54630 
diff
changeset
 | 
175  | 
(* FIXME ctxt vs. ctxt' (!?) *)  | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
176  | 
rule  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
177  | 
|> cterm_instantiate' [SOME (cert uncurry), NONE, SOME (cert curry), NONE, SOME (cert P_inst)]  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
178  | 
|> Tactic.rule_by_tactic ctxt  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
179  | 
(Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3 (* discharge U (C f) = f *)  | 
| 
54630
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
54405 
diff
changeset
 | 
180  | 
THEN Simplifier.simp_tac (put_simpset curry_K_ss ctxt) 4 (* simplify bot case *)  | 
| 
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
54405 
diff
changeset
 | 
181  | 
THEN Simplifier.full_simp_tac (put_simpset curry_uncurry_ss ctxt) 5) (* simplify induction step *)  | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
182  | 
|> (fn thm => thm OF [mono_thm, f_def])  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
183  | 
|> Conv.fconv_rule (Conv.concl_conv ~1 (* simplify conclusion *)  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54630 
diff
changeset
 | 
184  | 
         (Raw_Simplifier.rewrite ctxt false [mk_meta_eq @{thm Product_Type.curry_split}]))
 | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
185  | 
|> singleton (Variable.export ctxt' ctxt)  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
186  | 
end  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
187  | 
|
| 52727 | 188  | 
fun mk_curried_induct args ctxt inst_rule =  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
189  | 
let  | 
| 59580 | 190  | 
val cert = Proof_Context.cterm_of ctxt  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
191  | 
val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
192  | 
|
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
193  | 
val split_paired_all_conv =  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
194  | 
      Conv.every_conv (replicate (length args - 1) (Conv.rewr_conv @{thm split_paired_all}))
 | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
195  | 
|
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
196  | 
val split_params_conv =  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
197  | 
Conv.params_conv ~1 (fn ctxt' =>  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
198  | 
Conv.implies_conv split_paired_all_conv Conv.all_conv)  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
199  | 
|
| 
52521
 
a1c4f586e372
more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
 
krauss 
parents: 
52087 
diff
changeset
 | 
200  | 
val (P_var, x_var) =  | 
| 
51484
 
49eb8d73ae10
allow induction predicates with arbitrary arity (not just binary)
 
krauss 
parents: 
46961 
diff
changeset
 | 
201  | 
Thm.prop_of inst_rule |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop  | 
| 
52521
 
a1c4f586e372
more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
 
krauss 
parents: 
52087 
diff
changeset
 | 
202  | 
|> strip_comb |> apsnd hd  | 
| 
 
a1c4f586e372
more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
 
krauss 
parents: 
52087 
diff
changeset
 | 
203  | 
val P_rangeT = range_type (snd (Term.dest_Var P_var))  | 
| 
51484
 
49eb8d73ae10
allow induction predicates with arbitrary arity (not just binary)
 
krauss 
parents: 
46961 
diff
changeset
 | 
204  | 
val PT = map (snd o dest_Free) args ---> P_rangeT  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
205  | 
val x_inst = cert (foldl1 HOLogic.mk_prod args)  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
206  | 
val P_inst = cert (uncurry_n (length args) (Free (P, PT)))  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
207  | 
|
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
208  | 
val inst_rule' = inst_rule  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
209  | 
|> Tactic.rule_by_tactic ctxt  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
210  | 
(Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 4  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
211  | 
THEN Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
212  | 
THEN CONVERSION (split_params_conv ctxt  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
213  | 
then_conv (Conv.forall_conv (K split_paired_all_conv) ctxt)) 3)  | 
| 
52521
 
a1c4f586e372
more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
 
krauss 
parents: 
52087 
diff
changeset
 | 
214  | 
|> Thm.instantiate ([], [(cert P_var, P_inst), (cert x_var, x_inst)])  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51484 
diff
changeset
 | 
215  | 
|> Simplifier.full_simplify (put_simpset split_conv_ss ctxt)  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
216  | 
|> singleton (Variable.export ctxt' ctxt)  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
217  | 
in  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
218  | 
inst_rule'  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
219  | 
end;  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
220  | 
|
| 40107 | 221  | 
|
222  | 
(*** partial_function definition ***)  | 
|
223  | 
||
224  | 
fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy =  | 
|
225  | 
let  | 
|
| 43080 | 226  | 
val setup_data = the (lookup_mode lthy mode)  | 
| 40107 | 227  | 
handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".",  | 
228  | 
"Known modes are " ^ commas_quote (known_modes lthy) ^ "."]);  | 
|
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
229  | 
    val Setup_Data {fixp, mono, fixp_eq, fixp_induct, fixp_induct_user} = setup_data;
 | 
| 40107 | 230  | 
|
231  | 
val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;  | 
|
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
232  | 
val ((_, plain_eqn), args_ctxt) = Variable.focus eqn lthy;  | 
| 40107 | 233  | 
|
234  | 
val ((f_binding, fT), mixfix) = the_single fixes;  | 
|
235  | 
val fname = Binding.name_of f_binding;  | 
|
236  | 
||
| 59580 | 237  | 
val cert = Proof_Context.cterm_of lthy;  | 
| 40107 | 238  | 
val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn);  | 
239  | 
val (head, args) = strip_comb lhs;  | 
|
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
240  | 
val argnames = map (fst o dest_Free) args;  | 
| 40107 | 241  | 
val F = fold_rev lambda (head :: args) rhs;  | 
242  | 
||
243  | 
val arity = length args;  | 
|
244  | 
val (aTs, bTs) = chop arity (binder_types fT);  | 
|
245  | 
||
246  | 
val tupleT = foldl1 HOLogic.mk_prodT aTs;  | 
|
247  | 
val fT_uc = tupleT :: bTs ---> body_type fT;  | 
|
248  | 
val f_uc = Var ((fname, 0), fT_uc);  | 
|
249  | 
    val x_uc = Var (("x", 0), tupleT);
 | 
|
250  | 
val uncurry = lambda head (uncurry_n arity head);  | 
|
251  | 
val curry = lambda f_uc (curry_n arity f_uc);  | 
|
252  | 
||
253  | 
val F_uc =  | 
|
254  | 
lambda f_uc (uncurry_n arity (F $ curry_n arity f_uc));  | 
|
255  | 
||
256  | 
val mono_goal = apply_inst lthy mono (lambda f_uc (F_uc $ f_uc $ x_uc))  | 
|
257  | 
|> HOLogic.mk_Trueprop  | 
|
258  | 
|> Logic.all x_uc;  | 
|
259  | 
||
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
260  | 
val mono_thm = Goal.prove_internal lthy [] (cert mono_goal)  | 
| 40107 | 261  | 
(K (mono_tac lthy 1))  | 
| 52727 | 262  | 
val inst_mono_thm = Thm.forall_elim (cert x_uc) mono_thm  | 
| 40107 | 263  | 
|
264  | 
val f_def_rhs = curry_n arity (apply_inst lthy fixp F_uc);  | 
|
265  | 
val f_def_binding = Binding.conceal (Binding.name (Thm.def_name fname));  | 
|
266  | 
val ((f, (_, f_def)), lthy') = Local_Theory.define  | 
|
267  | 
((f_binding, mixfix), ((f_def_binding, []), f_def_rhs)) lthy;  | 
|
268  | 
||
269  | 
val eqn = HOLogic.mk_eq (list_comb (f, args),  | 
|
270  | 
Term.betapplys (F, f :: args))  | 
|
271  | 
|> HOLogic.mk_Trueprop;  | 
|
272  | 
||
273  | 
val unfold =  | 
|
274  | 
(cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq  | 
|
| 52727 | 275  | 
OF [inst_mono_thm, f_def])  | 
| 
52087
 
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
276  | 
|> Tactic.rule_by_tactic lthy' (Simplifier.simp_tac (put_simpset curry_uncurry_ss lthy') 1);  | 
| 40107 | 277  | 
|
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
278  | 
val specialized_fixp_induct =  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
279  | 
specialize_fixp_induct lthy' args fT fT_uc F curry uncurry inst_mono_thm f_def fixp_induct  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
280  | 
|> Drule.rename_bvars' (map SOME (fname :: fname :: argnames));  | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
281  | 
|
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
282  | 
val mk_raw_induct =  | 
| 52727 | 283  | 
cterm_instantiate' [SOME (cert uncurry), NONE, SOME (cert curry)]  | 
284  | 
#> mk_curried_induct args args_ctxt  | 
|
| 
52087
 
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
285  | 
#> singleton (Variable.export args_ctxt lthy')  | 
| 52727 | 286  | 
#> (fn thm => cterm_instantiate' [SOME (cert F)] thm OF [inst_mono_thm, f_def])  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
287  | 
#> Drule.rename_bvars' (map SOME (fname :: argnames @ argnames))  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
288  | 
|
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
289  | 
val raw_induct = Option.map mk_raw_induct fixp_induct_user  | 
| 40107 | 290  | 
val rec_rule = let open Conv in  | 
291  | 
Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>  | 
|
292  | 
CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
293  | 
        THEN resolve_tac lthy' @{thms refl} 1) end;
 | 
| 40107 | 294  | 
in  | 
295  | 
lthy'  | 
|
296  | 
|> Local_Theory.note (eq_abinding, [rec_rule])  | 
|
297  | 
|-> (fn (_, rec') =>  | 
|
| 40180 | 298  | 
Spec_Rules.add Spec_Rules.Equational ([f], rec')  | 
299  | 
#> Local_Theory.note ((Binding.qualify true fname (Binding.name "simps"), []), rec') #> snd)  | 
|
| 52727 | 300  | 
|> (Local_Theory.note ((Binding.qualify true fname (Binding.name "mono"), []), [mono_thm]) #> snd)  | 
| 
43083
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
301  | 
|> (case raw_induct of NONE => I | SOME thm =>  | 
| 
 
df41a5762c3d
generate raw induction rule as instance of generic rule with careful treatment of currying
 
krauss 
parents: 
43080 
diff
changeset
 | 
302  | 
Local_Theory.note ((Binding.qualify true fname (Binding.name "raw_induct"), []), [thm]) #> snd)  | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52727 
diff
changeset
 | 
303  | 
|> (Local_Theory.note ((Binding.qualify true fname (Binding.name "fixp_induct"), []), [specialized_fixp_induct]) #> snd)  | 
| 40107 | 304  | 
end;  | 
305  | 
||
306  | 
val add_partial_function = gen_add_partial_function Specification.check_spec;  | 
|
307  | 
val add_partial_function_cmd = gen_add_partial_function Specification.read_spec;  | 
|
308  | 
||
| 46949 | 309  | 
val mode = @{keyword "("} |-- Parse.xname --| @{keyword ")"};
 | 
| 40107 | 310  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
311  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
312  | 
  Outer_Syntax.local_theory @{command_spec "partial_function"} "define partial function"
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
313  | 
((mode -- (Parse.fixes -- (Parse.where_ |-- Parse_Spec.spec)))  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
314  | 
>> (fn (mode, (fixes, spec)) => add_partial_function_cmd mode fixes spec));  | 
| 40107 | 315  | 
|
316  | 
end  |