author | wenzelm |
Thu, 18 Apr 2013 17:07:01 +0200 | |
changeset 51717 | 9e7d1c139569 |
parent 44121 | 44adaa6db327 |
child 58468 | d1f6a38f9415 |
permissions | -rw-r--r-- |
24584 | 1 |
(* Title: HOL/Tools/split_rule.ML |
11037 | 2 |
Author: Stefan Berghofer, David von Oheimb, and Markus Wenzel, TU Muenchen |
3 |
||
15661
9ef583b08647
reverted renaming of Some/None in comments and strings;
wenzelm
parents:
15574
diff
changeset
|
4 |
Some tools for managing tupled arguments and abstractions in rules. |
11037 | 5 |
*) |
6 |
||
35365 | 7 |
signature SPLIT_RULE = |
11037 | 8 |
sig |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
9 |
val split_rule_var: Proof.context -> term -> thm -> thm |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
10 |
val split_rule: Proof.context -> thm -> thm |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
11 |
val complete_split_rule: Proof.context -> thm -> thm |
18708 | 12 |
val setup: theory -> theory |
11037 | 13 |
end; |
14 |
||
35365 | 15 |
structure Split_Rule: SPLIT_RULE = |
11037 | 16 |
struct |
17 |
||
18 |
(** split rules **) |
|
19 |
||
35365 | 20 |
fun internal_split_const (Ta, Tb, Tc) = |
21 |
Const (@{const_name Product_Type.internal_split}, |
|
22 |
[[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc); |
|
23 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
24 |
fun eval_internal_split ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
25 |
hol_simplify ctxt @{thms internal_split_def} o |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
26 |
hol_simplify ctxt @{thms internal_split_conv}; |
35365 | 27 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
28 |
fun remove_internal_split ctxt = eval_internal_split ctxt o split_all ctxt; |
11037 | 29 |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
30 |
|
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
31 |
(*In ap_split S T u, term u expects separate arguments for the factors of S, |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
32 |
with result type T. The call creates a new term expecting one argument |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
33 |
of type S.*) |
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37167
diff
changeset
|
34 |
fun ap_split (Type (@{type_name Product_Type.prod}, [T1, T2])) T3 u = |
11037 | 35 |
internal_split_const (T1, T2, T3) $ |
36 |
Abs ("v", T1, |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
37 |
ap_split T2 T3 |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
30722
diff
changeset
|
38 |
((ap_split T1 (HOLogic.flatten_tupleT T2 ---> T3) (incr_boundvars 1 u)) $ |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
39 |
Bound 0)) |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
40 |
| ap_split T T3 u = u; |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
41 |
|
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
42 |
(*Curries any Var of function type in the rule*) |
19735 | 43 |
fun split_rule_var' (t as Var (v, Type ("fun", [T1, T2]))) rl = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
30722
diff
changeset
|
44 |
let val T' = HOLogic.flatten_tupleT T1 ---> T2; |
11037 | 45 |
val newt = ap_split T1 T2 (Var (v, T')); |
19735 | 46 |
val cterm = Thm.cterm_of (Thm.theory_of_thm rl); |
11037 | 47 |
in Thm.instantiate ([], [(cterm t, cterm newt)]) rl end |
19735 | 48 |
| split_rule_var' t rl = rl; |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
49 |
|
11037 | 50 |
|
40388
cb9fd7dd641c
abolished obscure goal variant of [split_format] -- unused (cf. d1c14898fd04), unrelated to '(complete)' variant, and not at all canonical
krauss
parents:
37678
diff
changeset
|
51 |
(* complete splitting of partially split rules *) |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
52 |
|
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
53 |
fun ap_split' (T::Ts) U u = Abs ("v", T, ap_split' Ts U |
32288
168f31155c5e
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32287
diff
changeset
|
54 |
(ap_split T (maps HOLogic.flatten_tupleT Ts ---> U) |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
55 |
(incr_boundvars 1 u) $ Bound 0)) |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
56 |
| ap_split' _ _ u = u; |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
57 |
|
19735 | 58 |
fun complete_split_rule_var (t as Var (v, T), ts) (rl, vs) = |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
59 |
let |
19735 | 60 |
val cterm = Thm.cterm_of (Thm.theory_of_thm rl) |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
61 |
val (Us', U') = strip_type T; |
33957 | 62 |
val Us = take (length ts) Us'; |
63 |
val U = drop (length ts) Us' ---> U'; |
|
32288
168f31155c5e
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32287
diff
changeset
|
64 |
val T' = maps HOLogic.flatten_tupleT Us ---> U; |
19735 | 65 |
fun mk_tuple (v as Var ((a, _), T)) (xs, insts) = |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
66 |
let |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
30722
diff
changeset
|
67 |
val Ts = HOLogic.flatten_tupleT T; |
20071
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
19735
diff
changeset
|
68 |
val ys = Name.variant_list xs (replicate (length Ts) a); |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32288
diff
changeset
|
69 |
in (xs @ ys, (cterm v, cterm (HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths T) T |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
70 |
(map (Var o apfst (rpair 0)) (ys ~~ Ts))))::insts) |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
71 |
end |
19735 | 72 |
| mk_tuple _ x = x; |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
73 |
val newt = ap_split' Us U (Var (v, T')); |
19735 | 74 |
val cterm = Thm.cterm_of (Thm.theory_of_thm rl); |
75 |
val (vs', insts) = fold mk_tuple ts (vs, []); |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
76 |
in |
43333
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents:
40388
diff
changeset
|
77 |
(Drule.instantiate_normalize ([], [(cterm t, cterm newt)] @ insts) rl, vs') |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
78 |
end |
19735 | 79 |
| complete_split_rule_var _ x = x; |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
80 |
|
19735 | 81 |
fun collect_vars (Abs (_, _, t)) = collect_vars t |
82 |
| collect_vars t = |
|
83 |
(case strip_comb t of |
|
84 |
(v as Var _, ts) => cons (v, ts) |
|
85 |
| (t, ts) => fold collect_vars ts); |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
86 |
|
11037 | 87 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
88 |
fun split_rule_var ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
89 |
(Drule.export_without_context o remove_internal_split ctxt) oo split_rule_var'; |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
90 |
|
11037 | 91 |
(*curries ALL function variables occurring in a rule's conclusion*) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
92 |
fun split_rule ctxt rl = |
44121 | 93 |
fold_rev split_rule_var' (Misc_Legacy.term_vars (concl_of rl)) rl |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
94 |
|> remove_internal_split ctxt |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33957
diff
changeset
|
95 |
|> Drule.export_without_context; |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
96 |
|
19735 | 97 |
(*curries ALL function variables*) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
98 |
fun complete_split_rule ctxt rl = |
19735 | 99 |
let |
100 |
val prop = Thm.prop_of rl; |
|
101 |
val xs = Term.fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I) prop []; |
|
102 |
val vars = collect_vars prop []; |
|
103 |
in |
|
104 |
fst (fold_rev complete_split_rule_var vars (rl, xs)) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
105 |
|> remove_internal_split ctxt |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33957
diff
changeset
|
106 |
|> Drule.export_without_context |
33368 | 107 |
|> Rule_Cases.save rl |
19735 | 108 |
end; |
11037 | 109 |
|
110 |
||
111 |
(* attribute syntax *) |
|
112 |
||
113 |
val setup = |
|
30722
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30528
diff
changeset
|
114 |
Attrib.setup @{binding split_format} |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
115 |
(Scan.lift (Args.parens (Args.$$$ "complete") |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
116 |
>> K (Thm.rule_attribute (complete_split_rule o Context.proof_of)))) |
30528 | 117 |
"split pair-typed subterms in premises, or function arguments" #> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
118 |
Attrib.setup @{binding split_rule} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44121
diff
changeset
|
119 |
(Scan.succeed (Thm.rule_attribute (split_rule o Context.proof_of))) |
30528 | 120 |
"curries ALL function variables occurring in a rule's conclusion"; |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
121 |
|
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
122 |
end; |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff
changeset
|
123 |