author | wenzelm |
Wed, 08 Mar 2000 17:51:29 +0100 | |
changeset 8367 | 2d77b5a723f1 |
parent 5137 | 60205b0de9b9 |
child 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 |
||
8 |
open Hartog; |
|
9 |
||
5137 | 10 |
Goal "ALL a. Ord(a) --> a:X ==> P"; |
1123 | 11 |
by (res_inst_tac [("X1","{y:X. Ord(y)}")] (ON_class RS revcut_rl) 1); |
2469 | 12 |
by (Fast_tac 1); |
1299 | 13 |
qed "Ords_in_set"; |
1123 | 14 |
|
5137 | 15 |
Goalw [lepoll_def] "[| Ord(a); a lepoll X |] ==> \ |
1461 | 16 |
\ EX Y. Y<=X & (EX R. well_ord(Y,R) & ordertype(Y,R)=a)"; |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
17 |
by (etac exE 1); |
1123 | 18 |
by (REPEAT (resolve_tac [exI, conjI] 1)); |
19 |
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
|
20 |
by (rtac exI 1); |
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
21 |
by (rtac conjI 1); |
1123 | 22 |
by (eresolve_tac [well_ord_Memrel RSN (2, subset_refl RSN (2, |
1461 | 23 |
restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1 |
24 |
THEN (assume_tac 1)); |
|
1123 | 25 |
by (resolve_tac [subset_refl RSN (2, restrict_bij RS bij_converse_bij) RS |
1461 | 26 |
(well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS |
27 |
(ordertype_Memrel RSN (2, trans))] 1 |
|
28 |
THEN (REPEAT (assume_tac 1))); |
|
1299 | 29 |
qed "Ord_lepoll_imp_ex_well_ord"; |
1123 | 30 |
|
5137 | 31 |
Goal "[| Ord(a); a lepoll X |] ==> \ |
1461 | 32 |
\ EX Y. Y<=X & (EX R. R<=X*X & ordertype(Y,R)=a)"; |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
33 |
by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1)); |
3731 | 34 |
by Safe_tac; |
1123 | 35 |
by (REPEAT (ares_tac [exI, conjI] 1)); |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
36 |
by (etac ordertype_Int 2); |
2469 | 37 |
by (Fast_tac 1); |
1299 | 38 |
qed "Ord_lepoll_imp_eq_ordertype"; |
1123 | 39 |
|
5137 | 40 |
Goal "ALL a. Ord(a) --> a lepoll X ==> \ |
1461 | 41 |
\ ALL a. Ord(a) --> \ |
42 |
\ a:{a. Z:Pow(X)*Pow(X*X), EX Y R. Z=<Y,R> & ordertype(Y,R)=a}"; |
|
1123 | 43 |
by (REPEAT (resolve_tac [allI,impI] 1)); |
44 |
by (REPEAT (eresolve_tac [allE, impE] 1)); |
|
1196 | 45 |
by (assume_tac 1); |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
46 |
by (dtac Ord_lepoll_imp_eq_ordertype 1 THEN (assume_tac 1)); |
4091 | 47 |
by (fast_tac (claset() addSIs [ReplaceI] addEs [sym]) 1); |
1299 | 48 |
qed "Ords_lepoll_set_lemma"; |
1123 | 49 |
|
5137 | 50 |
Goal "ALL a. Ord(a) --> a lepoll X ==> P"; |
1123 | 51 |
by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1); |
1299 | 52 |
qed "Ords_lepoll_set"; |
1123 | 53 |
|
5068 | 54 |
Goal "EX a. Ord(a) & ~a lepoll X"; |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
55 |
by (rtac swap 1); |
2469 | 56 |
by (Fast_tac 1); |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
57 |
by (rtac Ords_lepoll_set 1); |
2469 | 58 |
by (Fast_tac 1); |
1299 | 59 |
qed "ex_Ord_not_lepoll"; |
1123 | 60 |
|
5068 | 61 |
Goalw [Hartog_def] "~ Hartog(A) lepoll A"; |
1123 | 62 |
by (resolve_tac [ex_Ord_not_lepoll RS exE] 1); |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
63 |
by (rtac LeastI 1); |
2469 | 64 |
by (REPEAT (Fast_tac 1)); |
1299 | 65 |
qed "HartogI"; |
1123 | 66 |
|
67 |
val HartogE = HartogI RS notE; |
|
68 |
||
5068 | 69 |
Goalw [Hartog_def] "Ord(Hartog(A))"; |
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1200
diff
changeset
|
70 |
by (rtac Ord_Least 1); |
1299 | 71 |
qed "Ord_Hartog"; |
1123 | 72 |
|
5137 | 73 |
Goalw [Hartog_def] "[| i < Hartog(A); ~ i lepoll A |] ==> P"; |
4091 | 74 |
by (fast_tac (claset() addEs [less_LeastE]) 1); |
1299 | 75 |
qed "less_HartogE1"; |
1123 | 76 |
|
5137 | 77 |
Goal "[| i < Hartog(A); i eqpoll Hartog(A) |] ==> P"; |
4091 | 78 |
by (fast_tac (claset() addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll |
1461 | 79 |
RS lepoll_trans RS HartogE]) 1); |
1299 | 80 |
qed "less_HartogE"; |
1123 | 81 |
|
5068 | 82 |
Goal "Card(Hartog(A))"; |
4091 | 83 |
by (fast_tac (claset() addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1); |
1299 | 84 |
qed "Card_Hartog"; |