src/HOL/Tools/split_rule.ML
author wenzelm
Tue, 22 Mar 2022 18:56:28 +0100
changeset 75297 fc4d07587695
parent 74282 c2ee8d993d6a
child 77879 dd222e2af01a
permissions -rw-r--r--
more robust errors -- on foreground process instead of background server;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24584
01e83ffa6c54 fixed title
haftmann
parents: 22278
diff changeset
     1
(*  Title:      HOL/Tools/split_rule.ML
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
     2
    Author:     Stefan Berghofer, David von Oheimb, and Markus Wenzel, TU Muenchen
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
     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
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
     5
*)
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
     6
35365
2fcd08c62495 modernized structure Split_Rule;
wenzelm
parents: 35021
diff changeset
     7
signature SPLIT_RULE =
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
     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
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    12
end;
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    13
35365
2fcd08c62495 modernized structure Split_Rule;
wenzelm
parents: 35021
diff changeset
    14
structure Split_Rule: SPLIT_RULE =
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    15
struct
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    16
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    17
(** split rules **)
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    18
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60642
diff changeset
    19
fun internal_case_prod_const (Ta, Tb, Tc) =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61853
diff changeset
    20
  Const (\<^const_name>\<open>Product_Type.internal_case_prod\<close>,
35365
2fcd08c62495 modernized structure Split_Rule;
wenzelm
parents: 35021
diff changeset
    21
    [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc);
2fcd08c62495 modernized structure Split_Rule;
wenzelm
parents: 35021
diff changeset
    22
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 44121
diff changeset
    23
fun eval_internal_split ctxt =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60642
diff changeset
    24
  hol_simplify ctxt @{thms internal_case_prod_def} o
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60642
diff changeset
    25
  hol_simplify ctxt @{thms internal_case_prod_conv};
35365
2fcd08c62495 modernized structure Split_Rule;
wenzelm
parents: 35021
diff changeset
    26
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 44121
diff changeset
    27
fun remove_internal_split ctxt = eval_internal_split ctxt o split_all ctxt;
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    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.*)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61853
diff changeset
    33
fun ap_split (Type (\<^type_name>\<open>Product_Type.prod\<close>, [T1, T2])) T3 u =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60642
diff changeset
    34
      internal_case_prod_const (T1, T2, T3) $
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    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: 30722
diff 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
haftmann
parents: 51717
diff changeset
    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
befdc10ebb42 clarified context;
wenzelm
parents: 59621
diff changeset
    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: 30722
diff changeset
    43
      let val T' = HOLogic.flatten_tupleT T1 ---> T2;
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    44
          val newt = ap_split T1 T2 (Var (v, T'));
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 67149
diff changeset
    45
      in Thm.instantiate (TVars.empty, Vars.make [(dest_Var t, Thm.cterm_of ctxt newt)]) rl end
60362
befdc10ebb42 clarified context;
wenzelm
parents: 59621
diff changeset
    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
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    48
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
    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: 32287
diff 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
5568b16aa477 clarified context;
wenzelm
parents: 60362
diff changeset
    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
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
    59
        val Us = take (length ts) Us';
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
    60
        val U = drop (length ts) Us' ---> U';
32288
168f31155c5e cleaned up abstract tuple operations and named them consistently
haftmann
parents: 32287
diff changeset
    61
        val T' = maps HOLogic.flatten_tupleT Us ---> U;
19735
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    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: 30722
diff changeset
    64
                val Ts = HOLogic.flatten_tupleT T;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19735
diff changeset
    65
                val ys = Name.variant_list xs (replicate (length Ts) a);
60363
5568b16aa477 clarified context;
wenzelm
parents: 60362
diff changeset
    66
              in
5568b16aa477 clarified context;
wenzelm
parents: 60362
diff changeset
    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: 60363
diff 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: 60363
diff 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: 60363
diff 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
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    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
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    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
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 67149
diff changeset
    76
        (Drule.instantiate_normalize
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 67149
diff changeset
    77
          (TVars.empty, Vars.make (((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
    78
      end
60363
5568b16aa477 clarified context;
wenzelm
parents: 60362
diff changeset
    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
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    81
fun collect_vars (Abs (_, _, t)) = collect_vars t
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    82
  | collect_vars t =
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    83
      (case strip_comb t of
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    84
        (v as Var _, ts) => cons (v, ts)
58468
haftmann
parents: 51717
diff changeset
    85
      | (_, ts) => fold collect_vars ts);
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff changeset
    86
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    87
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 44121
diff changeset
    88
fun split_rule_var ctxt =
60362
befdc10ebb42 clarified context;
wenzelm
parents: 59621
diff changeset
    89
  (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
    90
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
    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 =
60362
befdc10ebb42 clarified context;
wenzelm
parents: 59621
diff changeset
    93
  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: 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
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    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
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
    99
  let
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
   100
    val prop = Thm.prop_of rl;
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
   101
    val xs = Term.fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I) prop [];
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
   102
    val vars = collect_vars prop [];
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
   103
  in
60363
5568b16aa477 clarified context;
wenzelm
parents: 60362
diff changeset
   104
    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: 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
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 32342
diff changeset
   107
    |> Rule_Cases.save rl
19735
ff13585fbdab complete_split_rule: all Vars;
wenzelm
parents: 18728
diff changeset
   108
  end;
11037
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
   109
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
   110
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
   111
(* attribute syntax *)
37716a82a3d9 module setup;
wenzelm
parents: 11025
diff changeset
   112
58820
3ad2759acc52 modernized setup;
wenzelm
parents: 58468
diff changeset
   113
val _ =
3ad2759acc52 modernized setup;
wenzelm
parents: 58468
diff changeset
   114
  Theory.setup
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61853
diff changeset
   115
   (Attrib.setup \<^binding>\<open>split_format\<close>
58820
3ad2759acc52 modernized setup;
wenzelm
parents: 58468
diff changeset
   116
      (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: 61424
diff changeset
   117
        >> K (Thm.rule_attribute [] (complete_split_rule o Context.proof_of))))
58820
3ad2759acc52 modernized setup;
wenzelm
parents: 58468
diff changeset
   118
      "split pair-typed subterms in premises, or function arguments" #>
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61853
diff changeset
   119
    Attrib.setup \<^binding>\<open>split_rule\<close>
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: 61424
diff changeset
   120
      (Scan.succeed (Thm.rule_attribute [] (split_rule o Context.proof_of)))
58820
3ad2759acc52 modernized setup;
wenzelm
parents: 58468
diff changeset
   121
      "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
   122
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff changeset
   123
end;
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
diff changeset
   124