author | wenzelm |
Mon, 11 Sep 2023 19:30:48 +0200 | |
changeset 78659 | b5f3d1051b13 |
parent 75624 | 22d1c5f2b9f4 |
child 82248 | e8c96013ea8a |
permissions | -rw-r--r-- |
55061 | 1 |
(* Title: HOL/Tools/BNF/bnf_gfp_tactics.ML |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Dmitriy Traytel, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
3 |
Author: Andrei Popescu, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
5 |
Copyright 2012 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
7 |
Tactics for the codatatype construction. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
9 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
10 |
signature BNF_GFP_TACTICS = |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
11 |
sig |
55197 | 12 |
val mk_bis_Gr_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> thm list -> tactic |
51798 | 13 |
val mk_bis_O_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic |
55197 | 14 |
val mk_bis_Union_tac: Proof.context -> thm -> thm list -> tactic |
60728 | 15 |
val mk_bis_converse_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic |
16 |
val mk_bis_rel_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list -> |
|
51798 | 17 |
thm list list -> tactic |
55197 | 18 |
val mk_coalgT_tac: Proof.context -> int -> thm list -> thm list -> thm list list -> tactic |
60728 | 19 |
val mk_coalg_final_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list list -> |
20 |
thm list list -> tactic |
|
21 |
val mk_coalg_set_tac: Proof.context -> thm -> tactic |
|
55197 | 22 |
val mk_coind_wit_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> tactic |
60728 | 23 |
val mk_col_bd_tac: Proof.context -> int -> int -> cterm option list -> thm list -> thm list -> |
24 |
thm -> thm -> thm list list -> tactic |
|
55197 | 25 |
val mk_col_natural_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list -> |
26 |
thm list list -> tactic |
|
60728 | 27 |
val mk_congruent_str_final_tac: Proof.context -> int -> thm -> thm -> thm -> thm list -> tactic |
55197 | 28 |
val mk_corec_tac: Proof.context -> int -> thm list -> thm -> thm -> thm list -> tactic |
29 |
val mk_corec_unique_mor_tac: Proof.context -> thm list -> thm list -> thm -> tactic |
|
60728 | 30 |
val mk_dtor_coinduct_tac: Proof.context -> int -> thm -> thm -> thm list -> tactic |
58445 | 31 |
val mk_dtor_corec_transfer_tac: Proof.context -> int -> int -> thm list -> thm list -> thm list -> |
32 |
thm list -> tactic |
|
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
33 |
val mk_dtor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list -> |
51798 | 34 |
thm -> thm -> thm list -> thm list -> thm list list -> tactic |
55197 | 35 |
val mk_dtor_o_ctor_tac: Proof.context -> thm -> thm -> thm -> thm -> thm list -> tactic |
60728 | 36 |
val mk_equiv_lsbis_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> tactic |
56179 | 37 |
val mk_Jset_minimal_tac: Proof.context -> int -> thm -> tactic |
56113 | 38 |
val mk_col_minimal_tac: Proof.context -> int -> cterm option list -> thm list -> thm list -> |
55197 | 39 |
tactic |
60728 | 40 |
val mk_incl_lsbis_tac: Proof.context -> int -> int -> thm -> tactic |
41 |
val mk_length_Lev'_tac: Proof.context -> thm -> tactic |
|
51798 | 42 |
val mk_length_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic |
60728 | 43 |
val mk_map_comp0_tac: Proof.context -> thm list -> thm list -> thm -> tactic |
51798 | 44 |
val mk_mcong_tac: Proof.context -> int -> (int -> tactic) -> thm list -> thm list -> thm list -> |
55644 | 45 |
thm list list -> thm list list -> thm list list list -> thm list -> tactic |
60728 | 46 |
val mk_map_id0_tac: Proof.context -> thm list -> thm -> thm -> tactic |
47 |
val mk_map_tac: Proof.context -> int -> int -> thm -> thm -> thm -> thm -> tactic |
|
55197 | 48 |
val mk_dtor_map_unique_tac: Proof.context -> thm -> thm list -> tactic |
49 |
val mk_mor_Abs_tac: Proof.context -> thm list -> thm list -> tactic |
|
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
50 |
val mk_mor_Rep_tac: Proof.context -> thm list -> thm list -> thm list -> thm list list -> |
55197 | 51 |
thm list -> thm list -> tactic |
60728 | 52 |
val mk_mor_T_final_tac: Proof.context -> thm -> thm list -> thm list -> tactic |
53 |
val mk_mor_UNIV_tac: Proof.context -> thm list -> thm -> tactic |
|
55197 | 54 |
val mk_mor_beh_tac: Proof.context -> int -> thm -> thm -> thm list -> thm list -> thm list -> |
55 |
thm list -> thm list list -> thm list list -> thm list -> thm list -> thm list -> thm list -> |
|
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
56 |
thm list -> thm list -> thm list list -> thm list list list -> thm list list list -> |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
57 |
thm list list -> thm list -> thm list -> tactic |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
58 |
val mk_mor_comp_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> tactic |
60728 | 59 |
val mk_mor_elim_tac: Proof.context -> thm -> tactic |
60 |
val mk_mor_col_tac: Proof.context -> int -> int -> cterm option list -> int -> thm list -> |
|
61 |
thm list -> thm list -> thm list list -> thm list list -> tactic |
|
62 |
val mk_mor_incl_tac: Proof.context -> thm -> thm list -> tactic |
|
63 |
val mk_mor_str_tac: Proof.context -> 'a list -> thm -> tactic |
|
64 |
val mk_mor_unfold_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list -> |
|
65 |
thm list -> thm list -> thm list -> tactic |
|
66 |
val mk_raw_coind_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> |
|
67 |
thm -> thm list -> thm list -> thm list -> thm -> thm list -> tactic |
|
55197 | 68 |
val mk_rel_coinduct_tac: Proof.context -> thm list -> thm list -> thm list -> thm list list -> |
69 |
thm list -> thm list -> tactic |
|
55644 | 70 |
val mk_rel_coinduct_coind_tac: Proof.context -> bool -> int -> thm -> int list -> thm list -> |
71 |
thm list -> thm list -> thm list list -> thm list -> thm list -> thm list -> tactic |
|
55197 | 72 |
val mk_rel_coinduct_ind_tac: Proof.context -> int -> int list -> thm list -> thm list list -> |
73 |
int -> thm -> tactic |
|
60728 | 74 |
val mk_rv_last_tac: Proof.context -> ctyp option list -> cterm option list -> thm list -> |
75 |
thm list -> tactic |
|
51798 | 76 |
val mk_sbis_lsbis_tac: Proof.context -> thm list -> thm -> thm -> tactic |
77 |
val mk_set_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list -> |
|
78 |
thm list -> thm list list -> tactic |
|
75624 | 79 |
val mk_set_bd_tac: Proof.context -> thm -> thm -> thm -> thm -> tactic |
60728 | 80 |
val mk_set_Jset_incl_Jset_tac: Proof.context -> int -> thm -> int -> tactic |
51798 | 81 |
val mk_set_image_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> |
82 |
thm list -> thm list -> thm list list -> thm list list -> tactic |
|
60728 | 83 |
val mk_set_incl_Jset_tac: Proof.context -> thm -> tactic |
84 |
val mk_set_ge_tac: Proof.context -> int -> thm -> thm list -> tactic |
|
85 |
val mk_set_le_tac: Proof.context -> int -> thm -> thm list -> thm list list -> tactic |
|
86 |
val mk_set_map0_tac: Proof.context -> thm -> tactic |
|
87 |
val mk_unfold_unique_mor_tac: Proof.context -> thm list -> thm -> thm -> thm list -> tactic |
|
55197 | 88 |
val mk_unfold_transfer_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic |
89 |
val mk_wit_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm list -> tactic |
|
60728 | 90 |
val mk_le_rel_OO_tac: Proof.context -> thm -> thm list -> thm list -> tactic |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
91 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
92 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
93 |
structure BNF_GFP_Tactics : BNF_GFP_TACTICS = |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
94 |
struct |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
95 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
96 |
open BNF_Tactics |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
97 |
open BNF_Util |
51850
106afdf5806c
renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
blanchet
parents:
51798
diff
changeset
|
98 |
open BNF_FP_Util |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
99 |
open BNF_GFP_Util |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
100 |
|
63170 | 101 |
val fst_convol_fun_cong_sym = @{thm fst_convol[simplified convol_def]} RS fun_cong RS sym; |
49488
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
traytel
parents:
49463
diff
changeset
|
102 |
val list_inject_iffD1 = @{thm list.inject[THEN iffD1]}; |
49305 | 103 |
val nat_induct = @{thm nat_induct}; |
104 |
val o_apply_trans_sym = o_apply RS trans RS sym; |
|
105 |
val ord_eq_le_trans = @{thm ord_eq_le_trans}; |
|
106 |
val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans}; |
|
63170 | 107 |
val snd_convol_fun_cong_sym = @{thm snd_convol[simplified convol_def]} RS fun_cong RS sym; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57806
diff
changeset
|
108 |
val sum_case_cong_weak = @{thm sum.case_cong_weak}; |
49305 | 109 |
val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]}; |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
110 |
val Collect_splitD_set_mp = @{thm Collect_case_prodD[OF set_mp]}; |
52659
58b87aa4dc3b
eliminate duplicated theorems (thanks to "Auto solve_direct" in jEdit)
traytel
parents:
52635
diff
changeset
|
111 |
val rev_bspec = Drule.rotate_prems 1 bspec; |
67399 | 112 |
val Un_cong = @{thm arg_cong2[of _ _ _ _ "(\<union>)"]}; |
52749 | 113 |
val converse_shift = @{thm converse_subset_swap} RS iffD1; |
49305 | 114 |
|
60728 | 115 |
fun mk_coalg_set_tac ctxt coalg_def = |
116 |
dtac ctxt (coalg_def RS iffD1) 1 THEN |
|
117 |
REPEAT_DETERM (etac ctxt conjE 1) THEN |
|
60752 | 118 |
EVERY' [dtac ctxt rev_bspec, assume_tac ctxt] 1 THEN |
60757 | 119 |
REPEAT_DETERM (eresolve_tac ctxt [CollectE, conjE] 1) THEN assume_tac ctxt 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
120 |
|
60728 | 121 |
fun mk_mor_elim_tac ctxt mor_def = |
122 |
(dtac ctxt (mor_def RS iffD1) THEN' |
|
123 |
REPEAT o etac ctxt conjE THEN' |
|
124 |
TRY o rtac ctxt @{thm image_subsetI} THEN' |
|
125 |
etac ctxt bspec THEN' |
|
60752 | 126 |
assume_tac ctxt) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
127 |
|
60728 | 128 |
fun mk_mor_incl_tac ctxt mor_def map_ids = |
129 |
(rtac ctxt (mor_def RS iffD2) THEN' |
|
130 |
rtac ctxt conjI THEN' |
|
131 |
CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, etac ctxt set_mp, etac ctxt (id_apply RS @{thm ssubst_mem})])) |
|
56114 | 132 |
map_ids THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
133 |
CONJ_WRAP' (fn thm => |
60728 | 134 |
(EVERY' [rtac ctxt ballI, rtac ctxt (thm RS trans), rtac ctxt sym, rtac ctxt (id_apply RS arg_cong)])) map_ids) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
135 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
136 |
fun mk_mor_comp_tac ctxt mor_def mor_images morEs map_comp_ids = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
137 |
let |
60728 | 138 |
fun fbetw_tac image = EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS @{thm ssubst_mem}), etac ctxt image, |
60752 | 139 |
etac ctxt image, assume_tac ctxt]; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
140 |
fun mor_tac ((mor_image, morE), map_comp_id) = |
60728 | 141 |
EVERY' [rtac ctxt ballI, stac ctxt o_apply, rtac ctxt trans, rtac ctxt (map_comp_id RS sym), rtac ctxt trans, |
60752 | 142 |
etac ctxt (morE RS arg_cong), assume_tac ctxt, etac ctxt morE, |
143 |
etac ctxt mor_image, assume_tac ctxt]; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
144 |
in |
60728 | 145 |
(rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
146 |
CONJ_WRAP' fbetw_tac mor_images THEN' |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
147 |
CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
148 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
149 |
|
60728 | 150 |
fun mk_mor_UNIV_tac ctxt morEs mor_def = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
151 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
152 |
val n = length morEs; |
60728 | 153 |
fun mor_tac morE = EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, rtac ctxt trans, etac ctxt morE, |
154 |
rtac ctxt UNIV_I, rtac ctxt sym, rtac ctxt o_apply]; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
155 |
in |
60728 | 156 |
EVERY' [rtac ctxt iffI, CONJ_WRAP' mor_tac morEs, |
157 |
rtac ctxt (mor_def RS iffD2), rtac ctxt conjI, CONJ_WRAP' (K (rtac ctxt ballI THEN' rtac ctxt UNIV_I)) morEs, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
158 |
CONJ_WRAP' (fn i => |
60728 | 159 |
EVERY' [dtac ctxt (mk_conjunctN n i), rtac ctxt ballI, etac ctxt @{thm comp_eq_dest}]) (1 upto n)] 1 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
160 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
161 |
|
60728 | 162 |
fun mk_mor_str_tac ctxt ks mor_UNIV = |
163 |
(rtac ctxt (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac ctxt refl)) ks) 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
164 |
|
60728 | 165 |
fun mk_set_incl_Jset_tac ctxt rec_Suc = |
166 |
EVERY' (map (rtac ctxt) [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1}, sym, |
|
56113 | 167 |
rec_Suc]) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
168 |
|
60728 | 169 |
fun mk_set_Jset_incl_Jset_tac ctxt n rec_Suc i = |
170 |
EVERY' (map (rtac ctxt) [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, |
|
60752 | 171 |
rec_Suc, UnI2, mk_UnIN n i] @ [etac ctxt @{thm UN_I}, assume_tac ctxt]) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
172 |
|
56113 | 173 |
fun mk_col_minimal_tac ctxt m cts rec_0s rec_Sucs = |
60801 | 174 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 175 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
176 |
CONJ_WRAP' (fn thm => EVERY' |
60728 | 177 |
[rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt @{thm empty_subsetI}]) rec_0s, |
178 |
REPEAT_DETERM o rtac ctxt allI, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
179 |
CONJ_WRAP' (fn rec_Suc => EVERY' |
60728 | 180 |
[rtac ctxt ord_eq_le_trans, rtac ctxt rec_Suc, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
181 |
if m = 0 then K all_tac |
60728 | 182 |
else (rtac ctxt @{thm Un_least} THEN' Goal.assume_rule_tac ctxt), |
183 |
CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least})) |
|
184 |
(K (EVERY' [rtac ctxt @{thm UN_least}, REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], |
|
60752 | 185 |
rtac ctxt subset_trans, assume_tac ctxt, Goal.assume_rule_tac ctxt])) rec_0s]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
186 |
rec_Sucs] 1; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
187 |
|
56179 | 188 |
fun mk_Jset_minimal_tac ctxt n col_minimal = |
60728 | 189 |
(CONJ_WRAP' (K (EVERY' [rtac ctxt @{thm UN_least}, rtac ctxt rev_mp, rtac ctxt col_minimal, |
190 |
EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac ctxt impI, |
|
60752 | 191 |
REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], assume_tac ctxt])) (1 upto n)) 1 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
192 |
|
60728 | 193 |
fun mk_mor_col_tac ctxt m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss = |
60801 | 194 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 195 |
REPEAT_DETERM o rtac ctxt allI, |
196 |
CONJ_WRAP' (fn thm => EVERY' (map (rtac ctxt) [impI, thm RS trans, thm RS sym])) rec_0s, |
|
197 |
REPEAT_DETERM o rtac ctxt allI, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
198 |
CONJ_WRAP' |
53289 | 199 |
(fn (rec_Suc, (morE, ((passive_set_map0s, active_set_map0s), coalg_sets))) => |
60728 | 200 |
EVERY' [rtac ctxt impI, rtac ctxt (rec_Suc RS trans), rtac ctxt (rec_Suc RS trans RS sym), |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
201 |
if m = 0 then K all_tac |
60728 | 202 |
else EVERY' [rtac ctxt Un_cong, rtac ctxt @{thm box_equals}, |
203 |
rtac ctxt (nth passive_set_map0s (j - 1) RS sym), |
|
60752 | 204 |
rtac ctxt trans_fun_cong_image_id_id_apply, etac ctxt (morE RS arg_cong), |
205 |
assume_tac ctxt], |
|
60728 | 206 |
CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 Un_cong)) |
53289 | 207 |
(fn (i, (set_map0, coalg_set)) => |
60728 | 208 |
EVERY' [rtac ctxt sym, rtac ctxt trans, rtac ctxt (refl RSN (2, @{thm SUP_cong})), |
60752 | 209 |
etac ctxt (morE RS sym RS arg_cong RS trans), assume_tac ctxt, rtac ctxt set_map0, |
60728 | 210 |
rtac ctxt (@{thm UN_simps(10)} RS trans), rtac ctxt (refl RS @{thm SUP_cong}), |
60777 | 211 |
forward_tac ctxt [coalg_set], assume_tac ctxt, dtac ctxt set_mp, assume_tac ctxt, |
60752 | 212 |
rtac ctxt mp, rtac ctxt (mk_conjunctN n i), |
213 |
REPEAT_DETERM o etac ctxt allE, assume_tac ctxt, assume_tac ctxt]) |
|
53289 | 214 |
(rev ((1 upto n) ~~ (active_set_map0s ~~ coalg_sets)))]) |
215 |
(rec_Sucs ~~ (morEs ~~ (map (chop m) set_map0ss ~~ map (drop m) coalg_setss)))] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
216 |
|
60728 | 217 |
fun mk_bis_rel_tac ctxt m bis_def in_rels map_comp0s map_cong0s set_map0ss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
218 |
let |
56017 | 219 |
val n = length in_rels; |
220 |
val thms = ((1 upto n) ~~ map_comp0s ~~ map_cong0s ~~ set_map0ss ~~ in_rels); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
221 |
|
56017 | 222 |
fun mk_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) = |
60728 | 223 |
EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, dtac ctxt (mk_conjunctN n i), |
60752 | 224 |
etac ctxt allE, etac ctxt allE, etac ctxt impE, assume_tac ctxt, etac ctxt bexE, |
60757 | 225 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], |
60728 | 226 |
rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt (Drule.rotate_prems 1 conjI), |
227 |
CONJ_WRAP' (fn thm => EVERY' [rtac ctxt trans, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt map_cong0, |
|
228 |
REPEAT_DETERM_N m o rtac ctxt thm, REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), |
|
60752 | 229 |
assume_tac ctxt]) |
56017 | 230 |
@{thms fst_diag_id snd_diag_id}, |
60728 | 231 |
rtac ctxt CollectI, |
56017 | 232 |
CONJ_WRAP' (fn (i, thm) => |
233 |
if i <= m |
|
60728 | 234 |
then EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt subset_trans, |
235 |
etac ctxt @{thm image_mono}, rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, |
|
236 |
rtac ctxt @{thm case_prodI}, rtac ctxt refl] |
|
237 |
else EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt trans, rtac ctxt thm, |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
238 |
rtac ctxt trans_fun_cong_image_id_id_apply, etac ctxt @{thm Collect_case_prod_in_rel_leI}]) |
56017 | 239 |
(1 upto (m + n) ~~ set_map0s)]; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
240 |
|
56017 | 241 |
fun mk_only_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) = |
60728 | 242 |
EVERY' [dtac ctxt (mk_conjunctN n i), rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, |
60752 | 243 |
etac ctxt allE, etac ctxt allE, etac ctxt impE, assume_tac ctxt, |
60728 | 244 |
dtac ctxt (in_rel RS @{thm iffD1}), |
60757 | 245 |
REPEAT_DETERM o eresolve_tac ctxt ([CollectE, conjE, exE] @ |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
246 |
@{thms CollectE Collect_case_prod_in_rel_leE}), |
60728 | 247 |
rtac ctxt bexI, rtac ctxt conjI, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, |
248 |
REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong), |
|
249 |
REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), |
|
60752 | 250 |
assume_tac ctxt, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, |
60728 | 251 |
REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong), |
252 |
REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), |
|
253 |
rtac ctxt trans, rtac ctxt map_cong0, |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
254 |
REPEAT_DETERM_N m o EVERY' [rtac ctxt @{thm Collect_case_prodD}, etac ctxt set_mp, assume_tac ctxt], |
60728 | 255 |
REPEAT_DETERM_N n o rtac ctxt refl, |
60752 | 256 |
assume_tac ctxt, rtac ctxt CollectI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
257 |
CONJ_WRAP' (fn (i, thm) => |
60728 | 258 |
if i <= m then rtac ctxt subset_UNIV |
259 |
else EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt trans, rtac ctxt thm, |
|
60752 | 260 |
rtac ctxt trans_fun_cong_image_id_id_apply, assume_tac ctxt]) |
53289 | 261 |
(1 upto (m + n) ~~ set_map0s)]; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
262 |
in |
60728 | 263 |
EVERY' [rtac ctxt (bis_def RS trans), |
264 |
rtac ctxt iffI, etac ctxt conjE, etac ctxt conjI, CONJ_WRAP' mk_if_tac thms, |
|
265 |
etac ctxt conjE, etac ctxt conjI, CONJ_WRAP' mk_only_if_tac thms] 1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
266 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
267 |
|
60728 | 268 |
fun mk_bis_converse_tac ctxt m bis_rel rel_congs rel_converseps = |
269 |
EVERY' [rtac ctxt (bis_rel RS iffD2), dtac ctxt (bis_rel RS iffD1), |
|
270 |
REPEAT_DETERM o etac ctxt conjE, rtac ctxt conjI, |
|
271 |
CONJ_WRAP' (K (EVERY' [rtac ctxt converse_shift, etac ctxt subset_trans, |
|
272 |
rtac ctxt equalityD2, rtac ctxt @{thm converse_Times}])) rel_congs, |
|
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
273 |
CONJ_WRAP' (fn (rel_cong, rel_conversep) => |
60728 | 274 |
EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, |
275 |
rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), |
|
276 |
REPEAT_DETERM_N m o rtac ctxt @{thm conversep_eq}, |
|
277 |
REPEAT_DETERM_N (length rel_congs) o rtac ctxt @{thm conversep_in_rel}, |
|
278 |
rtac ctxt (rel_conversep RS sym RS @{thm eq_refl} RS @{thm predicate2D}), |
|
279 |
REPEAT_DETERM o etac ctxt allE, |
|
280 |
rtac ctxt @{thm conversepI}, etac ctxt mp, etac ctxt @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
281 |
|
57726 | 282 |
fun mk_bis_O_tac ctxt m bis_rel rel_congs le_rel_OOs = |
60728 | 283 |
EVERY' [rtac ctxt (bis_rel RS iffD2), REPEAT_DETERM o dtac ctxt (bis_rel RS iffD1), |
284 |
REPEAT_DETERM o etac ctxt conjE, rtac ctxt conjI, |
|
60752 | 285 |
CONJ_WRAP' (K (EVERY' [etac ctxt @{thm relcomp_subset_Sigma}, assume_tac ctxt])) rel_congs, |
57726 | 286 |
CONJ_WRAP' (fn (rel_cong, le_rel_OO) => |
60728 | 287 |
EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, |
288 |
rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), |
|
289 |
REPEAT_DETERM_N m o rtac ctxt @{thm eq_OO}, |
|
290 |
REPEAT_DETERM_N (length rel_congs) o rtac ctxt @{thm relcompp_in_rel}, |
|
291 |
rtac ctxt (le_rel_OO RS @{thm predicate2D}), |
|
292 |
etac ctxt @{thm relcompE}, |
|
293 |
REPEAT_DETERM o dtac ctxt prod_injectD, |
|
294 |
etac ctxt conjE, hyp_subst_tac ctxt, |
|
295 |
REPEAT_DETERM o etac ctxt allE, rtac ctxt @{thm relcomppI}, |
|
60752 | 296 |
etac ctxt mp, assume_tac ctxt, etac ctxt mp, assume_tac ctxt]) (rel_congs ~~ le_rel_OOs)] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
297 |
|
55197 | 298 |
fun mk_bis_Gr_tac ctxt bis_rel rel_Grps mor_images morEs coalg_ins = |
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
299 |
unfold_thms_tac ctxt (bis_rel :: @{thm eq_alt} :: @{thm in_rel_Gr} :: rel_Grps) THEN |
60728 | 300 |
EVERY' [rtac ctxt conjI, |
301 |
CONJ_WRAP' (fn thm => rtac ctxt (@{thm Gr_incl} RS iffD2) THEN' etac ctxt thm) mor_images, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
302 |
CONJ_WRAP' (fn (coalg_in, morE) => |
60728 | 303 |
EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt @{thm GrpI}, etac ctxt (morE RS trans), |
304 |
etac ctxt @{thm GrD1}, etac ctxt (@{thm GrD2} RS arg_cong), etac ctxt coalg_in, etac ctxt @{thm GrD1}]) |
|
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
305 |
(coalg_ins ~~ morEs)] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
306 |
|
55197 | 307 |
fun mk_bis_Union_tac ctxt bis_def in_monos = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
308 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
309 |
val n = length in_monos; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
310 |
val ks = 1 upto n; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
311 |
in |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
312 |
unfold_thms_tac ctxt [bis_def] THEN |
60728 | 313 |
EVERY' [rtac ctxt conjI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
314 |
CONJ_WRAP' (fn i => |
60752 | 315 |
EVERY' [rtac ctxt @{thm UN_least}, dtac ctxt bspec, assume_tac ctxt, |
60728 | 316 |
dtac ctxt conjunct1, etac ctxt (mk_conjunctN n i)]) ks, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
317 |
CONJ_WRAP' (fn (i, in_mono) => |
60752 | 318 |
EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, etac ctxt @{thm UN_E}, |
319 |
dtac ctxt bspec, assume_tac ctxt, |
|
60728 | 320 |
dtac ctxt conjunct2, dtac ctxt (mk_conjunctN n i), etac ctxt allE, etac ctxt allE, dtac ctxt mp, |
60752 | 321 |
assume_tac ctxt, etac ctxt bexE, rtac ctxt bexI, assume_tac ctxt, rtac ctxt in_mono, |
60728 | 322 |
REPEAT_DETERM_N n o etac ctxt @{thm SUP_upper2[OF _ subset_refl]}, |
60752 | 323 |
assume_tac ctxt]) (ks ~~ in_monos)] 1 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
324 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
325 |
|
51798 | 326 |
fun mk_sbis_lsbis_tac ctxt lsbis_defs bis_Union bis_cong = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
327 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
328 |
val n = length lsbis_defs; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
329 |
in |
60728 | 330 |
EVERY' [rtac ctxt (Thm.permute_prems 0 1 bis_cong), EVERY' (map (rtac ctxt) lsbis_defs), |
331 |
rtac ctxt bis_Union, rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, exE], |
|
332 |
hyp_subst_tac ctxt, etac ctxt bis_cong, EVERY' (map (rtac ctxt o mk_nth_conv n) (1 upto n))] 1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
333 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
334 |
|
60728 | 335 |
fun mk_incl_lsbis_tac ctxt n i lsbis_def = |
336 |
EVERY' [rtac ctxt @{thm xt1(3)}, rtac ctxt lsbis_def, rtac ctxt @{thm SUP_upper2}, rtac ctxt CollectI, |
|
60752 | 337 |
REPEAT_DETERM_N n o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, assume_tac ctxt, |
338 |
rtac ctxt equalityD2, rtac ctxt (mk_nth_conv n i)] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
339 |
|
60728 | 340 |
fun mk_equiv_lsbis_tac ctxt sbis_lsbis lsbis_incl incl_lsbis bis_Id_on bis_converse bis_O = |
341 |
EVERY' [rtac ctxt (@{thm equiv_def} RS iffD2), |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
342 |
|
60728 | 343 |
rtac ctxt conjI, rtac ctxt (@{thm refl_on_def} RS iffD2), |
344 |
rtac ctxt conjI, rtac ctxt lsbis_incl, rtac ctxt ballI, rtac ctxt set_mp, |
|
60752 | 345 |
rtac ctxt incl_lsbis, rtac ctxt bis_Id_on, assume_tac ctxt, etac ctxt @{thm Id_onI}, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
346 |
|
60728 | 347 |
rtac ctxt conjI, rtac ctxt (@{thm sym_def} RS iffD2), |
348 |
rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt set_mp, |
|
349 |
rtac ctxt incl_lsbis, rtac ctxt bis_converse, rtac ctxt sbis_lsbis, etac ctxt @{thm converseI}, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
350 |
|
60728 | 351 |
rtac ctxt (@{thm trans_def} RS iffD2), |
352 |
rtac ctxt allI, rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt impI, rtac ctxt set_mp, |
|
353 |
rtac ctxt incl_lsbis, rtac ctxt bis_O, rtac ctxt sbis_lsbis, rtac ctxt sbis_lsbis, |
|
60752 | 354 |
etac ctxt @{thm relcompI}, assume_tac ctxt] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
355 |
|
55197 | 356 |
fun mk_coalgT_tac ctxt m defs strT_defs set_map0ss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
357 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
358 |
val n = length strT_defs; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
359 |
val ks = 1 upto n; |
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
360 |
fun coalg_tac (i, (active_sets, def)) = |
60728 | 361 |
EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], |
362 |
hyp_subst_tac ctxt, rtac ctxt (def RS trans RS @{thm ssubst_mem}), etac ctxt (arg_cong RS trans), |
|
363 |
rtac ctxt (mk_sum_caseN n i), rtac ctxt CollectI, |
|
364 |
REPEAT_DETERM_N m o EVERY' [rtac ctxt conjI, rtac ctxt subset_UNIV], |
|
365 |
CONJ_WRAP' (fn (i, thm) => EVERY' [rtac ctxt (thm RS ord_eq_le_trans), |
|
366 |
rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt exI, rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, |
|
367 |
rtac ctxt conjI, |
|
368 |
rtac ctxt conjI, etac ctxt @{thm empty_Shift}, dtac ctxt set_rev_mp, |
|
369 |
etac ctxt equalityD1, etac ctxt CollectD, |
|
370 |
rtac ctxt ballI, |
|
371 |
CONJ_WRAP' (fn i => EVERY' [rtac ctxt ballI, etac ctxt CollectE, dtac ctxt @{thm ShiftD}, |
|
60752 | 372 |
dtac ctxt bspec, etac ctxt thin_rl, assume_tac ctxt, dtac ctxt (mk_conjunctN n i), |
60728 | 373 |
dtac ctxt bspec, rtac ctxt CollectI, etac ctxt @{thm set_mp[OF equalityD1[OF Succ_Shift]]}, |
374 |
REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac ctxt exI, |
|
375 |
rtac ctxt conjI, rtac ctxt (@{thm shift_def} RS fun_cong RS trans), |
|
60752 | 376 |
rtac ctxt (@{thm append_Cons} RS sym RS arg_cong RS trans), assume_tac ctxt, |
60728 | 377 |
CONJ_WRAP' (K (EVERY' [etac ctxt trans, rtac ctxt @{thm Collect_cong}, |
378 |
rtac ctxt @{thm eqset_imp_iff}, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm Succ_Shift}, |
|
379 |
rtac ctxt (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks, |
|
60752 | 380 |
dtac ctxt bspec, assume_tac ctxt, dtac ctxt (mk_conjunctN n i), dtac ctxt bspec, |
381 |
etac ctxt @{thm set_mp[OF equalityD1]}, assume_tac ctxt, |
|
60728 | 382 |
REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac ctxt exI, |
383 |
rtac ctxt conjI, rtac ctxt (@{thm shift_def} RS fun_cong RS trans), |
|
384 |
etac ctxt (@{thm append_Nil} RS sym RS arg_cong RS trans), |
|
60752 | 385 |
REPEAT_DETERM_N m o (rtac ctxt conjI THEN' assume_tac ctxt), |
60728 | 386 |
CONJ_WRAP' (K (EVERY' [etac ctxt trans, rtac ctxt @{thm Collect_cong}, |
387 |
rtac ctxt @{thm eqset_imp_iff}, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm Succ_Shift}, |
|
388 |
rtac ctxt (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)]; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
389 |
in |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
390 |
unfold_thms_tac ctxt defs THEN |
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
391 |
CONJ_WRAP' coalg_tac (ks ~~ (map (drop m) set_map0ss ~~ strT_defs)) 1 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
392 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
393 |
|
51798 | 394 |
fun mk_length_Lev_tac ctxt cts Lev_0s Lev_Sucs = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
395 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
396 |
val n = length Lev_0s; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
397 |
val ks = n downto 1; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
398 |
in |
60801 | 399 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 400 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
401 |
CONJ_WRAP' (fn Lev_0 => |
60728 | 402 |
EVERY' [rtac ctxt impI, dtac ctxt (Lev_0 RS equalityD1 RS set_mp), |
403 |
etac ctxt @{thm singletonE}, etac ctxt ssubst, rtac ctxt @{thm list.size(3)}]) Lev_0s, |
|
404 |
REPEAT_DETERM o rtac ctxt allI, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
405 |
CONJ_WRAP' (fn Lev_Suc => |
60728 | 406 |
EVERY' [rtac ctxt impI, dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), |
407 |
CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) |
|
51798 | 408 |
(fn i => |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
409 |
EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt, |
60728 | 410 |
rtac ctxt trans, rtac ctxt @{thm length_Cons}, rtac ctxt @{thm arg_cong[of _ _ Suc]}, |
60752 | 411 |
REPEAT_DETERM o etac ctxt allE, dtac ctxt (mk_conjunctN n i), |
412 |
etac ctxt mp, assume_tac ctxt]) ks]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
413 |
Lev_Sucs] 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
414 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
415 |
|
60728 | 416 |
fun mk_length_Lev'_tac ctxt length_Lev = |
60777 | 417 |
EVERY' [forward_tac ctxt [length_Lev], etac ctxt ssubst, assume_tac ctxt] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
418 |
|
60728 | 419 |
fun mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
420 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
421 |
val n = length rv_Nils; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
422 |
val ks = 1 upto n; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
423 |
in |
60801 | 424 |
EVERY' [rtac ctxt (Thm.instantiate' cTs cts @{thm list.induct}), |
60728 | 425 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
426 |
CONJ_WRAP' (fn rv_Cons => |
60728 | 427 |
CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac ctxt exI, |
428 |
rtac ctxt (@{thm append_Nil} RS arg_cong RS trans), |
|
429 |
rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS trans), rtac ctxt rv_Nil])) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
430 |
(ks ~~ rv_Nils)) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
431 |
rv_Conss, |
60728 | 432 |
REPEAT_DETERM o rtac ctxt allI, rtac ctxt (mk_sumEN n), |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
433 |
EVERY' (map (fn i => |
60728 | 434 |
CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac ctxt allE, dtac ctxt (mk_conjunctN n i), |
435 |
CONJ_WRAP' (fn i' => EVERY' [dtac ctxt (mk_conjunctN n i'), etac ctxt exE, rtac ctxt exI, |
|
436 |
rtac ctxt (@{thm append_Cons} RS arg_cong RS trans), rtac ctxt (rv_Cons RS trans), |
|
437 |
if n = 1 then K all_tac else etac ctxt (sum_case_cong_weak RS trans), |
|
60752 | 438 |
rtac ctxt (mk_sum_caseN n i RS trans), assume_tac ctxt]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
439 |
ks]) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
440 |
rv_Conss) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
441 |
ks)] 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
442 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
443 |
|
51798 | 444 |
fun mk_set_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
445 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
446 |
val n = length Lev_0s; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
447 |
val ks = 1 upto n; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
448 |
in |
60801 | 449 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 450 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
451 |
CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) => |
60728 | 452 |
EVERY' [rtac ctxt impI, dtac ctxt (Lev_0 RS equalityD1 RS set_mp), |
453 |
etac ctxt @{thm singletonE}, hyp_subst_tac ctxt, |
|
454 |
CONJ_WRAP' (fn i' => rtac ctxt impI THEN' dtac ctxt (sym RS trans) THEN' rtac ctxt rv_Nil THEN' |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
455 |
(if i = i' |
60728 | 456 |
then EVERY' [dtac ctxt (mk_InN_inject n i), hyp_subst_tac ctxt, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
457 |
CONJ_WRAP' (fn (i'', Lev_0'') => |
60728 | 458 |
EVERY' [rtac ctxt impI, rtac ctxt @{thm ssubst_mem[OF append_Nil]}, |
459 |
rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), rtac ctxt (mk_UnIN n i''), |
|
460 |
rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, |
|
461 |
etac ctxt conjI, rtac ctxt (Lev_0'' RS equalityD2 RS set_mp), |
|
462 |
rtac ctxt @{thm singletonI}]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
463 |
(ks ~~ Lev_0s)] |
60728 | 464 |
else etac ctxt (mk_InN_not_InM i' i RS notE))) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
465 |
ks]) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
466 |
((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils), |
60728 | 467 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
468 |
CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), from_to_sbds) => |
60728 | 469 |
EVERY' [rtac ctxt impI, dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), |
470 |
CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
471 |
(fn (i, from_to_sbd) => |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
472 |
EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt, |
60728 | 473 |
CONJ_WRAP' (fn i' => rtac ctxt impI THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
474 |
CONJ_WRAP' (fn i'' => |
60728 | 475 |
EVERY' [rtac ctxt impI, rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), |
476 |
rtac ctxt @{thm ssubst_mem[OF append_Cons]}, rtac ctxt (mk_UnIN n i), |
|
477 |
rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, |
|
60752 | 478 |
rtac ctxt conjI, assume_tac ctxt, dtac ctxt (sym RS trans RS sym), |
60728 | 479 |
rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS trans), |
480 |
etac ctxt (from_to_sbd RS arg_cong), REPEAT_DETERM o etac ctxt allE, |
|
60752 | 481 |
dtac ctxt (mk_conjunctN n i), dtac ctxt mp, assume_tac ctxt, |
482 |
dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, assume_tac ctxt, |
|
483 |
dtac ctxt (mk_conjunctN n i''), etac ctxt mp, assume_tac ctxt]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
484 |
ks) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
485 |
ks]) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
486 |
(rev (ks ~~ from_to_sbds))]) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
487 |
((Lev_Sucs ~~ rv_Conss) ~~ from_to_sbdss)] 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
488 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
489 |
|
51798 | 490 |
fun mk_set_image_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss to_sbd_injss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
491 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
492 |
val n = length Lev_0s; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
493 |
val ks = 1 upto n; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
494 |
in |
60801 | 495 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 496 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
497 |
CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) => |
60728 | 498 |
EVERY' [rtac ctxt impI, dtac ctxt (Lev_0 RS equalityD1 RS set_mp), |
499 |
etac ctxt @{thm singletonE}, hyp_subst_tac ctxt, |
|
500 |
CONJ_WRAP' (fn i' => rtac ctxt impI THEN' |
|
501 |
CONJ_WRAP' (fn i'' => rtac ctxt impI THEN' dtac ctxt (sym RS trans) THEN' rtac ctxt rv_Nil THEN' |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
502 |
(if i = i'' |
60728 | 503 |
then EVERY' [dtac ctxt @{thm ssubst_mem[OF sym[OF append_Nil]]}, |
504 |
dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), dtac ctxt (mk_InN_inject n i), |
|
51798 | 505 |
hyp_subst_tac ctxt, |
60728 | 506 |
CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
507 |
(fn k => REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' |
60728 | 508 |
dtac ctxt list_inject_iffD1 THEN' etac ctxt conjE THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
509 |
(if k = i' |
60728 | 510 |
then EVERY' [dtac ctxt (mk_InN_inject n k), hyp_subst_tac ctxt, etac ctxt imageI] |
511 |
else etac ctxt (mk_InN_not_InM i' k RS notE))) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
512 |
(rev ks)] |
60728 | 513 |
else etac ctxt (mk_InN_not_InM i'' i RS notE))) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
514 |
ks) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
515 |
ks]) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
516 |
((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils), |
60728 | 517 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
518 |
CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, to_sbd_injs)) => |
60728 | 519 |
EVERY' [rtac ctxt impI, dtac ctxt (Lev_Suc RS equalityD1 RS set_mp), |
520 |
CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
521 |
(fn (i, (from_to_sbd, to_sbd_inj)) => |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
522 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' hyp_subst_tac ctxt THEN' |
60728 | 523 |
CONJ_WRAP' (fn i' => rtac ctxt impI THEN' |
524 |
dtac ctxt @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN' |
|
525 |
dtac ctxt (Lev_Suc RS equalityD1 RS set_mp) THEN' |
|
526 |
CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn k => |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
527 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' |
60728 | 528 |
dtac ctxt list_inject_iffD1 THEN' etac ctxt conjE THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
529 |
(if k = i |
60728 | 530 |
then EVERY' [dtac ctxt (mk_InN_inject n i), |
531 |
dtac ctxt (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)), |
|
60752 | 532 |
assume_tac ctxt, assume_tac ctxt, hyp_subst_tac ctxt] THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
533 |
CONJ_WRAP' (fn i'' => |
60728 | 534 |
EVERY' [rtac ctxt impI, dtac ctxt (sym RS trans), |
535 |
rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS arg_cong RS trans), |
|
536 |
etac ctxt (from_to_sbd RS arg_cong), |
|
537 |
REPEAT_DETERM o etac ctxt allE, |
|
60752 | 538 |
dtac ctxt (mk_conjunctN n i), dtac ctxt mp, assume_tac ctxt, |
539 |
dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, assume_tac ctxt, |
|
60728 | 540 |
dtac ctxt (mk_conjunctN n i''), etac ctxt mp, etac ctxt sym]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
541 |
ks |
60728 | 542 |
else etac ctxt (mk_InN_not_InM i k RS notE))) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
543 |
(rev ks)) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
544 |
ks) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
545 |
(rev (ks ~~ (from_to_sbds ~~ to_sbd_injs)))]) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
546 |
((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ to_sbd_injss))] 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
547 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
548 |
|
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
549 |
fun mk_mor_beh_tac ctxt m mor_def mor_cong beh_defs carT_defs strT_defs isNode_defs to_sbd_injss |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
550 |
from_to_sbdss Lev_0s Lev_Sucs rv_Nils rv_Conss length_Levs length_Lev's rv_lastss set_Levsss |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
551 |
set_image_Levsss set_map0ss map_comp_ids map_cong0s = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
552 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
553 |
val n = length beh_defs; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
554 |
val ks = 1 upto n; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
555 |
|
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
556 |
fun fbetw_tac (i, (carT_def, (isNode_def, (Lev_0, (rv_Nil, ((length_Lev, length_Lev'), |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
557 |
(rv_lasts, (set_map0s, (set_Levss, set_image_Levss))))))))) = |
60728 | 558 |
EVERY' [rtac ctxt ballI, rtac ctxt (carT_def RS equalityD2 RS set_mp), |
559 |
rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, rtac ctxt conjI, |
|
560 |
rtac ctxt conjI, |
|
561 |
rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, rtac ctxt (Lev_0 RS equalityD2 RS set_mp), |
|
562 |
rtac ctxt @{thm singletonI}, |
|
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
563 |
(**) |
60728 | 564 |
rtac ctxt ballI, etac ctxt @{thm UN_E}, |
53289 | 565 |
CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_map0s, |
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
566 |
(set_Levs, set_image_Levs))))) => |
60728 | 567 |
EVERY' [rtac ctxt ballI, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
568 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}], |
60728 | 569 |
rtac ctxt (rev_mp OF [rv_last, impI]), etac ctxt exE, rtac ctxt (isNode_def RS iffD2), |
570 |
rtac ctxt exI, rtac ctxt conjI, |
|
571 |
if n = 1 then rtac ctxt refl |
|
572 |
else etac ctxt (sum_case_cong_weak RS trans) THEN' rtac ctxt (mk_sum_caseN n i), |
|
53289 | 573 |
CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) => |
60728 | 574 |
EVERY' [rtac ctxt (set_map0 RS trans), rtac ctxt equalityI, rtac ctxt @{thm image_subsetI}, |
575 |
rtac ctxt CollectI, rtac ctxt @{thm SuccI}, rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, etac ctxt set_Lev, |
|
60752 | 576 |
if n = 1 then rtac ctxt refl else assume_tac ctxt, assume_tac ctxt, rtac ctxt subsetI, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
577 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}], |
60728 | 578 |
REPEAT_DETERM_N 4 o etac ctxt thin_rl, |
579 |
rtac ctxt set_image_Lev, |
|
60752 | 580 |
assume_tac ctxt, dtac ctxt length_Lev, hyp_subst_tac ctxt, dtac ctxt length_Lev', |
60728 | 581 |
etac ctxt @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]}, |
60752 | 582 |
if n = 1 then rtac ctxt refl else assume_tac ctxt]) |
53289 | 583 |
(drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))]) |
584 |
(ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_map0ss ~~ |
|
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
585 |
(set_Levss ~~ set_image_Levss))))), |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
586 |
(*root isNode*) |
60728 | 587 |
rtac ctxt (isNode_def RS iffD2), rtac ctxt exI, rtac ctxt conjI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
588 |
CONVERSION (Conv.top_conv |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
589 |
(K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt), |
60728 | 590 |
if n = 1 then rtac ctxt refl else rtac ctxt (mk_sum_caseN n i), |
53289 | 591 |
CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) => |
60728 | 592 |
EVERY' [rtac ctxt (set_map0 RS trans), rtac ctxt equalityI, rtac ctxt @{thm image_subsetI}, |
593 |
rtac ctxt CollectI, rtac ctxt @{thm SuccI}, rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, rtac ctxt set_Lev, |
|
594 |
rtac ctxt (Lev_0 RS equalityD2 RS set_mp), rtac ctxt @{thm singletonI}, rtac ctxt rv_Nil, |
|
60752 | 595 |
assume_tac ctxt, rtac ctxt subsetI, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
596 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}], |
60728 | 597 |
rtac ctxt set_image_Lev, rtac ctxt (Lev_0 RS equalityD2 RS set_mp), |
598 |
rtac ctxt @{thm singletonI}, dtac ctxt length_Lev', |
|
599 |
etac ctxt @{thm set_mp[OF equalityD1[OF arg_cong[OF |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
600 |
trans[OF length_append_singleton arg_cong[of _ _ Suc, OF list.size(3)]]]]]}, |
60728 | 601 |
rtac ctxt rv_Nil]) |
53289 | 602 |
(drop m set_map0s ~~ (nth set_Levss (i - 1) ~~ nth set_image_Levss (i - 1)))]; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
603 |
|
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
604 |
fun mor_tac (i, (strT_def, ((Lev_Suc, (rv_Nil, rv_Cons)), |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
605 |
((map_comp_id, map_cong0), (length_Lev', (from_to_sbds, to_sbd_injs)))))) = |
60728 | 606 |
EVERY' [rtac ctxt ballI, rtac ctxt sym, rtac ctxt trans, rtac ctxt strT_def, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
607 |
CONVERSION (Conv.top_conv |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
608 |
(K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt), |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
609 |
if n = 1 then K all_tac |
60728 | 610 |
else (rtac ctxt (sum_case_cong_weak RS trans) THEN' |
611 |
rtac ctxt (mk_sum_caseN n i) THEN' rtac ctxt (mk_sum_caseN n i RS trans)), |
|
612 |
rtac ctxt (map_comp_id RS trans), rtac ctxt (map_cong0 OF replicate m refl), |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58446
diff
changeset
|
613 |
EVERY' (@{map 3} (fn i' => fn to_sbd_inj => fn from_to_sbd => |
60728 | 614 |
DETERM o EVERY' [rtac ctxt trans, rtac ctxt o_apply, rtac ctxt prod_injectI, rtac ctxt conjI, |
615 |
rtac ctxt trans, rtac ctxt @{thm Shift_def}, |
|
616 |
rtac ctxt equalityI, rtac ctxt subsetI, etac ctxt thin_rl, |
|
617 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm UN_E}], dtac ctxt length_Lev', dtac ctxt asm_rl, |
|
618 |
etac ctxt thin_rl, dtac ctxt @{thm set_rev_mp[OF _ equalityD1]}, |
|
619 |
rtac ctxt (@{thm length_Cons} RS arg_cong RS trans), rtac ctxt Lev_Suc, |
|
620 |
CONJ_WRAP_GEN' (etac ctxt (Thm.permute_prems 1 1 UnE)) (fn i'' => |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
621 |
EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], |
60728 | 622 |
dtac ctxt list_inject_iffD1, etac ctxt conjE, |
623 |
if i' = i'' then EVERY' [dtac ctxt (mk_InN_inject n i'), |
|
624 |
dtac ctxt (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)), |
|
60752 | 625 |
assume_tac ctxt, assume_tac ctxt, hyp_subst_tac ctxt, etac ctxt @{thm UN_I[OF UNIV_I]}] |
60728 | 626 |
else etac ctxt (mk_InN_not_InM i' i'' RS notE)]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
627 |
(rev ks), |
60728 | 628 |
rtac ctxt @{thm UN_least}, rtac ctxt subsetI, rtac ctxt CollectI, rtac ctxt @{thm UN_I[OF UNIV_I]}, |
629 |
rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), rtac ctxt (mk_UnIN n i'), rtac ctxt CollectI, |
|
60752 | 630 |
REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, etac ctxt conjI, assume_tac ctxt, |
60728 | 631 |
rtac ctxt trans, rtac ctxt @{thm shift_def}, rtac ctxt iffD2, rtac ctxt @{thm fun_eq_iff}, rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
632 |
CONVERSION (Conv.top_conv |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
633 |
(K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt), |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
634 |
if n = 1 then K all_tac |
60728 | 635 |
else rtac ctxt sum_case_cong_weak THEN' rtac ctxt (mk_sum_caseN n i' RS trans), |
636 |
SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac ctxt refl]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
637 |
ks to_sbd_injs from_to_sbds)]; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
638 |
in |
60728 | 639 |
(rtac ctxt mor_cong THEN' |
640 |
EVERY' (map (fn thm => rtac ctxt (thm RS @{thm ext})) beh_defs) THEN' |
|
641 |
rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN' |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
642 |
CONJ_WRAP' fbetw_tac |
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
643 |
(ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ ((length_Levs ~~ length_Lev's) ~~ |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
644 |
(rv_lastss ~~ (set_map0ss ~~ (set_Levsss ~~ set_image_Levsss))))))))) THEN' |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
645 |
CONJ_WRAP' mor_tac |
55577
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
646 |
(ks ~~ (strT_defs ~~ ((Lev_Sucs ~~ (rv_Nils ~~ rv_Conss)) ~~ |
a6c2379078c8
simplifications of internal codatatype construction
traytel
parents:
55541
diff
changeset
|
647 |
((map_comp_ids ~~ map_cong0s) ~~ |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
648 |
(length_Lev's ~~ (from_to_sbdss ~~ to_sbd_injss))))))) 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
649 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
650 |
|
60728 | 651 |
fun mk_congruent_str_final_tac ctxt m lsbisE map_comp_id map_cong0 equiv_LSBISs = |
652 |
EVERY' [rtac ctxt @{thm congruentI}, dtac ctxt lsbisE, |
|
60757 | 653 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, bexE], rtac ctxt (o_apply RS trans), |
60728 | 654 |
etac ctxt (sym RS arg_cong RS trans), rtac ctxt (map_comp_id RS trans), |
655 |
rtac ctxt (map_cong0 RS trans), REPEAT_DETERM_N m o rtac ctxt refl, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
656 |
EVERY' (map (fn equiv_LSBIS => |
60752 | 657 |
EVERY' [rtac ctxt @{thm equiv_proj}, rtac ctxt equiv_LSBIS, etac ctxt set_mp, assume_tac ctxt]) |
60728 | 658 |
equiv_LSBISs), rtac ctxt sym, rtac ctxt (o_apply RS trans), |
659 |
etac ctxt (sym RS arg_cong RS trans), rtac ctxt map_comp_id] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
660 |
|
60728 | 661 |
fun mk_coalg_final_tac ctxt m coalg_def congruent_str_finals equiv_LSBISs set_map0ss coalgT_setss = |
662 |
EVERY' [rtac ctxt (coalg_def RS iffD2), |
|
53289 | 663 |
CONJ_WRAP' (fn ((set_map0s, coalgT_sets), (equiv_LSBIS, congruent_str_final)) => |
60728 | 664 |
EVERY' [rtac ctxt @{thm univ_preserves}, rtac ctxt equiv_LSBIS, rtac ctxt congruent_str_final, |
665 |
rtac ctxt ballI, rtac ctxt @{thm ssubst_mem}, rtac ctxt o_apply, rtac ctxt CollectI, |
|
666 |
REPEAT_DETERM_N m o EVERY' [rtac ctxt conjI, rtac ctxt subset_UNIV], |
|
53289 | 667 |
CONJ_WRAP' (fn (equiv_LSBIS, (set_map0, coalgT_set)) => |
60728 | 668 |
EVERY' [rtac ctxt (set_map0 RS ord_eq_le_trans), |
669 |
rtac ctxt @{thm image_subsetI}, rtac ctxt iffD2, rtac ctxt @{thm proj_in_iff}, |
|
670 |
rtac ctxt equiv_LSBIS, etac ctxt set_rev_mp, etac ctxt coalgT_set]) |
|
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
671 |
(equiv_LSBISs ~~ (drop m set_map0s ~~ coalgT_sets))]) |
53289 | 672 |
((set_map0ss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
673 |
|
60728 | 674 |
fun mk_mor_T_final_tac ctxt mor_def congruent_str_finals equiv_LSBISs = |
675 |
EVERY' [rtac ctxt (mor_def RS iffD2), rtac ctxt conjI, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
676 |
CONJ_WRAP' (fn equiv_LSBIS => |
60752 | 677 |
EVERY' [rtac ctxt ballI, rtac ctxt iffD2, rtac ctxt @{thm proj_in_iff}, |
678 |
rtac ctxt equiv_LSBIS, assume_tac ctxt]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
679 |
equiv_LSBISs, |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
680 |
CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) => |
60728 | 681 |
EVERY' [rtac ctxt ballI, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm univ_commute}, rtac ctxt equiv_LSBIS, |
60752 | 682 |
rtac ctxt congruent_str_final, assume_tac ctxt, rtac ctxt o_apply]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
683 |
(equiv_LSBISs ~~ congruent_str_finals)] 1; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
684 |
|
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
685 |
fun mk_mor_Rep_tac ctxt defs Reps Abs_inverses coalg_final_setss map_comp_ids map_cong0Ls = |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
686 |
unfold_thms_tac ctxt defs THEN |
60728 | 687 |
EVERY' [rtac ctxt conjI, |
688 |
CONJ_WRAP' (fn thm => rtac ctxt ballI THEN' rtac ctxt thm) Reps, |
|
51761
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents:
51739
diff
changeset
|
689 |
CONJ_WRAP' (fn (Rep, ((map_comp_id, map_cong0L), coalg_final_sets)) => |
60728 | 690 |
EVERY' [rtac ctxt ballI, rtac ctxt (map_comp_id RS trans), rtac ctxt map_cong0L, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
691 |
EVERY' (map2 (fn Abs_inverse => fn coalg_final_set => |
60728 | 692 |
EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS trans), rtac ctxt Abs_inverse, |
693 |
etac ctxt set_rev_mp, rtac ctxt coalg_final_set, rtac ctxt Rep]) |
|
55541
fd9ea8ae28f6
syntactic simplifications of internal (co)datatype constructions
traytel
parents:
55414
diff
changeset
|
694 |
Abs_inverses coalg_final_sets)]) |
51761
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents:
51739
diff
changeset
|
695 |
(Reps ~~ ((map_comp_ids ~~ map_cong0Ls) ~~ coalg_final_setss))] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
696 |
|
55197 | 697 |
fun mk_mor_Abs_tac ctxt defs Abs_inverses = |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
698 |
unfold_thms_tac ctxt defs THEN |
60728 | 699 |
EVERY' [rtac ctxt conjI, |
700 |
CONJ_WRAP' (K (rtac ctxt ballI THEN' rtac ctxt UNIV_I)) Abs_inverses, |
|
701 |
CONJ_WRAP' (fn thm => rtac ctxt ballI THEN' etac ctxt (thm RS arg_cong RS sym)) Abs_inverses] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
702 |
|
60728 | 703 |
fun mk_mor_unfold_tac ctxt m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_cong0s = |
704 |
EVERY' [rtac ctxt iffD2, rtac ctxt mor_UNIV, |
|
51761
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents:
51739
diff
changeset
|
705 |
CONJ_WRAP' (fn ((Abs_inverse, morE), ((dtor_def, unfold_def), (map_comp_id, map_cong0))) => |
60728 | 706 |
EVERY' [rtac ctxt @{thm ext}, rtac ctxt (o_apply RS trans RS sym), rtac ctxt (dtor_def RS trans), |
707 |
rtac ctxt (unfold_def RS arg_cong RS trans), rtac ctxt (Abs_inverse RS arg_cong RS trans), |
|
708 |
rtac ctxt (morE RS arg_cong RS trans), rtac ctxt (map_comp_id RS trans), |
|
709 |
rtac ctxt (o_apply RS trans RS sym), rtac ctxt map_cong0, |
|
710 |
REPEAT_DETERM_N m o rtac ctxt refl, |
|
711 |
EVERY' (map (fn thm => rtac ctxt (thm RS trans) THEN' rtac ctxt (o_apply RS sym)) unfold_defs)]) |
|
51761
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
blanchet
parents:
51739
diff
changeset
|
712 |
((Abs_inverses ~~ morEs) ~~ ((dtor_defs ~~ unfold_defs) ~~ (map_comp_ids ~~ map_cong0s)))] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
713 |
|
60728 | 714 |
fun mk_raw_coind_tac ctxt bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
715 |
sbis_lsbis lsbis_incls incl_lsbiss equiv_LSBISs mor_Rep Rep_injects = |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
716 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
717 |
val n = length Rep_injects; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
718 |
in |
60777 | 719 |
EVERY' [rtac ctxt rev_mp, forward_tac ctxt [bis_def RS iffD1], |
60728 | 720 |
REPEAT_DETERM o etac ctxt conjE, rtac ctxt bis_cong, rtac ctxt bis_O, rtac ctxt bis_converse, |
60752 | 721 |
rtac ctxt bis_Gr, rtac ctxt tcoalg, rtac ctxt mor_Rep, rtac ctxt bis_O, assume_tac ctxt, |
722 |
rtac ctxt bis_Gr, rtac ctxt tcoalg, |
|
60728 | 723 |
rtac ctxt mor_Rep, REPEAT_DETERM_N n o etac ctxt @{thm relImage_Gr}, |
724 |
rtac ctxt impI, rtac ctxt rev_mp, rtac ctxt bis_cong, rtac ctxt bis_O, rtac ctxt bis_Gr, rtac ctxt coalgT, |
|
725 |
rtac ctxt mor_T_final, rtac ctxt bis_O, rtac ctxt sbis_lsbis, rtac ctxt bis_converse, rtac ctxt bis_Gr, rtac ctxt coalgT, |
|
726 |
rtac ctxt mor_T_final, EVERY' (map (fn thm => rtac ctxt (thm RS @{thm relInvImage_Gr})) lsbis_incls), |
|
727 |
rtac ctxt impI, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
728 |
CONJ_WRAP' (fn (Rep_inject, (equiv_LSBIS , (incl_lsbis, lsbis_incl))) => |
60728 | 729 |
EVERY' [rtac ctxt subset_trans, rtac ctxt @{thm relInvImage_UNIV_relImage}, rtac ctxt subset_trans, |
730 |
rtac ctxt @{thm relInvImage_mono}, rtac ctxt subset_trans, etac ctxt incl_lsbis, |
|
731 |
rtac ctxt ord_eq_le_trans, rtac ctxt @{thm sym[OF relImage_relInvImage]}, |
|
732 |
rtac ctxt @{thm xt1(3)}, rtac ctxt @{thm Sigma_cong}, |
|
733 |
rtac ctxt @{thm proj_image}, rtac ctxt @{thm proj_image}, rtac ctxt lsbis_incl, |
|
60752 | 734 |
rtac ctxt subset_trans, rtac ctxt @{thm relImage_mono}, rtac ctxt incl_lsbis, assume_tac ctxt, |
60728 | 735 |
rtac ctxt @{thm relImage_proj}, rtac ctxt equiv_LSBIS, rtac ctxt @{thm relInvImage_Id_on}, |
736 |
rtac ctxt Rep_inject]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
737 |
(Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
738 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
739 |
|
60728 | 740 |
fun mk_unfold_unique_mor_tac ctxt raw_coinds bis mor unfold_defs = |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
741 |
CONJ_WRAP' (fn (raw_coind, unfold_def) => |
60728 | 742 |
EVERY' [rtac ctxt @{thm ext}, etac ctxt (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac ctxt mor, |
743 |
rtac ctxt @{thm image2_eqI}, rtac ctxt refl, rtac ctxt (unfold_def RS arg_cong RS trans), |
|
744 |
rtac ctxt (o_apply RS sym), rtac ctxt UNIV_I]) (raw_coinds ~~ unfold_defs) 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
745 |
|
55197 | 746 |
fun mk_dtor_o_ctor_tac ctxt ctor_def unfold map_comp_id map_cong0L unfold_o_dtors = |
60728 | 747 |
unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply, |
748 |
rtac ctxt trans, rtac ctxt unfold, rtac ctxt trans, rtac ctxt map_comp_id, rtac ctxt trans, rtac ctxt map_cong0L, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
749 |
EVERY' (map (fn thm => |
60728 | 750 |
rtac ctxt ballI THEN' rtac ctxt (trans OF [thm RS fun_cong, id_apply])) unfold_o_dtors), |
751 |
rtac ctxt sym, rtac ctxt id_apply] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
752 |
|
55197 | 753 |
fun mk_corec_tac ctxt m corec_defs unfold map_cong0 corec_Inls = |
60728 | 754 |
unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac ctxt trans, rtac ctxt (o_apply RS arg_cong), |
755 |
rtac ctxt trans, rtac ctxt unfold, fo_rtac ctxt (@{thm sum.case(2)} RS arg_cong RS trans), |
|
756 |
rtac ctxt map_cong0, REPEAT_DETERM_N m o rtac ctxt refl, |
|
757 |
EVERY' (map (fn thm => rtac ctxt @{thm case_sum_expand_Inr} THEN' rtac ctxt thm) corec_Inls)] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
758 |
|
55197 | 759 |
fun mk_corec_unique_mor_tac ctxt corec_defs corec_Inls unfold_unique_mor = |
51739
3514b90d0a8b
(co)rec is (just as the (un)fold) the unique morphism;
traytel
parents:
51447
diff
changeset
|
760 |
unfold_thms_tac ctxt |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55393
diff
changeset
|
761 |
(corec_defs @ map (fn thm => thm RS @{thm case_sum_expand_Inr'}) corec_Inls) THEN |
60728 | 762 |
etac ctxt unfold_unique_mor 1; |
51739
3514b90d0a8b
(co)rec is (just as the (un)fold) the unique morphism;
traytel
parents:
51447
diff
changeset
|
763 |
|
60728 | 764 |
fun mk_dtor_coinduct_tac ctxt m raw_coind bis_rel rel_congs = |
765 |
EVERY' [rtac ctxt rev_mp, rtac ctxt raw_coind, rtac ctxt iffD2, rtac ctxt bis_rel, rtac ctxt conjI, |
|
766 |
CONJ_WRAP' (K (rtac ctxt @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) |
|
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
767 |
rel_congs, |
60728 | 768 |
CONJ_WRAP' (fn rel_cong => EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, |
769 |
REPEAT_DETERM o etac ctxt allE, rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}), |
|
770 |
REPEAT_DETERM_N m o rtac ctxt refl, |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
771 |
REPEAT_DETERM_N (length rel_congs) o rtac ctxt @{thm in_rel_Collect_case_prod_eq[symmetric]}, |
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
772 |
etac ctxt mp, etac ctxt CollectE, etac ctxt @{thm case_prodD}]) |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
773 |
rel_congs, |
60728 | 774 |
rtac ctxt impI, REPEAT_DETERM o etac ctxt conjE, |
775 |
CONJ_WRAP' (K (EVERY' [rtac ctxt impI, rtac ctxt @{thm IdD}, etac ctxt set_mp, |
|
776 |
rtac ctxt CollectI, etac ctxt @{thm case_prodI}])) rel_congs] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
777 |
|
60728 | 778 |
fun mk_map_tac ctxt m n map_arg_cong unfold map_comp map_cong0 = |
779 |
EVERY' [rtac ctxt @{thm ext}, rtac ctxt (o_apply RS trans RS sym), rtac ctxt (o_apply RS trans RS sym), |
|
780 |
rtac ctxt (unfold RS trans), rtac ctxt (Thm.permute_prems 0 1 (map_comp RS @{thm box_equals})), |
|
781 |
rtac ctxt map_cong0, |
|
782 |
REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong), |
|
783 |
REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong), |
|
784 |
rtac ctxt map_arg_cong, rtac ctxt (o_apply RS sym)] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
785 |
|
60728 | 786 |
fun mk_set_le_tac ctxt n Jset_minimal set_Jsets set_Jset_Jsetss = |
787 |
EVERY' [rtac ctxt Jset_minimal, |
|
788 |
REPEAT_DETERM_N n o rtac ctxt @{thm Un_upper1}, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
789 |
REPEAT_DETERM_N n o |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58446
diff
changeset
|
790 |
EVERY' (@{map 3} (fn i => fn set_Jset => fn set_Jset_Jsets => |
60728 | 791 |
EVERY' [rtac ctxt subsetI, rtac ctxt @{thm UnI2}, rtac ctxt (mk_UnIN n i), etac ctxt @{thm UN_I}, |
792 |
etac ctxt UnE, etac ctxt set_Jset, REPEAT_DETERM_N (n - 1) o etac ctxt UnE, |
|
60752 | 793 |
EVERY' (map (fn thm => EVERY' [etac ctxt @{thm UN_E}, etac ctxt thm, assume_tac ctxt]) set_Jset_Jsets)]) |
56179 | 794 |
(1 upto n) set_Jsets set_Jset_Jsetss)] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
795 |
|
60728 | 796 |
fun mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets = |
797 |
EVERY' [rtac ctxt @{thm Un_least}, rtac ctxt set_incl_Jset, |
|
798 |
REPEAT_DETERM_N (n - 1) o rtac ctxt @{thm Un_least}, |
|
799 |
EVERY' (map (fn thm => rtac ctxt @{thm UN_least} THEN' etac ctxt thm) set_Jset_incl_Jsets)] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
800 |
|
60728 | 801 |
fun mk_map_id0_tac ctxt maps unfold_unique unfold_dtor = |
802 |
EVERY' [rtac ctxt (unfold_unique RS trans), EVERY' (map (rtac ctxt o mk_sym) maps), |
|
803 |
rtac ctxt unfold_dtor] 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
804 |
|
60728 | 805 |
fun mk_map_comp0_tac ctxt maps map_comp0s map_unique = |
806 |
EVERY' [rtac ctxt map_unique, |
|
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54793
diff
changeset
|
807 |
EVERY' (map2 (fn map_thm => fn map_comp0 => |
60728 | 808 |
EVERY' (map (rtac ctxt) |
55067 | 809 |
[@{thm comp_assoc[symmetric]} RS trans, |
67399 | 810 |
@{thm arg_cong2[of _ _ _ _ "(\<circ>)"]} OF [map_thm, refl] RS trans, |
55067 | 811 |
@{thm comp_assoc[symmetric]} RS sym RS trans, map_thm RS arg_cong RS trans, |
812 |
@{thm comp_assoc[symmetric]} RS trans, |
|
67399 | 813 |
@{thm arg_cong2[of _ _ _ _ "(\<circ>)"]} OF [map_comp0 RS sym, refl]])) |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54793
diff
changeset
|
814 |
maps map_comp0s)] 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
815 |
|
56179 | 816 |
fun mk_mcong_tac ctxt m coinduct_tac map_comps dtor_maps map_cong0s set_map0ss set_Jsetss |
817 |
set_Jset_Jsetsss in_rels = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
818 |
let |
53288 | 819 |
val n = length map_comps; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
820 |
val ks = 1 upto n; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
821 |
in |
60728 | 822 |
EVERY' ([rtac ctxt rev_mp, coinduct_tac] @ |
56179 | 823 |
maps (fn ((((((map_comp_trans, dtor_maps_trans), map_cong0), set_map0s), set_Jsets), |
824 |
set_Jset_Jsetss), in_rel) => |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
825 |
[REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2], |
60728 | 826 |
REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], hyp_subst_tac ctxt, rtac ctxt exI, |
827 |
rtac ctxt (Drule.rotate_prems 1 conjI), |
|
828 |
rtac ctxt conjI, rtac ctxt map_comp_trans, rtac ctxt sym, rtac ctxt dtor_maps_trans, rtac ctxt map_cong0, |
|
829 |
REPEAT_DETERM_N m o (rtac ctxt o_apply_trans_sym THEN' rtac ctxt @{thm fst_conv}), |
|
830 |
REPEAT_DETERM_N n o rtac ctxt fst_convol_fun_cong_sym, |
|
831 |
rtac ctxt map_comp_trans, rtac ctxt sym, rtac ctxt dtor_maps_trans, rtac ctxt map_cong0, |
|
56179 | 832 |
EVERY' (maps (fn set_Jset => |
60728 | 833 |
[rtac ctxt o_apply_trans_sym, rtac ctxt (@{thm snd_conv} RS trans), etac ctxt CollectE, |
834 |
REPEAT_DETERM o etac ctxt conjE, etac ctxt bspec, etac ctxt set_Jset]) set_Jsets), |
|
835 |
REPEAT_DETERM_N n o rtac ctxt snd_convol_fun_cong_sym, |
|
836 |
rtac ctxt CollectI, |
|
837 |
EVERY' (map (fn set_map0 => EVERY' [rtac ctxt conjI, rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, |
|
838 |
rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, rtac ctxt refl]) |
|
55644 | 839 |
(take m set_map0s)), |
56179 | 840 |
CONJ_WRAP' (fn (set_map0, set_Jset_Jsets) => |
60728 | 841 |
EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, |
842 |
rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, rtac ctxt exI, rtac ctxt conjI, |
|
843 |
rtac ctxt CollectI, etac ctxt CollectE, |
|
844 |
REPEAT_DETERM o etac ctxt conjE, |
|
56179 | 845 |
CONJ_WRAP' (fn set_Jset_Jset => |
60752 | 846 |
EVERY' [rtac ctxt ballI, etac ctxt bspec, etac ctxt set_Jset_Jset, assume_tac ctxt]) set_Jset_Jsets, |
60728 | 847 |
rtac ctxt (conjI OF [refl, refl])]) |
56179 | 848 |
(drop m set_map0s ~~ set_Jset_Jsetss)]) |
53288 | 849 |
(map (fn th => th RS trans) map_comps ~~ map (fn th => th RS trans) dtor_maps ~~ |
56179 | 850 |
map_cong0s ~~ set_map0ss ~~ set_Jsetss ~~ set_Jset_Jsetsss ~~ in_rels) @ |
60728 | 851 |
[rtac ctxt impI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
852 |
CONJ_WRAP' (fn k => |
60728 | 853 |
EVERY' [rtac ctxt impI, dtac ctxt (mk_conjunctN n k), etac ctxt mp, rtac ctxt exI, rtac ctxt conjI, etac ctxt CollectI, |
854 |
rtac ctxt conjI, rtac ctxt refl, rtac ctxt refl]) ks]) 1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
855 |
end |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
856 |
|
55197 | 857 |
fun mk_dtor_map_unique_tac ctxt unfold_unique sym_map_comps = |
60728 | 858 |
rtac ctxt unfold_unique 1 THEN |
55067 | 859 |
unfold_thms_tac ctxt (sym_map_comps @ @{thms comp_assoc[symmetric] id_comp comp_id}) THEN |
60728 | 860 |
ALLGOALS (etac ctxt sym); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
861 |
|
55197 | 862 |
fun mk_col_natural_tac ctxt cts rec_0s rec_Sucs dtor_maps set_map0ss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
863 |
let |
49541 | 864 |
val n = length dtor_maps; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
865 |
in |
60801 | 866 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 867 |
REPEAT_DETERM o rtac ctxt allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s), |
868 |
CONJ_WRAP' (K (rtac ctxt @{thm image_empty})) rec_0s, |
|
869 |
REPEAT_DETERM o rtac ctxt allI, |
|
49541 | 870 |
CONJ_WRAP' (fn (rec_Suc, (dtor_map, set_nats)) => EVERY' |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
871 |
[SELECT_GOAL (unfold_thms_tac ctxt |
49541 | 872 |
(rec_Suc :: dtor_map :: set_nats @ @{thms image_Un image_UN UN_simps(10)})), |
60728 | 873 |
rtac ctxt Un_cong, rtac ctxt refl, |
874 |
CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 Un_cong)) |
|
875 |
(fn i => EVERY' [rtac ctxt @{thm SUP_cong[OF refl]}, |
|
876 |
REPEAT_DETERM o etac ctxt allE, etac ctxt (mk_conjunctN n i)]) (n downto 1)]) |
|
53289 | 877 |
(rec_Sucs ~~ (dtor_maps ~~ set_map0ss))] 1 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
878 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
879 |
|
60728 | 880 |
fun mk_set_map0_tac ctxt col_natural = |
881 |
EVERY' (map (rtac ctxt) [@{thm ext}, o_apply RS trans, sym, o_apply RS trans, |
|
882 |
@{thm image_UN} RS trans, refl RS @{thm SUP_cong}, col_natural]) 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
883 |
|
75624 | 884 |
fun mk_col_bd_tac ctxt m j cts rec_0s rec_Sucs sbd_regularCard sbd_Cinfinite set_sbdss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
885 |
let |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
886 |
val n = length rec_0s; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
887 |
in |
60801 | 888 |
EVERY' [rtac ctxt (Thm.instantiate' [] cts nat_induct), |
60728 | 889 |
REPEAT_DETERM o rtac ctxt allI, |
75624 | 890 |
CONJ_WRAP' (fn rec_0 => EVERY' (map (rtac ctxt) [@{thm ordIso_ordLess_trans}, |
891 |
@{thm card_of_ordIso_subst}, rec_0, @{thm Cinfinite_gt_empty}, sbd_Cinfinite])) rec_0s, |
|
60728 | 892 |
REPEAT_DETERM o rtac ctxt allI, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
893 |
CONJ_WRAP' (fn (rec_Suc, set_sbds) => EVERY' |
75624 | 894 |
[rtac ctxt @{thm ordIso_ordLess_trans}, rtac ctxt @{thm card_of_ordIso_subst}, |
895 |
rtac ctxt rec_Suc, rtac ctxt (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound_strict})), |
|
896 |
rtac ctxt (nth set_sbds (j - 1)), |
|
897 |
REPEAT_DETERM_N (n - 1) o rtac ctxt (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound_strict})), |
|
898 |
EVERY' (map2 (fn i => fn set_sbd => EVERY' [ |
|
899 |
rtac ctxt (@{thm card_of_UNION_ordLess_infinite_Field_regularCard} OF [sbd_regularCard, |
|
900 |
sbd_Cinfinite]), rtac ctxt set_sbd, rtac ctxt ballI, REPEAT_DETERM o etac ctxt allE, |
|
901 |
etac ctxt (mk_conjunctN n i)]) (1 upto n) (drop m set_sbds))]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
902 |
(rec_Sucs ~~ set_sbdss)] 1 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
903 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
904 |
|
75624 | 905 |
fun mk_set_bd_tac ctxt sbd_Cinfinite sbd_regularCard natLeq_ordLess_bd col_bd = |
906 |
EVERY' (map (rtac ctxt) [@{thm card_of_UNION_ordLess_infinite_Field_regularCard} OF [ |
|
907 |
sbd_regularCard, sbd_Cinfinite], @{thm ordIso_ordLess_trans}, @{thm card_of_nat}, |
|
908 |
natLeq_ordLess_bd, ballI, col_bd]) 1; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
909 |
|
60728 | 910 |
fun mk_le_rel_OO_tac ctxt coinduct rel_Jrels le_rel_OOs = |
911 |
EVERY' (rtac ctxt coinduct :: map2 (fn rel_Jrel => fn le_rel_OO => |
|
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54793
diff
changeset
|
912 |
let val Jrel_imp_rel = rel_Jrel RS iffD1; |
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54793
diff
changeset
|
913 |
in |
60728 | 914 |
EVERY' [rtac ctxt (le_rel_OO RS @{thm predicate2D}), etac ctxt @{thm relcomppE}, |
915 |
rtac ctxt @{thm relcomppI}, etac ctxt Jrel_imp_rel, etac ctxt Jrel_imp_rel] |
|
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54793
diff
changeset
|
916 |
end) |
57726 | 917 |
rel_Jrels le_rel_OOs) 1; |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
918 |
|
55197 | 919 |
fun mk_wit_tac ctxt n dtor_ctors dtor_set wit coind_wits = |
60728 | 920 |
ALLGOALS (TRY o (eresolve_tac ctxt coind_wits THEN' rtac ctxt refl)) THEN |
60752 | 921 |
REPEAT_DETERM (assume_tac ctxt 1 ORELSE |
60728 | 922 |
EVERY' [dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt dtor_set, |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
923 |
K (unfold_thms_tac ctxt dtor_ctors), |
60728 | 924 |
REPEAT_DETERM_N n o etac ctxt UnE, |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
925 |
REPEAT_DETERM o |
60728 | 926 |
(TRY o REPEAT_DETERM o etac ctxt UnE THEN' TRY o etac ctxt @{thm UN_E} THEN' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
927 |
(eresolve_tac ctxt wit ORELSE' |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
928 |
(dresolve_tac ctxt wit THEN' |
60728 | 929 |
(etac ctxt FalseE ORELSE' |
930 |
EVERY' [hyp_subst_tac ctxt, dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt dtor_set, |
|
931 |
K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac ctxt UnE]))))] 1); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
932 |
|
55197 | 933 |
fun mk_coind_wit_tac ctxt induct unfolds set_nats wits = |
60728 | 934 |
rtac ctxt induct 1 THEN ALLGOALS (TRY o rtac ctxt impI THEN' TRY o hyp_subst_tac ctxt) THEN |
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49501
diff
changeset
|
935 |
unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN |
60728 | 936 |
ALLGOALS (REPEAT_DETERM o etac ctxt imageE THEN' TRY o hyp_subst_tac ctxt) THEN |
49104 | 937 |
ALLGOALS (TRY o |
60728 | 938 |
FIRST' [rtac ctxt TrueI, rtac ctxt refl, etac ctxt (refl RSN (2, mp)), dresolve_tac ctxt wits THEN' etac ctxt FalseE]); |
58445 | 939 |
|
940 |
fun mk_dtor_corec_transfer_tac ctxt n m dtor_corec_defs dtor_unfold_transfer pre_T_map_transfers |
|
941 |
dtor_rels = |
|
58446 | 942 |
CONJ_WRAP (fn (dtor_corec_def, dtor_unfold_transfer) => |
60728 | 943 |
REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN |
58446 | 944 |
unfold_thms_tac ctxt [dtor_corec_def, o_apply] THEN |
60728 | 945 |
HEADGOAL (rtac ctxt (mk_rel_funDN (n + 1) dtor_unfold_transfer) THEN' |
58446 | 946 |
EVERY' (map2 (fn pre_T_map_transfer => fn dtor_rel => |
60728 | 947 |
etac ctxt (mk_rel_funDN_rotated 2 @{thm case_sum_transfer}) THEN' |
948 |
rtac ctxt (mk_rel_funDN 2 @{thm comp_transfer}) THEN' |
|
949 |
rtac ctxt (mk_rel_funDN (m + n) pre_T_map_transfer) THEN' |
|
950 |
REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer} THEN' |
|
951 |
REPEAT_DETERM_N n o rtac ctxt @{thm Inl_transfer} THEN' |
|
952 |
rtac ctxt rel_funI THEN' |
|
953 |
etac ctxt (dtor_rel RS iffD1)) pre_T_map_transfers dtor_rels) THEN' |
|
954 |
etac ctxt (mk_rel_funDN 1 @{thm Inr_transfer}))) |
|
58446 | 955 |
(dtor_corec_defs ~~ dtor_unfold_transfer); |
49104 | 956 |
|
53287 | 957 |
fun mk_dtor_rel_tac ctxt in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets dtor_inject |
53289 | 958 |
dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss = |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
959 |
let |
49544
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents:
49543
diff
changeset
|
960 |
val m = length dtor_set_incls; |
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
blanchet
parents:
49543
diff
changeset
|
961 |
val n = length dtor_set_set_inclss; |
53289 | 962 |
val (passive_set_map0s, active_set_map0s) = chop m set_map0s; |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
963 |
val in_Jrel = nth in_Jrels (i - 1); |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
964 |
val if_tac = |
60728 | 965 |
EVERY' [dtac ctxt (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], |
966 |
rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, |
|
53289 | 967 |
EVERY' (map2 (fn set_map0 => fn set_incl => |
60728 | 968 |
EVERY' [rtac ctxt conjI, rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, |
969 |
rtac ctxt ord_eq_le_trans, rtac ctxt trans_fun_cong_image_id_id_apply, |
|
970 |
etac ctxt (set_incl RS @{thm subset_trans})]) |
|
53289 | 971 |
passive_set_map0s dtor_set_incls), |
972 |
CONJ_WRAP' (fn (in_Jrel, (set_map0, dtor_set_set_incls)) => |
|
60728 | 973 |
EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, |
974 |
rtac ctxt @{thm case_prodI}, rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, |
|
60752 | 975 |
CONJ_WRAP' (fn thm => etac ctxt (thm RS @{thm subset_trans}) THEN' assume_tac ctxt) dtor_set_set_incls, |
60728 | 976 |
rtac ctxt conjI, rtac ctxt refl, rtac ctxt refl]) |
53289 | 977 |
(in_Jrels ~~ (active_set_map0s ~~ dtor_set_set_inclss)), |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
978 |
CONJ_WRAP' (fn conv => |
60728 | 979 |
EVERY' [rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0, |
980 |
REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]}, |
|
981 |
REPEAT_DETERM_N n o EVERY' (map (rtac ctxt) [trans, o_apply, conv]), |
|
60752 | 982 |
rtac ctxt trans, rtac ctxt sym, rtac ctxt dtor_map, rtac ctxt (dtor_inject RS iffD2), assume_tac ctxt]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
983 |
@{thms fst_conv snd_conv}]; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
984 |
val only_if_tac = |
60728 | 985 |
EVERY' [dtac ctxt (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], |
986 |
rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI, |
|
53289 | 987 |
CONJ_WRAP' (fn (dtor_set, passive_set_map0) => |
60728 | 988 |
EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt dtor_set, rtac ctxt @{thm Un_least}, |
989 |
rtac ctxt ord_eq_le_trans, rtac ctxt @{thm box_equals}, rtac ctxt passive_set_map0, |
|
60752 | 990 |
rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans_fun_cong_image_id_id_apply, assume_tac ctxt, |
60728 | 991 |
CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least})) |
992 |
(fn (active_set_map0, in_Jrel) => EVERY' [rtac ctxt ord_eq_le_trans, |
|
993 |
rtac ctxt @{thm SUP_cong[OF _ refl]}, rtac ctxt @{thm box_equals[OF _ _ refl]}, |
|
994 |
rtac ctxt active_set_map0, rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt @{thm UN_least}, |
|
995 |
dtac ctxt set_rev_mp, etac ctxt @{thm image_mono}, etac ctxt imageE, |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
996 |
dtac ctxt @{thm ssubst_mem[OF prod.collapse]}, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
997 |
REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE :: |
56765 | 998 |
@{thms case_prodE iffD1[OF prod.inject, elim_format]}), |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51850
diff
changeset
|
999 |
hyp_subst_tac ctxt, |
60728 | 1000 |
dtac ctxt (in_Jrel RS iffD1), |
1001 |
dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, |
|
60752 | 1002 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], assume_tac ctxt]) |
53289 | 1003 |
(rev (active_set_map0s ~~ in_Jrels))]) |
1004 |
(dtor_sets ~~ passive_set_map0s), |
|
60728 | 1005 |
rtac ctxt conjI, |
1006 |
REPEAT_DETERM_N 2 o EVERY'[rtac ctxt (dtor_inject RS iffD1), rtac ctxt trans, rtac ctxt dtor_map, |
|
1007 |
rtac ctxt @{thm box_equals}, rtac ctxt map_comp0, rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans, |
|
1008 |
rtac ctxt map_cong0, REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]}, |
|
60752 | 1009 |
EVERY' (map (fn in_Jrel => EVERY' [rtac ctxt trans, rtac ctxt o_apply, dtac ctxt set_rev_mp, assume_tac ctxt, |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60801
diff
changeset
|
1010 |
dtac ctxt @{thm ssubst_mem[OF prod.collapse]}, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1011 |
REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE :: |
56765 | 1012 |
@{thms case_prodE iffD1[OF prod.inject, elim_format]}), |
60728 | 1013 |
hyp_subst_tac ctxt, dtac ctxt (in_Jrel RS iffD1), |
60752 | 1014 |
dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt]) in_Jrels), |
1015 |
assume_tac ctxt]] |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1016 |
in |
60728 | 1017 |
EVERY' [rtac ctxt iffI, if_tac, only_if_tac] 1 |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1018 |
end; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1019 |
|
58445 | 1020 |
fun mk_rel_coinduct_coind_tac ctxt fst m coinduct ks map_comp0s map_congs map_arg_congs set_map0ss |
55644 | 1021 |
dtor_unfolds dtor_maps in_rels = |
1022 |
let |
|
1023 |
val n = length ks; |
|
1024 |
val fst_diag_nth = if fst then @{thm fst_diag_fst} else @{thm fst_diag_snd}; |
|
1025 |
val snd_diag_nth = if fst then @{thm snd_diag_fst} else @{thm snd_diag_snd}; |
|
51925 | 1026 |
in |
60728 | 1027 |
EVERY' [rtac ctxt coinduct, |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58446
diff
changeset
|
1028 |
EVERY' (@{map 8} (fn i => fn map_comp0 => fn map_cong => fn map_arg_cong => fn set_map0s => |
55644 | 1029 |
fn dtor_unfold => fn dtor_map => fn in_rel => |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1030 |
EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2], |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1031 |
REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], |
60752 | 1032 |
select_prem_tac ctxt (length ks) (dtac ctxt @{thm spec2}) i, dtac ctxt mp, assume_tac ctxt, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1033 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt, |
60728 | 1034 |
rtac ctxt exI, rtac ctxt (Drule.rotate_prems 1 conjI), rtac ctxt conjI, |
1035 |
rtac ctxt (map_comp0 RS trans), rtac ctxt (dtor_map RS trans RS sym), |
|
1036 |
rtac ctxt (dtor_unfold RS map_arg_cong RS trans), rtac ctxt (trans OF [map_comp0, map_cong]), |
|
1037 |
REPEAT_DETERM_N m o rtac ctxt (fst_diag_nth RS @{thm fun_cong[OF trans[OF o_id sym]]}), |
|
1038 |
REPEAT_DETERM_N n o (rtac ctxt @{thm sym[OF trans[OF o_apply]]} THEN' rtac ctxt @{thm fst_conv}), |
|
1039 |
rtac ctxt (map_comp0 RS trans), rtac ctxt (map_cong RS trans), |
|
1040 |
REPEAT_DETERM_N m o rtac ctxt (snd_diag_nth RS fun_cong), |
|
1041 |
REPEAT_DETERM_N n o (rtac ctxt @{thm trans[OF o_apply]} THEN' rtac ctxt @{thm snd_conv}), |
|
1042 |
etac ctxt (@{thm prod.case} RS map_arg_cong RS trans), |
|
55644 | 1043 |
SELECT_GOAL (unfold_thms_tac ctxt @{thms prod.case o_def fst_conv snd_conv}), |
60728 | 1044 |
rtac ctxt CollectI, |
53289 | 1045 |
CONJ_WRAP' (fn set_map0 => |
60728 | 1046 |
EVERY' [rtac ctxt (set_map0 RS ord_eq_le_trans), |
1047 |
rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, |
|
60752 | 1048 |
FIRST' [rtac ctxt refl, EVERY'[rtac ctxt exI, rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, |
1049 |
assume_tac ctxt, rtac ctxt (@{thm surjective_pairing} RS arg_cong)]]]) |
|
55644 | 1050 |
set_map0s]) |
1051 |
ks map_comp0s map_congs map_arg_congs set_map0ss dtor_unfolds dtor_maps in_rels)] 1 |
|
51925 | 1052 |
end; |
1053 |
||
55642
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents:
55606
diff
changeset
|
1054 |
val split_id_unfolds = @{thms prod.case image_id id_apply}; |
51925 | 1055 |
|
55197 | 1056 |
fun mk_rel_coinduct_ind_tac ctxt m ks unfolds set_map0ss j set_induct = |
51925 | 1057 |
let val n = length ks; |
1058 |
in |
|
60728 | 1059 |
rtac ctxt set_induct 1 THEN |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58446
diff
changeset
|
1060 |
EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i => |
60728 | 1061 |
EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt conjE, |
60752 | 1062 |
select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, assume_tac ctxt, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1063 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, Collect_splitD_set_mp, set_rev_mp], |
51925 | 1064 |
hyp_subst_tac ctxt, |
53289 | 1065 |
SELECT_GOAL (unfold_thms_tac ctxt ([unfold, nth set_map0s (j - 1)] @ split_id_unfolds)), |
60728 | 1066 |
rtac ctxt subset_refl]) |
53289 | 1067 |
unfolds set_map0ss ks) 1 THEN |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58446
diff
changeset
|
1068 |
EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i => |
53289 | 1069 |
EVERY' (map (fn set_map0 => |
60728 | 1070 |
EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt conjE, |
60752 | 1071 |
select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, assume_tac ctxt, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1072 |
REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt, |
53289 | 1073 |
SELECT_GOAL (unfold_thms_tac ctxt ([unfold, set_map0] @ split_id_unfolds)), |
60728 | 1074 |
etac ctxt imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp], |
60752 | 1075 |
rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, assume_tac ctxt, |
1076 |
rtac ctxt (@{thm surjective_pairing} RS arg_cong)]) |
|
53289 | 1077 |
(drop m set_map0s))) |
1078 |
unfolds set_map0ss ks) 1 |
|
51925 | 1079 |
end; |
1080 |
||
55197 | 1081 |
fun mk_rel_coinduct_tac ctxt CIHs in_rels in_Jrels helper_indss helper_coind1s helper_coind2s = |
51925 | 1082 |
let val n = length in_rels; |
1083 |
in |
|
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61424
diff
changeset
|
1084 |
Method.insert_tac ctxt CIHs 1 THEN |
51925 | 1085 |
unfold_thms_tac ctxt (@{thm choice_iff} :: @{thm ex_simps(6)[symmetric]} :: in_rels) THEN |
60728 | 1086 |
REPEAT_DETERM (etac ctxt exE 1) THEN |
51925 | 1087 |
CONJ_WRAP' (fn (in_Jrel, (helper_inds, (helper_coind1, helper_coind2))) => |
60728 | 1088 |
EVERY' [rtac ctxt @{thm predicate2I}, rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, |
1089 |
if null helper_inds then rtac ctxt UNIV_I |
|
1090 |
else rtac ctxt CollectI THEN' |
|
51925 | 1091 |
CONJ_WRAP' (fn helper_ind => |
60752 | 1092 |
EVERY' [rtac ctxt (helper_ind RS rev_mp), REPEAT_DETERM_N n o assume_tac ctxt, |
60728 | 1093 |
REPEAT_DETERM_N n o etac ctxt thin_rl, rtac ctxt impI, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58634
diff
changeset
|
1094 |
REPEAT_DETERM o resolve_tac ctxt [subsetI, CollectI, @{thm iffD2[OF split_beta]}], |
60752 | 1095 |
dtac ctxt bspec, assume_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp, conjE], |
60728 | 1096 |
etac ctxt (refl RSN (2, conjI))]) |
51925 | 1097 |
helper_inds, |
60728 | 1098 |
rtac ctxt conjI, |
60752 | 1099 |
rtac ctxt (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o assume_tac ctxt, |
1100 |
REPEAT_DETERM_N n o etac ctxt thin_rl, |
|
60728 | 1101 |
rtac ctxt impI, etac ctxt mp, rtac ctxt exI, etac ctxt (refl RSN (2, conjI)), |
60752 | 1102 |
rtac ctxt (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o assume_tac ctxt, |
1103 |
REPEAT_DETERM_N n o etac ctxt thin_rl, |
|
60728 | 1104 |
rtac ctxt impI, etac ctxt mp, rtac ctxt exI, etac ctxt (refl RSN (2, conjI))]) |
51925 | 1105 |
(in_Jrels ~~ (helper_indss ~~ (helper_coind1s ~~ helper_coind2s))) 1 |
1106 |
end; |
|
1107 |
||
55901
8c6d49dd8ae1
renamed a pair of low-level theorems to have c/dtor in their names (like the others)
blanchet
parents:
55644
diff
changeset
|
1108 |
fun mk_unfold_transfer_tac ctxt m ctor_rel_coinduct map_transfers unfolds = |
52731 | 1109 |
let |
1110 |
val n = length map_transfers; |
|
1111 |
in |
|
1112 |
unfold_thms_tac ctxt |
|
55945 | 1113 |
@{thms rel_fun_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN |
1114 |
unfold_thms_tac ctxt @{thms rel_fun_iff_geq_image2p} THEN |
|
52731 | 1115 |
HEADGOAL (EVERY' |
60728 | 1116 |
[REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctxt ctor_rel_coinduct, |
52731 | 1117 |
EVERY' (map (fn map_transfer => EVERY' |
60728 | 1118 |
[REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt @{thm image2pE}, hyp_subst_tac ctxt, |
52731 | 1119 |
SELECT_GOAL (unfold_thms_tac ctxt unfolds), |
60728 | 1120 |
rtac ctxt (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer), |
1121 |
REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer}, |
|
1122 |
REPEAT_DETERM_N n o rtac ctxt @{thm rel_fun_image2p}, |
|
1123 |
etac ctxt @{thm predicate2D}, etac ctxt @{thm image2pI}]) |
|
52731 | 1124 |
map_transfers)]) |
1125 |
end; |
|
1126 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1127 |
end; |