| author | nipkow | 
| Mon, 16 Oct 2017 08:00:28 +0200 | |
| changeset 66870 | f801b36d7c4e | 
| parent 62343 | 24106dc44def | 
| child 75624 | 22d1c5f2b9f4 | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_comp_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: Jasmin Blanchette, TU Muenchen  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
Copyright 2012  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
Tactics for composition of bounded natural functors.  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
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  | 
signature BNF_COMP_TACTICS =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 60728 | 11  | 
val mk_comp_bd_card_order_tac: Proof.context -> thm list -> thm -> tactic  | 
12  | 
val mk_comp_bd_cinfinite_tac: Proof.context -> thm -> thm -> tactic  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
13  | 
val mk_comp_in_alt_tac: Proof.context -> thm list -> tactic  | 
| 60728 | 14  | 
val mk_comp_map_comp0_tac: Proof.context -> thm -> thm -> thm list -> tactic  | 
| 55906 | 15  | 
val mk_comp_map_cong0_tac: Proof.context -> thm list -> thm list -> thm -> thm list -> tactic  | 
| 60728 | 16  | 
val mk_comp_map_id0_tac: Proof.context -> thm -> thm -> thm list -> tactic  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
17  | 
val mk_comp_set_alt_tac: Proof.context -> thm -> tactic  | 
| 55906 | 18  | 
val mk_comp_set_bd_tac: Proof.context -> thm -> thm option -> thm -> thm list -> tactic  | 
19  | 
val mk_comp_set_map0_tac: Proof.context -> thm -> thm -> thm -> thm -> thm list -> tactic  | 
|
20  | 
val mk_comp_wit_tac: Proof.context -> thm list -> thm list -> thm -> thm list -> tactic  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
21  | 
|
| 60728 | 22  | 
val kill_in_alt_tac: Proof.context -> tactic  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
49630 
diff
changeset
 | 
23  | 
val mk_kill_map_cong0_tac: Proof.context -> int -> int -> thm -> tactic  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
24  | 
|
| 60728 | 25  | 
val empty_natural_tac: Proof.context -> tactic  | 
26  | 
val lift_in_alt_tac: Proof.context -> tactic  | 
|
27  | 
val mk_lift_set_bd_tac: Proof.context -> thm -> tactic  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
28  | 
|
| 60728 | 29  | 
val mk_permute_in_alt_tac: Proof.context -> ''a list -> ''a list -> tactic  | 
| 49284 | 30  | 
|
| 60728 | 31  | 
val mk_le_rel_OO_tac: Proof.context -> thm -> thm -> thm list -> tactic  | 
32  | 
val mk_simple_rel_OO_Grp_tac: Proof.context -> thm -> thm -> tactic  | 
|
| 62324 | 33  | 
val mk_simple_pred_set_tac: Proof.context -> thm -> thm -> tactic  | 
| 60752 | 34  | 
val mk_simple_wit_tac: Proof.context -> thm list -> tactic  | 
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
35  | 
val mk_simplified_set_tac: Proof.context -> thm -> tactic  | 
| 60757 | 36  | 
val bd_ordIso_natLeq_tac: Proof.context -> tactic  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
37  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
39  | 
structure BNF_Comp_Tactics : BNF_COMP_TACTICS =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
40  | 
struct  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
42  | 
open BNF_Util  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
43  | 
open BNF_Tactics  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
45  | 
val arg_cong_Union = @{thm arg_cong[of _ _ Union]};
 | 
| 55067 | 46  | 
val comp_eq_dest_lhs = @{thm comp_eq_dest_lhs};
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
47  | 
val trans_image_cong_o_apply = @{thm trans[OF image_cong[OF o_apply refl]]};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
48  | 
val trans_o_apply = @{thm trans[OF o_apply]};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
51  | 
(* Composition *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
52  | 
|
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51761 
diff
changeset
 | 
53  | 
fun mk_comp_set_alt_tac ctxt collect_set_map =  | 
| 55067 | 54  | 
  unfold_thms_tac ctxt @{thms comp_assoc} THEN
 | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51761 
diff
changeset
 | 
55  | 
unfold_thms_tac ctxt [collect_set_map RS sym] THEN  | 
| 60728 | 56  | 
rtac ctxt refl 1;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
|
| 60728 | 58  | 
fun mk_comp_map_id0_tac ctxt Gmap_id0 Gmap_cong0 map_id0s =  | 
59  | 
  EVERY' ([rtac ctxt @{thm ext}, rtac ctxt (Gmap_cong0 RS trans)] @
 | 
|
60  | 
map (fn thm => rtac ctxt (thm RS fun_cong)) map_id0s @ [rtac ctxt (Gmap_id0 RS fun_cong)]) 1;  | 
|
| 49630 | 61  | 
|
| 60728 | 62  | 
fun mk_comp_map_comp0_tac ctxt Gmap_comp0 Gmap_cong0 map_comp0s =  | 
63  | 
  EVERY' ([rtac ctxt @{thm ext}, rtac ctxt sym, rtac ctxt trans_o_apply,
 | 
|
64  | 
rtac ctxt (Gmap_comp0 RS sym RS comp_eq_dest_lhs RS trans), rtac ctxt Gmap_cong0] @  | 
|
65  | 
map (fn thm => rtac ctxt (thm RS sym RS fun_cong)) map_comp0s) 1;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
66  | 
|
| 55906 | 67  | 
fun mk_comp_set_map0_tac ctxt set'_eq_set Gmap_comp0 Gmap_cong0 Gset_map0 set_map0s =  | 
68  | 
unfold_thms_tac ctxt [set'_eq_set] THEN  | 
|
| 60728 | 69  | 
  EVERY' ([rtac ctxt @{thm ext}] @
 | 
70  | 
replicate 3 (rtac ctxt trans_o_apply) @  | 
|
71  | 
[rtac ctxt (arg_cong_Union RS trans),  | 
|
72  | 
     rtac ctxt (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans),
 | 
|
73  | 
rtac ctxt (Gmap_comp0 RS sym RS comp_eq_dest_lhs RS trans),  | 
|
74  | 
rtac ctxt Gmap_cong0] @  | 
|
75  | 
map (fn thm => rtac ctxt (thm RS fun_cong)) set_map0s @  | 
|
76  | 
[rtac ctxt (Gset_map0 RS comp_eq_dest_lhs), rtac ctxt sym, rtac ctxt trans_o_apply,  | 
|
77  | 
rtac ctxt trans_image_cong_o_apply, rtac ctxt trans_image_cong_o_apply,  | 
|
| 61760 | 78  | 
     rtac ctxt (@{thm image_cong} OF [Gset_map0 RS comp_eq_dest_lhs RS arg_cong_Union, refl]
 | 
79  | 
RS trans),  | 
|
| 60728 | 80  | 
     rtac ctxt @{thm trans[OF comp_eq_dest[OF Union_natural[symmetric]]]}, rtac ctxt arg_cong_Union,
 | 
81  | 
     rtac ctxt @{thm trans[OF comp_eq_dest_lhs[OF image_o_collect[symmetric]]]},
 | 
|
82  | 
     rtac ctxt @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @
 | 
|
83  | 
     [REPEAT_DETERM_N (length set_map0s) o EVERY' [rtac ctxt @{thm trans[OF image_insert]},
 | 
|
| 61760 | 84  | 
        rtac ctxt @{thm arg_cong2[of _ _ _ _ insert]}, rtac ctxt @{thm ext},
 | 
85  | 
rtac ctxt trans_o_apply, rtac ctxt trans_image_cong_o_apply,  | 
|
86  | 
        rtac ctxt @{thm trans[OF image_image]}, rtac ctxt @{thm sym[OF trans[OF o_apply]]},
 | 
|
87  | 
        rtac ctxt @{thm image_cong[OF refl o_apply]}],
 | 
|
| 60728 | 88  | 
     rtac ctxt @{thm image_empty}]) 1;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
89  | 
|
| 55906 | 90  | 
fun mk_comp_map_cong0_tac ctxt set'_eq_sets comp_set_alts map_cong0 map_cong0s =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
91  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
92  | 
val n = length comp_set_alts;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
93  | 
in  | 
| 55906 | 94  | 
unfold_thms_tac ctxt set'_eq_sets THEN  | 
| 60728 | 95  | 
(if n = 0 then rtac ctxt refl 1  | 
96  | 
else rtac ctxt map_cong0 1 THEN  | 
|
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
49630 
diff
changeset
 | 
97  | 
EVERY' (map_index (fn (i, map_cong0) =>  | 
| 60728 | 98  | 
rtac ctxt map_cong0 THEN' EVERY' (map_index (fn (k, set_alt) =>  | 
99  | 
          EVERY' [select_prem_tac ctxt n (dtac ctxt @{thm meta_spec}) (k + 1), etac ctxt meta_mp,
 | 
|
100  | 
rtac ctxt (equalityD2 RS set_mp), rtac ctxt (set_alt RS fun_cong RS trans),  | 
|
101  | 
            rtac ctxt trans_o_apply, rtac ctxt (@{thm collect_def} RS arg_cong_Union),
 | 
|
| 61760 | 102  | 
            rtac ctxt @{thm UnionI}, rtac ctxt @{thm UN_I},
 | 
103  | 
            REPEAT_DETERM_N i o rtac ctxt @{thm insertI2}, rtac ctxt @{thm insertI1},
 | 
|
104  | 
            rtac ctxt (o_apply RS equalityD2 RS set_mp), etac ctxt @{thm imageI}, assume_tac ctxt])
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
105  | 
comp_set_alts))  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
49630 
diff
changeset
 | 
106  | 
map_cong0s) 1)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
107  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
108  | 
|
| 60728 | 109  | 
fun mk_comp_bd_card_order_tac ctxt Fbd_card_orders Gbd_card_order =  | 
110  | 
  rtac ctxt @{thm natLeq_card_order} 1 ORELSE
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
111  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
112  | 
val (card_orders, last_card_order) = split_last Fbd_card_orders;  | 
| 60728 | 113  | 
    fun gen_before thm = rtac ctxt @{thm card_order_csum} THEN' rtac ctxt thm;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
114  | 
in  | 
| 60728 | 115  | 
    (rtac ctxt @{thm card_order_cprod} THEN'
 | 
116  | 
WRAP' gen_before (K (K all_tac)) card_orders (rtac ctxt last_card_order) THEN'  | 
|
117  | 
rtac ctxt Gbd_card_order) 1  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
118  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
119  | 
|
| 60728 | 120  | 
fun mk_comp_bd_cinfinite_tac ctxt Fbd_cinfinite Gbd_cinfinite =  | 
121  | 
  (rtac ctxt @{thm natLeq_cinfinite} ORELSE'
 | 
|
122  | 
   rtac ctxt @{thm cinfinite_cprod} THEN'
 | 
|
123  | 
   ((K (TRY ((rtac ctxt @{thm cinfinite_csum} THEN' rtac ctxt disjI1) 1)) THEN'
 | 
|
124  | 
     ((rtac ctxt @{thm cinfinite_csum} THEN' rtac ctxt disjI1 THEN' rtac ctxt Fbd_cinfinite) ORELSE'
 | 
|
125  | 
rtac ctxt Fbd_cinfinite)) ORELSE'  | 
|
126  | 
rtac ctxt Fbd_cinfinite) THEN'  | 
|
127  | 
rtac ctxt Gbd_cinfinite) 1;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
128  | 
|
| 55906 | 129  | 
fun mk_comp_set_bd_tac ctxt set'_eq_set bd_ordIso_natLeq_opt comp_set_alt Gset_Fset_bds =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
130  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
val (bds, last_bd) = split_last Gset_Fset_bds;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
132  | 
fun gen_before bd =  | 
| 60728 | 133  | 
      rtac ctxt ctrans THEN' rtac ctxt @{thm Un_csum} THEN'
 | 
134  | 
      rtac ctxt ctrans THEN' rtac ctxt @{thm csum_mono} THEN'
 | 
|
135  | 
rtac ctxt bd;  | 
|
136  | 
    fun gen_after _ = rtac ctxt @{thm ordIso_imp_ordLeq} THEN' rtac ctxt @{thm cprod_csum_distrib1};
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
137  | 
in  | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
138  | 
(case bd_ordIso_natLeq_opt of  | 
| 60728 | 139  | 
      SOME thm => rtac ctxt (thm RSN (2, @{thm ordLeq_ordIso_trans})) 1
 | 
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
140  | 
| NONE => all_tac) THEN  | 
| 55906 | 141  | 
unfold_thms_tac ctxt [set'_eq_set, comp_set_alt] THEN  | 
| 60728 | 142  | 
    rtac ctxt @{thm comp_set_bd_Union_o_collect} 1 THEN
 | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49490 
diff
changeset
 | 
143  | 
    unfold_thms_tac ctxt @{thms Union_image_insert Union_image_empty Union_Un_distrib o_apply} THEN
 | 
| 60728 | 144  | 
(rtac ctxt ctrans THEN'  | 
145  | 
WRAP' gen_before gen_after bds (rtac ctxt last_bd) THEN'  | 
|
146  | 
     rtac ctxt @{thm ordIso_imp_ordLeq} THEN'
 | 
|
147  | 
     rtac ctxt @{thm cprod_com}) 1
 | 
|
| 
48975
 
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  | 
|
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
150  | 
val comp_in_alt_thms = @{thms o_apply collect_def image_insert image_empty Union_insert UN_insert
 | 
| 
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
151  | 
UN_empty Union_empty Un_empty_right Union_Un_distrib Un_subset_iff conj_subset_def UN_image_subset  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
152  | 
conj_assoc};  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
154  | 
fun mk_comp_in_alt_tac ctxt comp_set_alts =  | 
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
155  | 
unfold_thms_tac ctxt comp_set_alts THEN  | 
| 
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
156  | 
unfold_thms_tac ctxt comp_in_alt_thms THEN  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49490 
diff
changeset
 | 
157  | 
  unfold_thms_tac ctxt @{thms set_eq_subset} THEN
 | 
| 60728 | 158  | 
rtac ctxt conjI 1 THEN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
REPEAT_DETERM (  | 
| 60728 | 160  | 
    rtac ctxt @{thm subsetI} 1 THEN
 | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49490 
diff
changeset
 | 
161  | 
    unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN
 | 
| 60728 | 162  | 
(REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
163  | 
REPEAT_DETERM (CHANGED ((  | 
| 60752 | 164  | 
(rtac ctxt conjI THEN' (assume_tac ctxt ORELSE' rtac ctxt subset_UNIV)) ORELSE'  | 
165  | 
assume_tac ctxt ORELSE'  | 
|
| 60728 | 166  | 
(rtac ctxt subset_UNIV)) 1)) ORELSE rtac ctxt subset_UNIV 1));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
167  | 
|
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
168  | 
val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def UN_insert UN_empty Un_empty_right
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
169  | 
Union_image_insert Union_image_empty};  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
170  | 
|
| 55906 | 171  | 
fun mk_comp_wit_tac ctxt set'_eq_sets Gwit_thms collect_set_map Fwit_thms =  | 
172  | 
unfold_thms_tac ctxt set'_eq_sets THEN  | 
|
| 60728 | 173  | 
  ALLGOALS (dtac ctxt @{thm in_Union_o_assoc}) THEN
 | 
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
174  | 
unfold_thms_tac ctxt [collect_set_map] THEN  | 
| 
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
175  | 
unfold_thms_tac ctxt comp_wit_thms THEN  | 
| 60752 | 176  | 
REPEAT_DETERM ((assume_tac ctxt ORELSE'  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58181 
diff
changeset
 | 
177  | 
    REPEAT_DETERM o eresolve_tac ctxt @{thms UnionE UnE} THEN'
 | 
| 60728 | 178  | 
etac ctxt imageE THEN' TRY o dresolve_tac ctxt Gwit_thms THEN'  | 
179  | 
(etac ctxt FalseE ORELSE'  | 
|
| 51798 | 180  | 
hyp_subst_tac ctxt THEN'  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58181 
diff
changeset
 | 
181  | 
dresolve_tac ctxt Fwit_thms THEN'  | 
| 60752 | 182  | 
(etac ctxt FalseE ORELSE' assume_tac ctxt))) 1);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
185  | 
(* Kill operation *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
186  | 
|
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
49630 
diff
changeset
 | 
187  | 
fun mk_kill_map_cong0_tac ctxt n m map_cong0 =  | 
| 60728 | 188  | 
(rtac ctxt map_cong0 THEN' EVERY' (replicate n (rtac ctxt refl)) THEN'  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
189  | 
EVERY' (replicate m (Goal.assume_rule_tac ctxt))) 1;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
190  | 
|
| 60728 | 191  | 
fun kill_in_alt_tac ctxt =  | 
192  | 
  ((rtac ctxt @{thm Collect_cong} THEN' rtac ctxt iffI) 1 THEN
 | 
|
193  | 
REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN  | 
|
194  | 
REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE'  | 
|
195  | 
rtac ctxt conjI THEN' rtac ctxt subset_UNIV) 1)) THEN  | 
|
| 60752 | 196  | 
(rtac ctxt subset_UNIV ORELSE' assume_tac ctxt) 1 THEN  | 
| 60728 | 197  | 
REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN  | 
| 60752 | 198  | 
REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' assume_tac ctxt) 1))) ORELSE  | 
| 60728 | 199  | 
  ((rtac ctxt @{thm UNIV_eq_I} THEN' rtac ctxt CollectI) 1 THEN
 | 
200  | 
REPEAT_DETERM (TRY (rtac ctxt conjI 1) THEN rtac ctxt subset_UNIV 1));  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
203  | 
(* Lift operation *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
204  | 
|
| 60728 | 205  | 
fun empty_natural_tac ctxt = rtac ctxt @{thm empty_natural} 1;
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
206  | 
|
| 60728 | 207  | 
fun mk_lift_set_bd_tac ctxt bd_Card_order =  | 
208  | 
  (rtac ctxt @{thm Card_order_empty} THEN' rtac ctxt bd_Card_order) 1;
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
209  | 
|
| 60728 | 210  | 
fun lift_in_alt_tac ctxt =  | 
211  | 
  ((rtac ctxt @{thm Collect_cong} THEN' rtac ctxt iffI) 1 THEN
 | 
|
212  | 
REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN  | 
|
| 60752 | 213  | 
REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' assume_tac ctxt) 1)) THEN  | 
| 60728 | 214  | 
REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN  | 
215  | 
REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE'  | 
|
216  | 
    rtac ctxt conjI THEN' rtac ctxt @{thm empty_subsetI}) 1)) THEN
 | 
|
| 60752 | 217  | 
  (rtac ctxt @{thm empty_subsetI} ORELSE' assume_tac ctxt) 1) ORELSE
 | 
| 60728 | 218  | 
  ((rtac ctxt sym THEN' rtac ctxt @{thm UNIV_eq_I} THEN' rtac ctxt CollectI) 1 THEN
 | 
219  | 
    REPEAT_DETERM (TRY (rtac ctxt conjI 1) THEN rtac ctxt @{thm empty_subsetI} 1));
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
220  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
222  | 
(* Permute operation *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
223  | 
|
| 60728 | 224  | 
fun mk_permute_in_alt_tac ctxt src dest =  | 
225  | 
  (rtac ctxt @{thm Collect_cong} THEN'
 | 
|
| 61760 | 226  | 
  mk_rotate_eq_tac ctxt (rtac ctxt refl) trans @{thm conj_assoc} @{thm conj_commute}
 | 
227  | 
    @{thm conj_cong} dest src) 1;
 | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
228  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
229  | 
|
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
230  | 
(* Miscellaneous *)  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
231  | 
|
| 60728 | 232  | 
fun mk_le_rel_OO_tac ctxt outer_le_rel_OO outer_rel_mono inner_le_rel_OOs =  | 
| 62324 | 233  | 
  HEADGOAL (EVERY' (map (rtac ctxt) (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono ::
 | 
234  | 
inner_le_rel_OOs)));  | 
|
| 49284 | 235  | 
|
| 60728 | 236  | 
fun mk_simple_rel_OO_Grp_tac ctxt rel_OO_Grp in_alt_thm =  | 
| 62324 | 237  | 
  HEADGOAL (rtac ctxt (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]));
 | 
238  | 
||
239  | 
fun mk_simple_pred_set_tac ctxt pred_set in_alt_thm =  | 
|
240  | 
HEADGOAL (rtac ctxt (pred_set RS trans)) THEN  | 
|
241  | 
  unfold_thms_tac ctxt @{thms Ball_Collect UNIV_def} THEN
 | 
|
242  | 
  HEADGOAL (rtac ctxt (unfold_thms ctxt @{thms UNIV_def} in_alt_thm RS @{thm Collect_inj} RS sym));
 | 
|
| 49463 | 243  | 
|
| 60752 | 244  | 
fun mk_simple_wit_tac ctxt wit_thms =  | 
| 60757 | 245  | 
  ALLGOALS (assume_tac ctxt ORELSE' eresolve_tac ctxt (@{thm emptyE} :: wit_thms));
 | 
| 49284 | 246  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
247  | 
val csum_thms =  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
248  | 
  @{thms csum_cong1 csum_cong2 csum_cong  csum_dup[OF natLeq_cinfinite natLeq_Card_order]};
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
249  | 
val cprod_thms =  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
250  | 
  @{thms cprod_cong1 cprod_cong2 cprod_cong cprod_dup[OF natLeq_cinfinite natLeq_Card_order]};
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
251  | 
|
| 55906 | 252  | 
val simplified_set_simps =  | 
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
253  | 
  @{thms collect_def[abs_def] UN_insert UN_empty Un_empty_right Un_empty_left
 | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
62324 
diff
changeset
 | 
254  | 
o_def Union_Un_distrib UN_empty2 UN_singleton id_bnf_def};  | 
| 55906 | 255  | 
|
| 
55930
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
256  | 
fun mk_simplified_set_tac ctxt collect_set_map =  | 
| 
 
25a90cebbbe5
more careful simplification of sets (cf. abf91ebd0820)---yields smaller terms
 
traytel 
parents: 
55906 
diff
changeset
 | 
257  | 
  unfold_thms_tac ctxt (collect_set_map :: @{thms comp_assoc}) THEN
 | 
| 60728 | 258  | 
unfold_thms_tac ctxt simplified_set_simps THEN rtac ctxt refl 1;  | 
| 55906 | 259  | 
|
| 60757 | 260  | 
fun bd_ordIso_natLeq_tac ctxt =  | 
261  | 
HEADGOAL (REPEAT_DETERM o resolve_tac ctxt  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
262  | 
    (@{thm ordIso_refl[OF natLeq_Card_order]} :: csum_thms @ cprod_thms));
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55707 
diff
changeset
 | 
263  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
264  | 
end;  |