| author | wenzelm | 
| Mon, 30 May 2016 20:58:16 +0200 | |
| changeset 63183 | 4d04e14d7ab8 | 
| parent 61853 | fb7756087101 | 
| child 67149 | e61557884799 | 
| 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 | ||
| 35365 | 7 | signature SPLIT_RULE = | 
| 11037 | 8 | sig | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 9 | val split_rule_var: Proof.context -> term -> thm -> thm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 10 | val split_rule: Proof.context -> thm -> thm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 11 | val complete_split_rule: Proof.context -> thm -> thm | 
| 11037 | 12 | end; | 
| 13 | ||
| 35365 | 14 | structure Split_Rule: SPLIT_RULE = | 
| 11037 | 15 | struct | 
| 16 | ||
| 17 | (** split rules **) | |
| 18 | ||
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60642diff
changeset | 19 | fun internal_case_prod_const (Ta, Tb, Tc) = | 
| 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60642diff
changeset | 20 |   Const (@{const_name Product_Type.internal_case_prod},
 | 
| 35365 | 21 | [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc); | 
| 22 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 23 | fun eval_internal_split ctxt = | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60642diff
changeset | 24 |   hol_simplify ctxt @{thms internal_case_prod_def} o
 | 
| 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60642diff
changeset | 25 |   hol_simplify ctxt @{thms internal_case_prod_conv};
 | 
| 35365 | 26 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 27 | fun remove_internal_split ctxt = eval_internal_split ctxt o split_all ctxt; | 
| 11037 | 28 | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 29 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 30 | (*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 | 31 | 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 | 32 | of type S.*) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37167diff
changeset | 33 | fun ap_split (Type (@{type_name Product_Type.prod}, [T1, T2])) T3 u =
 | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60642diff
changeset | 34 | internal_case_prod_const (T1, T2, T3) $ | 
| 11037 | 35 |       Abs ("v", T1,
 | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 36 | ap_split T2 T3 | 
| 32287 
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
 haftmann parents: 
30722diff
changeset | 37 | ((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 | 38 | Bound 0)) | 
| 58468 | 39 | | ap_split _ T3 u = u; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 40 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 41 | (*Curries any Var of function type in the rule*) | 
| 60362 | 42 | fun split_rule_var' ctxt (t as Var (v, Type ("fun", [T1, T2]))) rl =
 | 
| 32287 
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
 haftmann parents: 
30722diff
changeset | 43 | let val T' = HOLogic.flatten_tupleT T1 ---> T2; | 
| 11037 | 44 | val newt = ap_split T1 T2 (Var (v, T')); | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60363diff
changeset | 45 | in Thm.instantiate ([], [(dest_Var t, Thm.cterm_of ctxt newt)]) rl end | 
| 60362 | 46 | | split_rule_var' _ _ rl = rl; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 47 | |
| 11037 | 48 | |
| 40388 
cb9fd7dd641c
abolished obscure goal variant of [split_format] -- unused (cf. d1c14898fd04), unrelated to '(complete)' variant, and not at all canonical
 krauss parents: 
37678diff
changeset | 49 | (* complete splitting of partially split rules *) | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 50 | |
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 51 | 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: 
32287diff
changeset | 52 | (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 | 53 | (incr_boundvars 1 u) $ Bound 0)) | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 54 | | ap_split' _ _ u = u; | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 55 | |
| 60363 | 56 | fun complete_split_rule_var ctxt (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 | 57 | let | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 58 | val (Us', U') = strip_type T; | 
| 33957 | 59 | val Us = take (length ts) Us'; | 
| 60 | val U = drop (length ts) Us' ---> U'; | |
| 32288 
168f31155c5e
cleaned up abstract tuple operations and named them consistently
 haftmann parents: 
32287diff
changeset | 61 | val T' = maps HOLogic.flatten_tupleT Us ---> U; | 
| 19735 | 62 | 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 | 63 | let | 
| 32287 
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
 haftmann parents: 
30722diff
changeset | 64 | val Ts = HOLogic.flatten_tupleT T; | 
| 20071 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 wenzelm parents: 
19735diff
changeset | 65 | val ys = Name.variant_list xs (replicate (length Ts) a); | 
| 60363 | 66 | in | 
| 67 | (xs @ ys, | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60363diff
changeset | 68 | (dest_Var v, | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60363diff
changeset | 69 | Thm.cterm_of ctxt (HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths T) T | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60363diff
changeset | 70 | (map (Var o apfst (rpair 0)) (ys ~~ Ts)))) :: insts) | 
| 11025 
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 (vs', insts) = fold mk_tuple ts (vs, []); | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 75 | in | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60363diff
changeset | 76 | (Drule.instantiate_normalize ([], ((v, T), Thm.cterm_of ctxt newt) :: insts) rl, vs') | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 77 | end | 
| 60363 | 78 | | complete_split_rule_var _ _ x = x; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 79 | |
| 19735 | 80 | fun collect_vars (Abs (_, _, t)) = collect_vars t | 
| 81 | | collect_vars t = | |
| 82 | (case strip_comb t of | |
| 83 | (v as Var _, ts) => cons (v, ts) | |
| 58468 | 84 | | (_, ts) => fold collect_vars ts); | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 85 | |
| 11037 | 86 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 87 | fun split_rule_var ctxt = | 
| 60362 | 88 | (Drule.export_without_context o remove_internal_split ctxt) oo split_rule_var' ctxt; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 89 | |
| 11037 | 90 | (*curries ALL function variables occurring in a rule's conclusion*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 91 | fun split_rule ctxt rl = | 
| 60362 | 92 | fold_rev (split_rule_var' ctxt) (Misc_Legacy.term_vars (Thm.concl_of rl)) rl | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 93 | |> 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: 
33957diff
changeset | 94 | |> Drule.export_without_context; | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: diff
changeset | 95 | |
| 19735 | 96 | (*curries ALL function variables*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 97 | fun complete_split_rule ctxt rl = | 
| 19735 | 98 | let | 
| 99 | val prop = Thm.prop_of rl; | |
| 100 | val xs = Term.fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I) prop []; | |
| 101 | val vars = collect_vars prop []; | |
| 102 | in | |
| 60363 | 103 | fst (fold_rev (complete_split_rule_var ctxt) vars (rl, xs)) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44121diff
changeset | 104 | |> 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: 
33957diff
changeset | 105 | |> Drule.export_without_context | 
| 33368 | 106 | |> Rule_Cases.save rl | 
| 19735 | 107 | end; | 
| 11037 | 108 | |
| 109 | ||
| 110 | (* attribute syntax *) | |
| 111 | ||
| 58820 | 112 | val _ = | 
| 113 | Theory.setup | |
| 114 |    (Attrib.setup @{binding split_format}
 | |
| 115 | (Scan.lift (Args.parens (Args.$$$ "complete") | |
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61424diff
changeset | 116 | >> K (Thm.rule_attribute [] (complete_split_rule o Context.proof_of)))) | 
| 58820 | 117 | "split pair-typed subterms in premises, or function arguments" #> | 
| 118 |     Attrib.setup @{binding split_rule}
 | |
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61424diff
changeset | 119 | (Scan.succeed (Thm.rule_attribute [] (split_rule o Context.proof_of))) | 
| 58820 | 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 |