| author | wenzelm | 
| Tue, 20 Nov 2012 22:52:04 +0100 | |
| changeset 50136 | a96bd08258a2 | 
| parent 49623 | 1a0f25c38629 | 
| child 51761 | 4c9f08836d87 | 
| permissions | -rw-r--r-- | 
| 49509 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 blanchet parents: 
49506diff
changeset | 1 | (* Title: HOL/BNF/Tools/bnf_def_tactics.ML | 
| 49284 | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 3 | Author: Jasmin Blanchette, TU Muenchen | |
| 4 | Copyright 2012 | |
| 5 | ||
| 6 | Tactics for definition of bounded natural functors. | |
| 7 | *) | |
| 8 | ||
| 9 | signature BNF_DEF_TACTICS = | |
| 10 | sig | |
| 49623 | 11 | val mk_collect_set_natural_tac: thm list -> tactic | 
| 49284 | 12 | val mk_id': thm -> thm | 
| 13 | val mk_comp': thm -> thm | |
| 14 | val mk_in_mono_tac: int -> tactic | |
| 15 | val mk_map_wppull_tac: thm -> thm -> thm -> thm -> thm list -> tactic | |
| 16 | val mk_set_natural': thm -> thm | |
| 17 | ||
| 49506 | 18 | val mk_srel_Gr_tac: thm list -> thm -> thm -> thm -> thm -> thm list -> | 
| 49284 | 19 |     {prems: thm list, context: Proof.context} -> tactic
 | 
| 49506 | 20 |   val mk_srel_Id_tac: int -> thm -> thm -> {prems: 'a, context: Proof.context} -> tactic
 | 
| 21 | val mk_srel_O_tac: thm list -> thm -> thm -> thm -> thm -> thm list -> | |
| 49284 | 22 |     {prems: thm list, context: Proof.context} -> tactic
 | 
| 49506 | 23 |   val mk_in_srel_tac: thm list -> int -> {prems: 'b, context: Proof.context} -> tactic
 | 
| 24 | val mk_srel_converse_tac: thm -> tactic | |
| 25 | val mk_srel_converse_le_tac: thm list -> thm -> thm -> thm -> thm list -> | |
| 49284 | 26 |     {prems: thm list, context: Proof.context} -> tactic
 | 
| 49506 | 27 |   val mk_srel_mono_tac: thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
 | 
| 49284 | 28 | end; | 
| 29 | ||
| 30 | structure BNF_Def_Tactics : BNF_DEF_TACTICS = | |
| 31 | struct | |
| 32 | ||
| 33 | open BNF_Util | |
| 34 | open BNF_Tactics | |
| 35 | ||
| 36 | fun mk_id' id = mk_trans (fun_cong OF [id]) @{thm id_apply};
 | |
| 37 | fun mk_comp' comp = @{thm o_eq_dest_lhs} OF [mk_sym comp];
 | |
| 38 | fun mk_set_natural' set_natural = set_natural RS @{thm pointfreeE};
 | |
| 49490 | 39 | fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1 | 
| 49284 | 40 | else (rtac subsetI THEN' | 
| 41 | rtac CollectI) 1 THEN | |
| 42 | REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN | |
| 43 | REPEAT_DETERM_N (n - 1) | |
| 44 | ((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN | |
| 45 | (etac subset_trans THEN' atac) 1; | |
| 46 | ||
| 49623 | 47 | fun mk_collect_set_natural_tac set_naturals = | 
| 48 |   (rtac (@{thm collect_o} RS trans) THEN' rtac @{thm arg_cong[of _ _ collect]} THEN'
 | |
| 49 | EVERY' (map (fn set_natural => | |
| 50 |     rtac (mk_trans @{thm image_insert} @{thm arg_cong2[of _ _ _ _ insert]}) THEN'
 | |
| 51 | rtac set_natural) set_naturals) THEN' | |
| 52 |   rtac @{thm image_empty}) 1;
 | |
| 49284 | 53 | |
| 54 | fun mk_map_wppull_tac map_id map_cong map_wpull map_comp set_naturals = | |
| 55 | if null set_naturals then | |
| 56 |     EVERY' [rtac @{thm wppull_id}, rtac map_wpull, rtac map_id, rtac map_id] 1
 | |
| 57 |   else EVERY' [REPEAT_DETERM o etac conjE, REPEAT_DETERM o dtac @{thm wppull_thePull},
 | |
| 58 |     REPEAT_DETERM o etac exE, rtac @{thm wpull_wppull}, rtac map_wpull,
 | |
| 59 |     REPEAT_DETERM o rtac @{thm wpull_thePull}, rtac ballI,
 | |
| 60 | REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac CollectI, | |
| 61 |     CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
 | |
| 62 |       rtac @{thm image_subsetI}, rtac conjunct1, etac bspec, etac set_mp, atac])
 | |
| 63 | set_naturals, | |
| 64 | CONJ_WRAP' (fn thm => EVERY' [rtac (map_comp RS trans), rtac (map_comp RS trans), | |
| 65 | rtac (map_comp RS trans RS sym), rtac map_cong, | |
| 66 | REPEAT_DETERM_N (length set_naturals) o EVERY' [rtac (o_apply RS trans), | |
| 67 | rtac (o_apply RS trans RS sym), rtac (o_apply RS trans), rtac thm, | |
| 68 | rtac conjunct2, etac bspec, etac set_mp, atac]]) [conjunct1, conjunct2]] 1; | |
| 69 | ||
| 49506 | 70 | fun mk_srel_Gr_tac srel_O_Grs map_id map_cong map_id' map_comp set_naturals | 
| 49284 | 71 |   {context = ctxt, prems = _} =
 | 
| 72 | let | |
| 73 | val n = length set_naturals; | |
| 74 | in | |
| 75 | if null set_naturals then | |
| 49506 | 76 |       unfold_thms_tac ctxt srel_O_Grs THEN EVERY' [rtac @{thm Gr_UNIV_id}, rtac map_id] 1
 | 
| 77 |     else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
 | |
| 49284 | 78 | EVERY' [rtac equalityI, rtac subsetI, | 
| 79 |         REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 80 | REPEAT_DETERM o dtac Pair_eqD, | 
| 49284 | 81 | REPEAT_DETERM o etac conjE, hyp_subst_tac, | 
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 82 | rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, | 
| 49284 | 83 | rtac sym, rtac trans, rtac map_comp, rtac map_cong, | 
| 84 |         REPEAT_DETERM_N n o EVERY' [dtac @{thm set_rev_mp}, atac,
 | |
| 85 | REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac, | |
| 86 |           rtac (o_apply RS trans), rtac (@{thm fst_conv} RS arg_cong RS trans),
 | |
| 87 |           rtac (@{thm snd_conv} RS sym)],
 | |
| 88 | rtac CollectI, | |
| 89 |         CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
 | |
| 90 |           rtac @{thm image_subsetI}, dtac @{thm set_rev_mp}, atac,
 | |
| 91 | REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac, | |
| 92 |           stac @{thm fst_conv}, atac]) set_naturals,
 | |
| 93 |         rtac @{thm subrelI}, etac CollectE, REPEAT_DETERM o eresolve_tac [exE, conjE],
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 94 | REPEAT_DETERM o dtac Pair_eqD, | 
| 49284 | 95 | REPEAT_DETERM o etac conjE, hyp_subst_tac, | 
| 96 |         rtac @{thm relcompI}, rtac @{thm converseI},
 | |
| 49495 | 97 | EVERY' (map2 (fn convol => fn map_id => | 
| 98 | EVERY' [rtac CollectI, rtac exI, rtac conjI, | |
| 99 | rtac Pair_eqI, rtac conjI, rtac refl, rtac sym, | |
| 100 | rtac (box_equals OF [map_cong, map_comp RS sym, map_id]), | |
| 101 | REPEAT_DETERM_N n o rtac (convol RS fun_cong), | |
| 102 | REPEAT_DETERM o eresolve_tac [CollectE, conjE], | |
| 103 | rtac CollectI, | |
| 104 | CONJ_WRAP' (fn thm => | |
| 105 |               EVERY' [rtac @{thm ord_eq_le_trans}, rtac thm, rtac @{thm image_subsetI},
 | |
| 106 |                 rtac @{thm convol_memI[of id _ "%x. x", OF id_apply refl]}, etac set_mp, atac])
 | |
| 107 | set_naturals]) | |
| 108 |           @{thms fst_convol snd_convol} [map_id', refl])] 1
 | |
| 49284 | 109 | end; | 
| 110 | ||
| 49506 | 111 | fun mk_srel_Id_tac n srel_Gr map_id {context = ctxt, prems = _} =
 | 
| 112 |   unfold_thms_tac ctxt [srel_Gr, @{thm Id_alt}] THEN
 | |
| 49621 | 113 | (if n = 0 then rtac refl 1 | 
| 114 |   else EVERY' [rtac @{thm arg_cong2[of _ _ _ _ Gr]},
 | |
| 115 | rtac equalityI, rtac subset_UNIV, rtac subsetI, rtac CollectI, | |
| 116 | CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac map_id] 1); | |
| 49284 | 117 | |
| 49506 | 118 | fun mk_srel_mono_tac srel_O_Grs in_mono {context = ctxt, prems = _} =
 | 
| 119 | unfold_thms_tac ctxt srel_O_Grs THEN | |
| 49284 | 120 |     EVERY' [rtac @{thm relcomp_mono}, rtac @{thm iffD2[OF converse_mono]},
 | 
| 121 |       rtac @{thm Gr_mono}, rtac in_mono, REPEAT_DETERM o atac,
 | |
| 122 |       rtac @{thm Gr_mono}, rtac in_mono, REPEAT_DETERM o atac] 1;
 | |
| 123 | ||
| 49506 | 124 | fun mk_srel_converse_le_tac srel_O_Grs srel_Id map_cong map_comp set_naturals | 
| 49284 | 125 |   {context = ctxt, prems = _} =
 | 
| 126 | let | |
| 127 | val n = length set_naturals; | |
| 128 | in | |
| 129 | if null set_naturals then | |
| 49506 | 130 |       unfold_thms_tac ctxt [srel_Id] THEN rtac equalityD2 1 THEN rtac @{thm converse_Id} 1
 | 
| 131 |     else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
 | |
| 49284 | 132 |       EVERY' [rtac @{thm subrelI},
 | 
| 133 |         REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 134 | REPEAT_DETERM o dtac Pair_eqD, | 
| 49284 | 135 |         REPEAT_DETERM o etac conjE, hyp_subst_tac, rtac @{thm converseI},
 | 
| 136 |         rtac @{thm relcompI}, rtac @{thm converseI},
 | |
| 137 | EVERY' (map (fn thm => EVERY' [rtac CollectI, rtac exI, | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 138 | rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, rtac trans, | 
| 49284 | 139 | rtac map_cong, REPEAT_DETERM_N n o rtac thm, | 
| 140 | rtac (map_comp RS sym), rtac CollectI, | |
| 141 |           CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
 | |
| 142 |             etac @{thm flip_rel}]) set_naturals]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1
 | |
| 143 | end; | |
| 144 | ||
| 49506 | 145 | fun mk_srel_converse_tac le_converse = | 
| 49284 | 146 |   EVERY' [rtac equalityI, rtac le_converse, rtac @{thm xt1(6)}, rtac @{thm converse_shift},
 | 
| 147 |     rtac le_converse, REPEAT_DETERM o stac @{thm converse_converse}, rtac subset_refl] 1;
 | |
| 148 | ||
| 49506 | 149 | fun mk_srel_O_tac srel_O_Grs srel_Id map_cong map_wppull map_comp set_naturals | 
| 49284 | 150 |   {context = ctxt, prems = _} =
 | 
| 151 | let | |
| 152 | val n = length set_naturals; | |
| 153 | fun in_tac nthO_in = rtac CollectI THEN' | |
| 154 |         CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
 | |
| 155 |           rtac @{thm image_subsetI}, rtac nthO_in, etac set_mp, atac]) set_naturals;
 | |
| 156 | in | |
| 49506 | 157 |     if null set_naturals then unfold_thms_tac ctxt [srel_Id] THEN rtac (@{thm Id_O_R} RS sym) 1
 | 
| 158 |     else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
 | |
| 49284 | 159 |       EVERY' [rtac equalityI, rtac @{thm subrelI},
 | 
| 160 |         REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 161 | REPEAT_DETERM o dtac Pair_eqD, | 
| 49284 | 162 | REPEAT_DETERM o etac conjE, hyp_subst_tac, | 
| 163 |         rtac @{thm relcompI}, rtac @{thm relcompI}, rtac @{thm converseI},
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 164 | rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, | 
| 49284 | 165 | rtac sym, rtac trans, rtac map_comp, rtac sym, rtac map_cong, | 
| 166 |         REPEAT_DETERM_N n o rtac @{thm fst_fstO},
 | |
| 167 |         in_tac @{thm fstO_in},
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 168 | rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, | 
| 49284 | 169 | rtac sym, rtac trans, rtac map_comp, rtac map_cong, | 
| 170 | REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, rtac ballE, rtac subst, | |
| 171 |           rtac @{thm csquare_def}, rtac @{thm csquare_fstO_sndO}, atac, etac notE,
 | |
| 172 | etac set_mp, atac], | |
| 173 |         in_tac @{thm fstO_in},
 | |
| 174 |         rtac @{thm relcompI}, rtac @{thm converseI},
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 175 | rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, | 
| 49284 | 176 | rtac sym, rtac trans, rtac map_comp, rtac map_cong, | 
| 177 | REPEAT_DETERM_N n o rtac o_apply, | |
| 178 |         in_tac @{thm sndO_in},
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 179 | rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, | 
| 49284 | 180 | rtac sym, rtac trans, rtac map_comp, rtac sym, rtac map_cong, | 
| 181 |         REPEAT_DETERM_N n o rtac @{thm snd_sndO},
 | |
| 182 |         in_tac @{thm sndO_in},
 | |
| 183 |         rtac @{thm subrelI},
 | |
| 184 |         REPEAT_DETERM o eresolve_tac [CollectE, @{thm relcompE}, @{thm converseE}],
 | |
| 185 | REPEAT_DETERM o eresolve_tac [exE, conjE], | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 186 | REPEAT_DETERM o dtac Pair_eqD, | 
| 49284 | 187 | REPEAT_DETERM o etac conjE, hyp_subst_tac, | 
| 188 |         rtac allE, rtac subst, rtac @{thm wppull_def}, rtac map_wppull,
 | |
| 189 |         CONJ_WRAP' (K (rtac @{thm wppull_fstO_sndO})) set_naturals,
 | |
| 190 | etac allE, etac impE, etac conjI, etac conjI, atac, | |
| 191 | REPEAT_DETERM o eresolve_tac [bexE, conjE], | |
| 192 |         rtac @{thm relcompI}, rtac @{thm converseI},
 | |
| 193 | EVERY' (map (fn thm => EVERY' [rtac CollectI, rtac exI, | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49463diff
changeset | 194 | rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, rtac sym, rtac trans, | 
| 49284 | 195 | rtac trans, rtac map_cong, REPEAT_DETERM_N n o rtac thm, | 
| 196 |           rtac (map_comp RS sym), atac, atac]) [@{thm fst_fstO}, @{thm snd_sndO}])] 1
 | |
| 197 | end; | |
| 198 | ||
| 49506 | 199 | fun mk_in_srel_tac srel_O_Grs m {context = ctxt, prems = _} =
 | 
| 49284 | 200 | let | 
| 201 | val ls' = replicate (Int.max (1, m)) (); | |
| 202 | in | |
| 49506 | 203 | unfold_thms_tac ctxt (srel_O_Grs @ | 
| 49284 | 204 |       @{thms Gr_def converse_unfold relcomp_unfold mem_Collect_eq prod.cases Pair_eq}) THEN
 | 
| 205 | EVERY' [rtac iffI, REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac, rtac exI, | |
| 206 | rtac conjI, CONJ_WRAP' (K atac) ls', rtac conjI, rtac refl, rtac refl, | |
| 207 | REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI, rtac conjI, | |
| 208 |       REPEAT_DETERM_N 2 o EVERY' [rtac exI, rtac conjI, etac @{thm conjI[OF refl sym]},
 | |
| 209 | CONJ_WRAP' (K atac) ls']] 1 | |
| 210 | end; | |
| 211 | ||
| 212 | end; |