src/HOL/BNF/Tools/bnf_ctr_sugar_tactics.ML
author blanchet
Thu, 06 Jun 2013 21:22:04 +0200
changeset 52325 a74e0a4741df
parent 51937 db22d73e6c3e
child 52966 3777ba39c660
permissions -rw-r--r--
tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51742
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_ctr_sugar_tactics.ML
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
     3
    Copyright   2012
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
     4
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51742
diff changeset
     5
Tactics for wrapping existing freely generated type's constructors.
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
     6
*)
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
     7
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51742
diff changeset
     8
signature BNF_CTR_SUGAR_TACTICS =
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
     9
sig
49148
93f281430e77 fixed "mk_alternate_disc_def_tac" in the case where the constructors are swapped compared with the common Nil/Cons case
blanchet
parents: 49137
diff changeset
    10
  val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
    11
  val mk_case_cong_tac: Proof.context -> thm -> thm list -> tactic
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49586
diff changeset
    12
  val mk_case_conv_tac: Proof.context -> int -> thm -> thm list -> thm list list -> thm list list ->
49153
c15a7123605c made "mk_case_eq_tac" work in the case where the first constructor is an "alternate" constructor (in which case its discriminator paradoxically starts with a negation)
blanchet
parents: 49151
diff changeset
    13
    tactic
49118
b815fa776b91 renamed theorem
blanchet
parents: 49116
diff changeset
    14
  val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
    15
  val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
    16
  val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
    17
    thm list list list -> thm list list list -> tactic
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51797
diff changeset
    18
  val mk_half_disc_exclude_tac: Proof.context -> int -> thm -> thm -> tactic
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    19
  val mk_nchotomy_tac: int -> thm -> tactic
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49118
diff changeset
    20
  val mk_other_half_disc_exclude_tac: thm -> tactic
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
    21
  val mk_split_tac: Proof.context ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
    22
    thm -> thm list -> thm list list -> thm list list list -> tactic
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
    23
  val mk_split_asm_tac: Proof.context -> thm -> tactic
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49122
diff changeset
    24
  val mk_unique_disc_def_tac: int -> thm -> tactic
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    25
end;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    26
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51742
diff changeset
    27
structure BNF_Ctr_Sugar_Tactics : BNF_CTR_SUGAR_TACTICS =
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    28
struct
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    29
49050
9f4a7ed344b4 optimized "mk_case_disc_tac"
blanchet
parents: 49049
diff changeset
    30
open BNF_Util
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
    31
open BNF_Tactics
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    32
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    33
val meta_mp = @{thm meta_mp};
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    34
49212
ca59649170b0 more sugar on codatatypes
blanchet
parents: 49174
diff changeset
    35
fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P});
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
    36
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    37
fun mk_nchotomy_tac n exhaust =
52325
blanchet
parents: 51937
diff changeset
    38
  HEADGOAL (rtac allI THEN' rtac exhaust THEN'
blanchet
parents: 51937
diff changeset
    39
   EVERY'  (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n)));
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
    40
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    41
fun mk_unique_disc_def_tac m uexhaust =
52325
blanchet
parents: 51937
diff changeset
    42
  HEADGOAL (EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl]);
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49122
diff changeset
    43
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    44
fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust =
52325
blanchet
parents: 51937
diff changeset
    45
  HEADGOAL (EVERY' ([rtac (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans),
49630
9f6ca87ab405 tuned tactics
traytel
parents: 49594
diff changeset
    46
    rtac @{thm iffI_np}, REPEAT_DETERM o etac exE,
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51797
diff changeset
    47
    hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI,
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    48
    rtac distinct, rtac uexhaust] @
49174
41790d616f63 by default, only generate one discriminator for a two-value datatype
blanchet
parents: 49168
diff changeset
    49
    (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac])
52325
blanchet
parents: 51937
diff changeset
    50
     |> k = 1 ? swap |> op @)));
49116
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49075
diff changeset
    51
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51797
diff changeset
    52
fun mk_half_disc_exclude_tac ctxt m discD disc' =
52325
blanchet
parents: 51937
diff changeset
    53
  HEADGOAL (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
blanchet
parents: 51937
diff changeset
    54
    rtac disc');
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49022
diff changeset
    55
52325
blanchet
parents: 51937
diff changeset
    56
fun mk_other_half_disc_exclude_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half);
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49022
diff changeset
    57
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
    58
fun mk_disc_exhaust_tac n exhaust discIs =
52325
blanchet
parents: 51937
diff changeset
    59
  HEADGOAL (rtac exhaust THEN'
blanchet
parents: 51937
diff changeset
    60
    EVERY' (map2 (fn k => fn discI =>
blanchet
parents: 51937
diff changeset
    61
      dtac discI THEN' select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) discIs));
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
    62
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49463
diff changeset
    63
fun mk_collapse_tac ctxt m discD sels =
52325
blanchet
parents: 51937
diff changeset
    64
  HEADGOAL (dtac discD THEN'
blanchet
parents: 51937
diff changeset
    65
    (if m = 0 then
blanchet
parents: 51937
diff changeset
    66
       atac
blanchet
parents: 51937
diff changeset
    67
     else
blanchet
parents: 51937
diff changeset
    68
       REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
blanchet
parents: 51937
diff changeset
    69
       SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl));
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
    70
51742
b5ff7393642d fix bugs in expand tactic w.r.t. datatypes with "needless" discriminators (e.g. lists with is_Nil instead of ~= Nil)
blanchet
parents: 51717
diff changeset
    71
fun mk_expand_tac ctxt n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss
b5ff7393642d fix bugs in expand tactic w.r.t. datatypes with "needless" discriminators (e.g. lists with is_Nil instead of ~= Nil)
blanchet
parents: 51717
diff changeset
    72
    disc_excludesss' =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    73
  if ms = [0] then
52325
blanchet
parents: 51937
diff changeset
    74
    HEADGOAL (rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN'
blanchet
parents: 51937
diff changeset
    75
      TRY o EVERY' [rtac udisc_exhaust, atac, rtac vdisc_exhaust, atac])
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    76
  else
51742
b5ff7393642d fix bugs in expand tactic w.r.t. datatypes with "needless" discriminators (e.g. lists with is_Nil instead of ~= Nil)
blanchet
parents: 51717
diff changeset
    77
    let val ks = 1 upto n in
52325
blanchet
parents: 51937
diff changeset
    78
      HEADGOAL (rtac udisc_exhaust THEN'
blanchet
parents: 51937
diff changeset
    79
        EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' =>
blanchet
parents: 51937
diff changeset
    80
            fn uuncollapse =>
blanchet
parents: 51937
diff changeset
    81
          EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac,
blanchet
parents: 51937
diff changeset
    82
            rtac sym, rtac vdisc_exhaust,
blanchet
parents: 51937
diff changeset
    83
            EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse =>
blanchet
parents: 51937
diff changeset
    84
              EVERY'
blanchet
parents: 51937
diff changeset
    85
                (if k' = k then
blanchet
parents: 51937
diff changeset
    86
                   [rtac (vuncollapse RS trans), TRY o atac] @
blanchet
parents: 51937
diff changeset
    87
                   (if m = 0 then
blanchet
parents: 51937
diff changeset
    88
                      [rtac refl]
blanchet
parents: 51937
diff changeset
    89
                    else
blanchet
parents: 51937
diff changeset
    90
                      [if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac],
blanchet
parents: 51937
diff changeset
    91
                       REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE,
blanchet
parents: 51937
diff changeset
    92
                       asm_simp_tac (ss_only [] ctxt)])
blanchet
parents: 51937
diff changeset
    93
                 else
blanchet
parents: 51937
diff changeset
    94
                   [dtac (the_single (if k = n then disc_excludes else disc_excludes')),
blanchet
parents: 51937
diff changeset
    95
                    etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
blanchet
parents: 51937
diff changeset
    96
                    atac, atac]))
blanchet
parents: 51937
diff changeset
    97
              ks disc_excludess disc_excludess' uncollapses)])
blanchet
parents: 51937
diff changeset
    98
          ks ms disc_excludesss disc_excludesss' uncollapses))
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
    99
    end;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   100
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49586
diff changeset
   101
fun mk_case_conv_tac ctxt n uexhaust cases discss' selss =
52325
blanchet
parents: 51937
diff changeset
   102
  HEADGOAL (rtac uexhaust THEN'
blanchet
parents: 51937
diff changeset
   103
    EVERY' (map3 (fn casex => fn if_discs => fn sels =>
blanchet
parents: 51937
diff changeset
   104
        EVERY' [hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)),
blanchet
parents: 51937
diff changeset
   105
          rtac casex])
blanchet
parents: 51937
diff changeset
   106
      cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss));
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   107
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   108
fun mk_case_cong_tac ctxt uexhaust cases =
52325
blanchet
parents: 51937
diff changeset
   109
  HEADGOAL (rtac uexhaust THEN'
blanchet
parents: 51937
diff changeset
   110
    EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases));
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49032
diff changeset
   111
49586
d5e342ffe91e parameterized "subst_tac"
blanchet
parents: 49510
diff changeset
   112
(* TODO: More precise "simp_thms"; get rid of "blast_tac" *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   113
fun mk_split_tac ctxt uexhaust cases injectss distinctsss =
52325
blanchet
parents: 51937
diff changeset
   114
  HEADGOAL (rtac uexhaust) THEN
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51797
diff changeset
   115
  ALLGOALS (fn k => (hyp_subst_tac ctxt THEN'
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49463
diff changeset
   116
     simp_tac (ss_only (@{thms simp_thms} @ cases @ nth injectss (k - 1) @
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51686
diff changeset
   117
       flat (nth distinctsss (k - 1))) ctxt)) k) THEN
51937
db22d73e6c3e proper context;
wenzelm
parents: 51798
diff changeset
   118
  ALLGOALS (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   119
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   120
val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   121
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   122
fun mk_split_asm_tac ctxt split =
52325
blanchet
parents: 51937
diff changeset
   123
  HEADGOAL (rtac (split RS trans)) THEN unfold_thms_tac ctxt split_asm_thms THEN
blanchet
parents: 51937
diff changeset
   124
  HEADGOAL (rtac refl);
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   125
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff changeset
   126
end;