author | blanchet |
Sat, 27 Apr 2013 11:37:50 +0200 | |
changeset 51797 | 182454c06a80 |
parent 51742 | src/HOL/BNF/Tools/bnf_wrap_tactics.ML@b5ff7393642d |
child 51798 | ad3a241def73 |
permissions | -rw-r--r-- |
51797 | 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 | 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 | 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 | 14 |
val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic |
49029 | 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 |
49122 | 18 |
val mk_half_disc_exclude_tac: 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 | 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 | 23 |
val mk_split_asm_tac: Proof.context -> thm -> tactic |
49137 | 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 | 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 | 30 |
open BNF_Util |
49029 | 31 |
open BNF_Tactics |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
32 |
|
49486 | 33 |
val meta_mp = @{thm meta_mp}; |
34 |
||
49212 | 35 |
fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P}); |
49031 | 36 |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
37 |
fun mk_nchotomy_tac n exhaust = |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
38 |
(rtac allI THEN' rtac exhaust THEN' |
49022 | 39 |
EVERY' (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n))) 1; |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
40 |
|
49486 | 41 |
fun mk_unique_disc_def_tac m uexhaust = |
42 |
EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl] 1; |
|
49137 | 43 |
|
49486 | 44 |
fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust = |
49630 | 45 |
EVERY' ([rtac (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans), |
46 |
rtac @{thm iffI_np}, REPEAT_DETERM o etac exE, |
|
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49486
diff
changeset
|
47 |
hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI, |
49486 | 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]) |
41790d616f63
by default, only generate one discriminator for a two-value datatype
blanchet
parents:
49168
diff
changeset
|
50 |
|> k = 1 ? swap |> op @)) 1; |
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 |
|
49484 | 52 |
fun mk_half_disc_exclude_tac m discD disc' = |
53 |
(dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN' rtac disc') 1; |
|
49028 | 54 |
|
49484 | 55 |
fun mk_other_half_disc_exclude_tac half = (etac @{thm contrapos_pn} THEN' etac half) 1; |
49028 | 56 |
|
49029 | 57 |
fun mk_disc_exhaust_tac n exhaust discIs = |
58 |
(rtac exhaust THEN' |
|
59 |
EVERY' (map2 (fn k => fn discI => |
|
49486 | 60 |
dtac discI THEN' select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) discIs)) 1; |
49029 | 61 |
|
49484 | 62 |
fun mk_collapse_tac ctxt m discD sels = |
49030 | 63 |
(dtac discD THEN' |
64 |
(if m = 0 then |
|
65 |
atac |
|
66 |
else |
|
49484 | 67 |
REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN' |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49486
diff
changeset
|
68 |
SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl)) 1; |
49030 | 69 |
|
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
|
70 |
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
|
71 |
disc_excludesss' = |
49486 | 72 |
if ms = [0] then |
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
|
73 |
(rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN' |
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
|
74 |
TRY o EVERY' [rtac udisc_exhaust, atac, rtac vdisc_exhaust, atac]) 1 |
49486 | 75 |
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
|
76 |
let val ks = 1 upto n in |
49486 | 77 |
(rtac udisc_exhaust THEN' |
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
|
78 |
EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' => |
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
|
79 |
fn uuncollapse => |
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
|
80 |
EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac, |
49486 | 81 |
rtac sym, rtac vdisc_exhaust, |
82 |
EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse => |
|
83 |
EVERY' |
|
84 |
(if k' = k then |
|
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
|
85 |
[rtac (vuncollapse RS trans), TRY o atac] @ |
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
|
86 |
(if m = 0 then |
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
|
87 |
[rtac refl] |
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
|
88 |
else |
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
|
89 |
[if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac], |
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
|
90 |
REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, |
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
|
91 |
asm_simp_tac (ss_only [] ctxt)]) |
49486 | 92 |
else |
93 |
[dtac (the_single (if k = n then disc_excludes else disc_excludes')), |
|
94 |
etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}), |
|
95 |
atac, atac])) |
|
96 |
ks disc_excludess disc_excludess' uncollapses)]) |
|
97 |
ks ms disc_excludesss disc_excludesss' uncollapses)) 1 |
|
98 |
end; |
|
99 |
||
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49586
diff
changeset
|
100 |
fun mk_case_conv_tac ctxt n uexhaust cases discss' selss = |
49486 | 101 |
(rtac uexhaust THEN' |
49484 | 102 |
EVERY' (map3 (fn casex => fn if_discs => fn sels => |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49486
diff
changeset
|
103 |
EVERY' [hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)), rtac casex]) |
49484 | 104 |
cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss)) 1; |
49031 | 105 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51686
diff
changeset
|
106 |
fun mk_case_cong_tac ctxt uexhaust cases = |
49486 | 107 |
(rtac uexhaust THEN' |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51686
diff
changeset
|
108 |
EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases)) 1; |
49043 | 109 |
|
51686 | 110 |
val naked_ctxt = @{theory_context HOL}; |
49043 | 111 |
|
49586 | 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 = |
49486 | 114 |
rtac uexhaust 1 THEN |
49157 | 115 |
ALLGOALS (fn k => (hyp_subst_tac THEN' |
49484 | 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 |
49045
7d9631754bba
minor fixes (for compatibility with existing datatype package)
blanchet
parents:
49044
diff
changeset
|
118 |
ALLGOALS (blast_tac naked_ctxt); |
49032 | 119 |
|
49044 | 120 |
val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex}; |
121 |
||
122 |
fun mk_split_asm_tac ctxt split = |
|
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49486
diff
changeset
|
123 |
rtac (split RS trans) 1 THEN unfold_thms_tac ctxt split_asm_thms THEN rtac refl 1; |
49044 | 124 |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
125 |
end; |