| author | hoelzl |
| Fri, 16 Jan 2015 10:59:15 +0100 | |
| changeset 59415 | 854fe701c984 |
| parent 59271 | c448752e8b9d |
| child 59498 | 50b60f501b05 |
| permissions | -rw-r--r-- |
| 54701 | 1 |
(* Title: HOL/Tools/Ctr_Sugar/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 |
| 54397 | 3 |
Copyright 2012, 2013 |
|
49020
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 |
|
|
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
8 |
signature CTR_SUGAR_GENERAL_TACTICS = |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
9 |
sig |
|
54905
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents:
54844
diff
changeset
|
10 |
val clean_blast_tac: Proof.context -> int -> tactic |
|
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
11 |
val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
12 |
val unfold_thms_tac: Proof.context -> thm list -> tactic |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
13 |
end; |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
14 |
|
| 54006 | 15 |
signature CTR_SUGAR_TACTICS = |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
16 |
sig |
|
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
17 |
include CTR_SUGAR_GENERAL_TACTICS |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
18 |
|
|
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
|
19 |
val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
20 |
val mk_case_tac: Proof.context -> int -> int -> thm -> thm list -> thm list list -> tactic |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51686
diff
changeset
|
21 |
val mk_case_cong_tac: Proof.context -> thm -> thm list -> tactic |
| 59266 | 22 |
val mk_case_distrib_tac: Proof.context -> cterm -> thm -> thm list -> tactic |
| 54491 | 23 |
val mk_case_eq_if_tac: Proof.context -> int -> thm -> thm list -> thm list list -> |
24 |
thm list list -> tactic |
|
| 49118 | 25 |
val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic |
| 59271 | 26 |
val mk_disc_eq_case_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> |
27 |
tactic |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
28 |
val mk_exhaust_disc_tac: int -> thm -> thm list -> tactic |
|
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
29 |
val mk_exhaust_sel_tac: int -> thm -> thm list -> tactic |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51686
diff
changeset
|
30 |
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
|
31 |
thm list list list -> thm list list list -> tactic |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
32 |
val mk_half_distinct_disc_tac: Proof.context -> int -> thm -> thm -> tactic |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
33 |
val mk_nchotomy_tac: int -> thm -> tactic |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
34 |
val mk_other_half_distinct_disc_tac: thm -> tactic |
| 54491 | 35 |
val mk_split_tac: Proof.context -> thm -> thm list -> thm list list -> thm list list -> |
36 |
thm list list list -> tactic |
|
| 49044 | 37 |
val mk_split_asm_tac: Proof.context -> thm -> tactic |
| 49137 | 38 |
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
|
39 |
end; |
|
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
40 |
|
| 54006 | 41 |
structure Ctr_Sugar_Tactics : CTR_SUGAR_TACTICS = |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
42 |
struct |
|
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
43 |
|
|
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
44 |
open Ctr_Sugar_Util |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
45 |
|
| 49486 | 46 |
val meta_mp = @{thm meta_mp};
|
47 |
||
|
54905
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents:
54844
diff
changeset
|
48 |
fun clean_blast_tac ctxt = blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt);
|
|
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents:
54844
diff
changeset
|
49 |
|
|
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
50 |
fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl, |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
51 |
tac, REPEAT_DETERM_N (n - k) o etac thin_rl]); |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
52 |
|
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
53 |
fun unfold_thms_tac _ [] = all_tac |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
54 |
| unfold_thms_tac ctxt thms = Local_Defs.unfold_tac ctxt (distinct Thm.eq_thm_prop thms); |
|
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
55 |
|
| 49212 | 56 |
fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P});
|
| 49031 | 57 |
|
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
58 |
fun mk_nchotomy_tac n exhaust = |
| 52325 | 59 |
HEADGOAL (rtac allI THEN' rtac exhaust THEN' |
| 57525 | 60 |
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
|
61 |
|
| 49486 | 62 |
fun mk_unique_disc_def_tac m uexhaust = |
| 52325 | 63 |
HEADGOAL (EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl]); |
| 49137 | 64 |
|
| 49486 | 65 |
fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust = |
| 52325 | 66 |
HEADGOAL (EVERY' ([rtac (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans),
|
| 49630 | 67 |
rtac @{thm iffI_np}, REPEAT_DETERM o etac exE,
|
| 51798 | 68 |
hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI, |
| 49486 | 69 |
rtac distinct, rtac uexhaust] @ |
|
49174
41790d616f63
by default, only generate one discriminator for a two-value datatype
blanchet
parents:
49168
diff
changeset
|
70 |
(([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac]) |
| 52325 | 71 |
|> 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
|
72 |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
73 |
fun mk_half_distinct_disc_tac ctxt m discD disc' = |
| 52325 | 74 |
HEADGOAL (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN' |
75 |
rtac disc'); |
|
| 49028 | 76 |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
77 |
fun mk_other_half_distinct_disc_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half);
|
| 49028 | 78 |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
79 |
fun mk_exhaust_disc_or_sel_tac n exhaust destIs = |
| 52325 | 80 |
HEADGOAL (rtac exhaust THEN' |
| 53916 | 81 |
EVERY' (map2 (fn k => fn destI => dtac destI THEN' |
82 |
select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) destIs)); |
|
83 |
||
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
84 |
val mk_exhaust_disc_tac = mk_exhaust_disc_or_sel_tac; |
| 53916 | 85 |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
86 |
fun mk_exhaust_sel_tac n exhaust_disc collapses = |
|
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
87 |
mk_exhaust_disc_or_sel_tac n exhaust_disc collapses ORELSE |
|
53920
c6de7f20c845
made tactic more robust in case somebody specified a discriminator for a one-constructor type
blanchet
parents:
53917
diff
changeset
|
88 |
HEADGOAL (etac meta_mp THEN' resolve_tac collapses); |
| 49029 | 89 |
|
| 49484 | 90 |
fun mk_collapse_tac ctxt m discD sels = |
| 52325 | 91 |
HEADGOAL (dtac discD THEN' |
92 |
(if m = 0 then |
|
93 |
atac |
|
94 |
else |
|
95 |
REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN' |
|
96 |
SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl)); |
|
| 49030 | 97 |
|
| 59271 | 98 |
fun mk_disc_eq_case_tac ctxt ct exhaust discs distincts cases = |
99 |
HEADGOAL Goal.conjunction_tac THEN |
|
100 |
ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW |
|
101 |
(hyp_subst_tac ctxt THEN' |
|
102 |
SELECT_GOAL (unfold_thms_tac ctxt (@{thms not_True_eq_False not_False_eq_True} @ cases @
|
|
103 |
((refl :: discs @ distincts) RL [eqTrueI, eqFalseI]))) THEN' |
|
104 |
resolve_tac @{thms TrueI True_not_False False_not_True}));
|
|
105 |
||
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
106 |
fun mk_expand_tac ctxt n ms uexhaust_disc vexhaust_disc uncollapses distinct_discsss |
|
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
107 |
distinct_discsss' = |
| 49486 | 108 |
if ms = [0] then |
| 52325 | 109 |
HEADGOAL (rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN'
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
110 |
TRY o EVERY' [rtac uexhaust_disc, atac, rtac vexhaust_disc, atac]) |
| 49486 | 111 |
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
|
112 |
let val ks = 1 upto n in |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
113 |
HEADGOAL (rtac uexhaust_disc THEN' |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
57983
diff
changeset
|
114 |
EVERY' (@{map 5} (fn k => fn m => fn distinct_discss => fn distinct_discss' => fn uuncollapse =>
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
115 |
EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac, rtac sym, rtac vexhaust_disc, |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
57983
diff
changeset
|
116 |
EVERY' (@{map 4} (fn k' => fn distinct_discs => fn distinct_discs' => fn vuncollapse =>
|
| 52325 | 117 |
EVERY' |
118 |
(if k' = k then |
|
119 |
[rtac (vuncollapse RS trans), TRY o atac] @ |
|
120 |
(if m = 0 then |
|
121 |
[rtac refl] |
|
122 |
else |
|
123 |
[if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac], |
|
124 |
REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, |
|
125 |
asm_simp_tac (ss_only [] ctxt)]) |
|
126 |
else |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
127 |
[dtac (the_single (if k = n then distinct_discs else distinct_discs')), |
| 52325 | 128 |
etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
|
129 |
atac, atac])) |
|
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
130 |
ks distinct_discss distinct_discss' uncollapses)]) |
|
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57525
diff
changeset
|
131 |
ks ms distinct_discsss distinct_discsss' uncollapses)) |
| 49486 | 132 |
end; |
133 |
||
| 52967 | 134 |
fun mk_case_same_ctr_tac ctxt injects = |
135 |
REPEAT_DETERM o etac exE THEN' etac conjE THEN' |
|
136 |
(case injects of |
|
137 |
[] => atac |
|
138 |
| [inject] => dtac (inject RS iffD1) THEN' REPEAT_DETERM o etac conjE THEN' |
|
139 |
hyp_subst_tac ctxt THEN' rtac refl); |
|
140 |
||
141 |
fun mk_case_distinct_ctrs_tac ctxt distincts = |
|
142 |
REPEAT_DETERM o etac exE THEN' etac conjE THEN' full_simp_tac (ss_only distincts ctxt); |
|
143 |
||
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
144 |
fun mk_case_tac ctxt n k case_def injects distinctss = |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
145 |
let |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
146 |
val case_def' = mk_unabs_def (n + 1) (case_def RS meta_eq_to_obj_eq); |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
147 |
val ks = 1 upto n; |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
148 |
in |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
149 |
HEADGOAL (rtac (case_def' RS trans) THEN' rtac @{thm the_equality} THEN'
|
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
150 |
rtac (mk_disjIN n k) THEN' REPEAT_DETERM o rtac exI THEN' rtac conjI THEN' rtac refl THEN' |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52967
diff
changeset
|
151 |
rtac refl THEN' |
| 52967 | 152 |
EVERY' (map2 (fn k' => fn distincts => |
153 |
(if k' < n then etac disjE else K all_tac) THEN' |
|
154 |
(if k' = k then mk_case_same_ctr_tac ctxt injects |
|
155 |
else mk_case_distinct_ctrs_tac ctxt distincts)) ks distinctss)) |
|
156 |
end; |
|
157 |
||
| 59266 | 158 |
fun mk_case_distrib_tac ctxt ct exhaust cases = |
159 |
HEADGOAL (rtac (cterm_instantiate_pos [SOME ct] exhaust)) THEN |
|
160 |
ALLGOALS (hyp_subst_tac ctxt THEN' SELECT_GOAL (unfold_thms_tac ctxt cases) THEN' rtac refl); |
|
161 |
||
| 54491 | 162 |
fun mk_case_cong_tac ctxt uexhaust cases = |
163 |
HEADGOAL (rtac uexhaust THEN' |
|
164 |
EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases)); |
|
165 |
||
166 |
fun mk_case_eq_if_tac ctxt n uexhaust cases discss' selss = |
|
| 52325 | 167 |
HEADGOAL (rtac uexhaust THEN' |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
57983
diff
changeset
|
168 |
EVERY' (@{map 3} (fn casex => fn if_discs => fn sels =>
|
| 52325 | 169 |
EVERY' [hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)), |
170 |
rtac casex]) |
|
171 |
cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss)); |
|
| 49031 | 172 |
|
| 53917 | 173 |
fun mk_split_tac ctxt uexhaust cases selss injectss distinctsss = |
| 52325 | 174 |
HEADGOAL (rtac uexhaust) THEN |
| 51798 | 175 |
ALLGOALS (fn k => (hyp_subst_tac ctxt THEN' |
| 53917 | 176 |
simp_tac (ss_only (@{thms simp_thms} @ cases @ nth selss (k - 1) @ nth injectss (k - 1) @
|
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51686
diff
changeset
|
177 |
flat (nth distinctsss (k - 1))) ctxt)) k) THEN |
|
54905
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
blanchet
parents:
54844
diff
changeset
|
178 |
ALLGOALS (clean_blast_tac ctxt); |
| 49032 | 179 |
|
| 49044 | 180 |
val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
|
181 |
||
182 |
fun mk_split_asm_tac ctxt split = |
|
| 52325 | 183 |
HEADGOAL (rtac (split RS trans)) THEN unfold_thms_tac ctxt split_asm_thms THEN |
184 |
HEADGOAL (rtac refl); |
|
| 49044 | 185 |
|
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
diff
changeset
|
186 |
end; |
|
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
187 |
|
| 54491 | 188 |
structure Ctr_Sugar_General_Tactics : CTR_SUGAR_GENERAL_TACTICS = Ctr_Sugar_Tactics; |