991
|
1 |
(* Title: ZF/AC/AC_Equiv.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Krzysztof Gr`abczewski
|
|
4 |
|
|
5 |
*)
|
|
6 |
|
|
7 |
|
|
8 |
open AC_Equiv;
|
|
9 |
|
1071
|
10 |
val WO_defs = [WO1_def, WO2_def, WO3_def, WO4_def, WO5_def, WO6_def, WO8_def];
|
|
11 |
|
|
12 |
val AC_defs = [AC0_def, AC1_def, AC2_def, AC3_def, AC4_def, AC5_def,
|
|
13 |
AC6_def, AC7_def, AC8_def, AC9_def, AC10_def, AC11_def,
|
|
14 |
AC12_def, AC13_def, AC14_def, AC15_def, AC16_def,
|
|
15 |
AC17_def, AC18_def, AC19_def];
|
|
16 |
|
991
|
17 |
val AC_aux_defs = [pairwise_disjoint_def, sets_of_size_between_def];
|
|
18 |
|
|
19 |
(* ******************************************** *)
|
|
20 |
|
|
21 |
(* Theorems analogous to ones presented in "ZF/Ordinal.ML" *)
|
|
22 |
|
|
23 |
(* lemma for nat_le_imp_lepoll *)
|
|
24 |
val [prem] = goalw Cardinal.thy [lepoll_def]
|
|
25 |
"m:nat ==> ALL n: nat. m le n --> m lepoll n";
|
|
26 |
by (nat_ind_tac "m" [prem] 1);
|
|
27 |
by (fast_tac (ZF_cs addSIs [le_imp_subset RS id_subset_inj]) 1);
|
|
28 |
by (rtac ballI 1);
|
|
29 |
by (eres_inst_tac [("n","n")] natE 1);
|
|
30 |
by (asm_simp_tac (ZF_ss addsimps [inj_def, succI1 RS Pi_empty2]) 1);
|
|
31 |
by (fast_tac (ZF_cs addSDs [le0D]) 1);
|
|
32 |
by (fast_tac (ZF_cs addSIs [le_imp_subset RS id_subset_inj]) 1);
|
|
33 |
val nat_le_imp_lepoll_lemma = result();
|
|
34 |
|
|
35 |
(* used in : AC10-AC15.ML WO1-WO6.ML WO6WO1.ML*)
|
|
36 |
val nat_le_imp_lepoll = nat_le_imp_lepoll_lemma RS bspec RS mp |> standard;
|
|
37 |
|
|
38 |
(* ********************************************************************** *)
|
|
39 |
(* lemmas concerning FOL and pure ZF theory *)
|
|
40 |
(* ********************************************************************** *)
|
|
41 |
|
|
42 |
(* The following two theorms are useful when rewriting only one instance *)
|
|
43 |
(* of a definition *)
|
|
44 |
(* first one for definitions of formulae and the second for terms *)
|
|
45 |
|
|
46 |
val prems = goal ZF.thy "(A == B) ==> A <-> B";
|
|
47 |
by (asm_simp_tac (ZF_ss addsimps prems) 1);
|
|
48 |
val def_imp_iff = result();
|
|
49 |
|
|
50 |
val prems = goal ZF.thy "(A == B) ==> P(A) <-> P(B)";
|
|
51 |
by (asm_simp_tac (ZF_ss addsimps prems) 1);
|
|
52 |
val def_imp_iff_P = result();
|
|
53 |
|
|
54 |
(* used only in lemmas4-7.ML *)
|
|
55 |
(*Note modified statement and proof*)
|
|
56 |
goal ZF.thy "!!X. X~=0 ==> (A->X)~=0";
|
|
57 |
by (fast_tac (ZF_cs addSIs [equals0I,lam_type] addSEs [equals0D]) 1);
|
|
58 |
val fun_space_not0I = result();
|
|
59 |
|
|
60 |
(* used only in WO1_DC.ML *)
|
|
61 |
(*Note simpler proof*)
|
|
62 |
goal ZF.thy "!!A f g. [| ALL x:A. f`x=g`x; f:Df->Cf; g:Dg->Cg; \
|
|
63 |
\ A<=Df; A<=Dg |] ==> f``A=g``A";
|
|
64 |
by (asm_simp_tac (ZF_ss addsimps [image_fun]) 1);
|
|
65 |
val images_eq = result();
|
|
66 |
|
|
67 |
(* used in : AC10-AC15.ML AC16WO4.ML WO6WO1.ML *)
|
|
68 |
(*I don't know where to put this one.*)
|
|
69 |
goal Cardinal.thy
|
|
70 |
"!!m A B. [| A lepoll succ(m); B<=A; B~=0 |] ==> A-B lepoll m";
|
|
71 |
by (resolve_tac [not_emptyE] 1 THEN (atac 1));
|
|
72 |
by (forward_tac [singleton_subsetI] 1);
|
|
73 |
by (forward_tac [subsetD] 1 THEN (atac 1));
|
|
74 |
by (res_inst_tac [("A2","A")]
|
1057
|
75 |
(Diff_sing_lepoll RSN (2, subset_imp_lepoll RS lepoll_trans)) 1
|
991
|
76 |
THEN (REPEAT (atac 2)));
|
|
77 |
by (fast_tac ZF_cs 1);
|
|
78 |
val Diff_lepoll = result();
|
|
79 |
|
|
80 |
(* ********************************************************************** *)
|
|
81 |
(* lemmas concerning lepoll and eqpoll relations *)
|
|
82 |
(* ********************************************************************** *)
|
|
83 |
|
|
84 |
(* ********************************************************************** *)
|
|
85 |
(* Theorems concerning ordinals *)
|
|
86 |
(* ********************************************************************** *)
|
|
87 |
|
|
88 |
(* lemma for ordertype_Int *)
|
|
89 |
goalw Cardinal.thy [rvimage_def] "rvimage(A,id(A),r) = r Int A*A";
|
|
90 |
by (resolve_tac [equalityI] 1);
|
|
91 |
by (step_tac ZF_cs 1);
|
|
92 |
by (dres_inst_tac [("P","%a. <id(A)`xb,a>:r")] (id_conv RS subst) 1
|
|
93 |
THEN (atac 1));
|
|
94 |
by (dres_inst_tac [("P","%a. <a,ya>:r")] (id_conv RS subst) 1
|
|
95 |
THEN (REPEAT (atac 1)));
|
|
96 |
by (fast_tac (ZF_cs addIs [id_conv RS ssubst]) 1);
|
|
97 |
val rvimage_id = result();
|
|
98 |
|
|
99 |
(* used only in Hartog.ML *)
|
|
100 |
goal Cardinal.thy
|
|
101 |
"!!A r. well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)";
|
|
102 |
by (res_inst_tac [("P","%a. ordertype(A,a)=ordertype(A,r)")]
|
|
103 |
(rvimage_id RS subst) 1);
|
|
104 |
by (eresolve_tac [id_bij RS bij_ordertype_vimage] 1);
|
|
105 |
val ordertype_Int = result();
|
|
106 |
|
|
107 |
(* used only in AC16_lemmas.ML *)
|
|
108 |
goalw CardinalArith.thy [InfCard_def]
|
|
109 |
"!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)";
|
|
110 |
by (asm_simp_tac (ZF_ss addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1);
|
|
111 |
val Inf_Card_is_InfCard = result();
|