| author | paulson |
| Tue, 26 Jun 2001 16:54:39 +0200 | |
| changeset 11380 | e76366922751 |
| parent 11317 | 7f9e4c389318 |
| permissions | -rw-r--r-- |
| 1461 | 1 |
(* Title: ZF/AC/Hartog.ML |
| 1123 | 2 |
ID: $Id$ |
| 1461 | 3 |
Author: Krzysztof Grabczewski |
| 1123 | 4 |
|
5 |
Some proofs on the Hartogs function. |
|
6 |
*) |
|
7 |
||
| 11317 | 8 |
Goal "\\<forall>a. Ord(a) --> a \\<in> X ==> P"; |
9 |
by (res_inst_tac [("X1","{y \\<in> X. Ord(y)}")] (ON_class RS revcut_rl) 1);
|
|
| 2469 | 10 |
by (Fast_tac 1); |
| 1299 | 11 |
qed "Ords_in_set"; |
| 1123 | 12 |
|
| 5137 | 13 |
Goalw [lepoll_def] "[| Ord(a); a lepoll X |] ==> \ |
| 11317 | 14 |
\ \\<exists>Y. Y \\<subseteq> X & (\\<exists>R. well_ord(Y,R) & ordertype(Y,R)=a)"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
15 |
by (etac exE 1); |
| 1123 | 16 |
by (REPEAT (resolve_tac [exI, conjI] 1)); |
17 |
by (eresolve_tac [inj_is_fun RS fun_is_rel RS image_subset] 1); |
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
18 |
by (rtac exI 1); |
|
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
19 |
by (rtac conjI 1); |
| 1123 | 20 |
by (eresolve_tac [well_ord_Memrel RSN (2, subset_refl RSN (2, |
| 1461 | 21 |
restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1 |
22 |
THEN (assume_tac 1)); |
|
| 1123 | 23 |
by (resolve_tac [subset_refl RSN (2, restrict_bij RS bij_converse_bij) RS |
| 1461 | 24 |
(well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS |
25 |
(ordertype_Memrel RSN (2, trans))] 1 |
|
26 |
THEN (REPEAT (assume_tac 1))); |
|
| 1299 | 27 |
qed "Ord_lepoll_imp_ex_well_ord"; |
| 1123 | 28 |
|
| 5137 | 29 |
Goal "[| Ord(a); a lepoll X |] ==> \ |
| 11317 | 30 |
\ \\<exists>Y. Y \\<subseteq> X & (\\<exists>R. R \\<subseteq> X*X & ordertype(Y,R)=a)"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
31 |
by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1)); |
| 3731 | 32 |
by Safe_tac; |
| 1123 | 33 |
by (REPEAT (ares_tac [exI, conjI] 1)); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
34 |
by (etac ordertype_Int 2); |
| 2469 | 35 |
by (Fast_tac 1); |
| 1299 | 36 |
qed "Ord_lepoll_imp_eq_ordertype"; |
| 1123 | 37 |
|
| 11317 | 38 |
Goal "\\<forall>a. Ord(a) --> a lepoll X ==> \ |
39 |
\ \\<forall>a. Ord(a) --> \ |
|
40 |
\ a:{a. Z \\<in> Pow(X)*Pow(X*X), \\<exists>Y R. Z=<Y,R> & ordertype(Y,R)=a}";
|
|
| 1123 | 41 |
by (REPEAT (resolve_tac [allI,impI] 1)); |
42 |
by (REPEAT (eresolve_tac [allE, impE] 1)); |
|
| 1196 | 43 |
by (assume_tac 1); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
44 |
by (dtac Ord_lepoll_imp_eq_ordertype 1 THEN (assume_tac 1)); |
| 4091 | 45 |
by (fast_tac (claset() addSIs [ReplaceI] addEs [sym]) 1); |
| 1299 | 46 |
qed "Ords_lepoll_set_lemma"; |
| 1123 | 47 |
|
| 11317 | 48 |
Goal "\\<forall>a. Ord(a) --> a lepoll X ==> P"; |
| 1123 | 49 |
by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1); |
| 1299 | 50 |
qed "Ords_lepoll_set"; |
| 1123 | 51 |
|
| 11317 | 52 |
Goal "\\<exists>a. Ord(a) & ~a lepoll X"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
53 |
by (rtac swap 1); |
| 2469 | 54 |
by (Fast_tac 1); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
55 |
by (rtac Ords_lepoll_set 1); |
| 2469 | 56 |
by (Fast_tac 1); |
| 1299 | 57 |
qed "ex_Ord_not_lepoll"; |
| 1123 | 58 |
|
| 5068 | 59 |
Goalw [Hartog_def] "~ Hartog(A) lepoll A"; |
| 1123 | 60 |
by (resolve_tac [ex_Ord_not_lepoll RS exE] 1); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
61 |
by (rtac LeastI 1); |
| 2469 | 62 |
by (REPEAT (Fast_tac 1)); |
| 1299 | 63 |
qed "HartogI"; |
| 1123 | 64 |
|
65 |
val HartogE = HartogI RS notE; |
|
66 |
||
| 5068 | 67 |
Goalw [Hartog_def] "Ord(Hartog(A))"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
68 |
by (rtac Ord_Least 1); |
| 1299 | 69 |
qed "Ord_Hartog"; |
| 1123 | 70 |
|
| 5137 | 71 |
Goalw [Hartog_def] "[| i < Hartog(A); ~ i lepoll A |] ==> P"; |
| 4091 | 72 |
by (fast_tac (claset() addEs [less_LeastE]) 1); |
| 1299 | 73 |
qed "less_HartogE1"; |
| 1123 | 74 |
|
| 5137 | 75 |
Goal "[| i < Hartog(A); i eqpoll Hartog(A) |] ==> P"; |
| 4091 | 76 |
by (fast_tac (claset() addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll |
| 1461 | 77 |
RS lepoll_trans RS HartogE]) 1); |
| 1299 | 78 |
qed "less_HartogE"; |
| 1123 | 79 |
|
| 5068 | 80 |
Goal "Card(Hartog(A))"; |
| 4091 | 81 |
by (fast_tac (claset() addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1); |
| 1299 | 82 |
qed "Card_Hartog"; |