| author | immler@in.tum.de | 
| Sun, 28 Jun 2009 15:01:28 +0200 | |
| changeset 31834 | b7f1e86d9f04 | 
| parent 30722 | 623d4831c8cf | 
| child 32287 | 65d5c5b30747 | 
| 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: 
15574diff
changeset | 4 | Some tools for managing tupled arguments and abstractions in rules. | 
| 11037 | 5 | *) | 
| 6 | ||
| 7 | signature BASIC_SPLIT_RULE = | |
| 8 | sig | |
| 9 | val split_rule: thm -> thm | |
| 10 | val complete_split_rule: thm -> thm | |
| 11 | end; | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 12 | |
| 11037 | 13 | signature SPLIT_RULE = | 
| 14 | sig | |
| 15 | include BASIC_SPLIT_RULE | |
| 19735 | 16 | val split_rule_var: term -> thm -> thm | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
26352diff
changeset | 17 | val split_rule_goal: Proof.context -> string list list -> thm -> thm | 
| 18708 | 18 | val setup: theory -> theory | 
| 11037 | 19 | end; | 
| 20 | ||
| 21 | structure SplitRule: SPLIT_RULE = | |
| 22 | struct | |
| 23 | ||
| 24 | ||
| 19735 | 25 | |
| 11037 | 26 | (** theory context references **) | 
| 27 | ||
| 11838 | 28 | val split_conv = thm "split_conv"; | 
| 29 | val fst_conv = thm "fst_conv"; | |
| 30 | val snd_conv = thm "snd_conv"; | |
| 31 | ||
| 11037 | 32 | fun internal_split_const (Ta, Tb, Tc) = | 
| 33 |   Const ("Product_Type.internal_split", [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc);
 | |
| 34 | ||
| 35 | val internal_split_def = thm "internal_split_def"; | |
| 36 | val internal_split_conv = thm "internal_split_conv"; | |
| 37 | ||
| 38 | ||
| 39 | ||
| 40 | (** split rules **) | |
| 41 | ||
| 42 | val eval_internal_split = hol_simplify [internal_split_def] o hol_simplify [internal_split_conv]; | |
| 43 | val remove_internal_split = eval_internal_split o split_all; | |
| 44 | ||
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 45 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 46 | (*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 | 47 | 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 | 48 | of type S.*) | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 49 | fun ap_split (Type ("*", [T1, T2])) T3 u =
 | 
| 11037 | 50 | internal_split_const (T1, T2, T3) $ | 
| 51 |       Abs ("v", T1,
 | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 52 | ap_split T2 T3 | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 53 | ((ap_split T1 (HOLogic.prodT_factors T2 ---> T3) (incr_boundvars 1 u)) $ | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 54 | Bound 0)) | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 55 | | ap_split T T3 u = u; | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 56 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 57 | (*Curries any Var of function type in the rule*) | 
| 19735 | 58 | fun split_rule_var' (t as Var (v, Type ("fun", [T1, T2]))) rl =
 | 
| 11037 | 59 | let val T' = HOLogic.prodT_factors T1 ---> T2; | 
| 60 | val newt = ap_split T1 T2 (Var (v, T')); | |
| 19735 | 61 | val cterm = Thm.cterm_of (Thm.theory_of_thm rl); | 
| 11037 | 62 | in Thm.instantiate ([], [(cterm t, cterm newt)]) rl end | 
| 19735 | 63 | | split_rule_var' t rl = rl; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 64 | |
| 11037 | 65 | |
| 66 | (* complete splitting of partially splitted rules *) | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 67 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 68 | fun ap_split' (T::Ts) U u = Abs ("v", T, ap_split' Ts U
 | 
| 15570 | 69 | (ap_split T (List.concat (map HOLogic.prodT_factors Ts) ---> U) | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 70 | (incr_boundvars 1 u) $ Bound 0)) | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 71 | | ap_split' _ _ u = u; | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 72 | |
| 19735 | 73 | 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 | 74 | let | 
| 19735 | 75 | 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 | 76 | val (Us', U') = strip_type T; | 
| 15570 | 77 | val Us = Library.take (length ts, Us'); | 
| 78 | val U = Library.drop (length ts, Us') ---> U'; | |
| 79 | val T' = List.concat (map HOLogic.prodT_factors Us) ---> U; | |
| 19735 | 80 | 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 | 81 | let | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 82 | val Ts = HOLogic.prodT_factors T; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19735diff
changeset | 83 | val ys = Name.variant_list xs (replicate (length Ts) a); | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 84 | in (xs @ ys, (cterm v, cterm (HOLogic.mk_tuple T | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 85 | (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 | 86 | end | 
| 19735 | 87 | | mk_tuple _ x = x; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 88 | val newt = ap_split' Us U (Var (v, T')); | 
| 19735 | 89 | val cterm = Thm.cterm_of (Thm.theory_of_thm rl); | 
| 90 | 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 | 91 | in | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 92 | (instantiate ([], [(cterm t, cterm newt)] @ insts) rl, vs') | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 93 | end | 
| 19735 | 94 | | complete_split_rule_var _ x = x; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 95 | |
| 19735 | 96 | fun collect_vars (Abs (_, _, t)) = collect_vars t | 
| 97 | | collect_vars t = | |
| 98 | (case strip_comb t of | |
| 99 | (v as Var _, ts) => cons (v, ts) | |
| 100 | | (t, ts) => fold collect_vars ts); | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 101 | |
| 11037 | 102 | |
| 19735 | 103 | val split_rule_var = (Drule.standard o remove_internal_split) oo split_rule_var'; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 104 | |
| 11037 | 105 | (*curries ALL function variables occurring in a rule's conclusion*) | 
| 106 | fun split_rule rl = | |
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
27882diff
changeset | 107 | fold_rev split_rule_var' (OldTerm.term_vars (concl_of rl)) rl | 
| 11037 | 108 | |> remove_internal_split | 
| 109 | |> Drule.standard; | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 110 | |
| 19735 | 111 | (*curries ALL function variables*) | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 112 | fun complete_split_rule rl = | 
| 19735 | 113 | let | 
| 114 | val prop = Thm.prop_of rl; | |
| 115 | val xs = Term.fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I) prop []; | |
| 116 | val vars = collect_vars prop []; | |
| 117 | in | |
| 118 | fst (fold_rev complete_split_rule_var vars (rl, xs)) | |
| 119 | |> remove_internal_split | |
| 120 | |> Drule.standard | |
| 121 | |> RuleCases.save rl | |
| 122 | end; | |
| 11037 | 123 | |
| 124 | ||
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
26352diff
changeset | 125 | fun pair_tac ctxt s = | 
| 27239 | 126 |   res_inst_tac ctxt [(("p", 0), s)] @{thm PairE}
 | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
26352diff
changeset | 127 | THEN' hyp_subst_tac | 
| 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
26352diff
changeset | 128 | THEN' K prune_params_tac; | 
| 26352 | 129 | |
| 11037 | 130 | val split_rule_ss = HOL_basic_ss addsimps [split_conv, fst_conv, snd_conv]; | 
| 131 | ||
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
26352diff
changeset | 132 | fun split_rule_goal ctxt xss rl = | 
| 11037 | 133 | let | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
26352diff
changeset | 134 | fun one_split i s = Tactic.rule_by_tactic (pair_tac ctxt s i); | 
| 19735 | 135 | fun one_goal (i, xs) = fold (one_split (i + 1)) xs; | 
| 11037 | 136 | in | 
| 18050 | 137 | rl | 
| 19735 | 138 | |> fold_index one_goal xss | 
| 18050 | 139 | |> Simplifier.full_simplify split_rule_ss | 
| 19735 | 140 | |> Drule.standard | 
| 11037 | 141 | |> RuleCases.save rl | 
| 142 | end; | |
| 143 | ||
| 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: 
30528diff
changeset | 144 | |
| 11037 | 145 | (* attribute syntax *) | 
| 146 | ||
| 27882 
eaa9fef9f4c1
Args.name_source(_position) for proper position information;
 wenzelm parents: 
27809diff
changeset | 147 | (* FIXME dynamically scoped due to Args.name_source/pair_tac *) | 
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15574diff
changeset | 148 | |
| 11037 | 149 | 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: 
30528diff
changeset | 150 |   Attrib.setup @{binding split_format}
 | 
| 
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: 
30528diff
changeset | 151 | (Scan.lift | 
| 
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: 
30528diff
changeset | 152 | (Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute (K complete_split_rule)) || | 
| 
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: 
30528diff
changeset | 153 | OuterParse.and_list1 (Scan.repeat Args.name_source) | 
| 
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: 
30528diff
changeset | 154 | >> (fn xss => Thm.rule_attribute (fn context => | 
| 
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: 
30528diff
changeset | 155 | split_rule_goal (Context.proof_of context) xss)))) | 
| 30528 | 156 | "split pair-typed subterms in premises, or function arguments" #> | 
| 157 |   Attrib.setup @{binding split_rule} (Scan.succeed (Thm.rule_attribute (K split_rule)))
 | |
| 158 | "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 | 159 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 160 | end; | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 161 | |
| 11037 | 162 | structure BasicSplitRule: BASIC_SPLIT_RULE = SplitRule; | 
| 163 | open BasicSplitRule; |