src/ZF/AC/Hartog.ML
author wenzelm
Wed, 09 Jul 1997 17:00:34 +0200
changeset 3511 da4dd8b7ced4
parent 2469 b50b8c0eec01
child 3731 71366483323b
permissions -rw-r--r--
removed obsolete init_pps and init_thy_reader;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
     1
(*  Title:      ZF/AC/Hartog.ML
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
     3
    Author:     Krzysztof Grabczewski
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     4
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     5
  Some proofs on the Hartogs function.
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     6
*)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     7
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     8
open Hartog;
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     9
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    10
goal thy "!!X. ALL a. Ord(a) --> a:X ==> P";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    11
by (res_inst_tac [("X1","{y:X. Ord(y)}")] (ON_class RS revcut_rl) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    12
by (Fast_tac 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    13
qed "Ords_in_set";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    14
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    15
goalw thy [lepoll_def] "!!X. [| Ord(a); a lepoll X |] ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    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
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    18
by (REPEAT (resolve_tac [exI, conjI] 1));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    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
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    22
by (eresolve_tac [well_ord_Memrel RSN (2, subset_refl RSN (2, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    23
        restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    24
        THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    25
by (resolve_tac [subset_refl RSN (2, restrict_bij RS bij_converse_bij) RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    26
        (well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    27
        (ordertype_Memrel RSN (2, trans))] 1
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    28
        THEN (REPEAT (assume_tac 1)));
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    29
qed "Ord_lepoll_imp_ex_well_ord";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    30
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    31
goal thy "!!X. [| Ord(a); a lepoll X |] ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    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));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    34
by (Step_tac 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    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
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    37
by (Fast_tac 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    38
qed "Ord_lepoll_imp_eq_ordertype";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    39
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    40
goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    41
\       ALL a. Ord(a) -->  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    42
\       a:{a. Z:Pow(X)*Pow(X*X), EX Y R. Z=<Y,R> & ordertype(Y,R)=a}";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    43
by (REPEAT (resolve_tac [allI,impI] 1));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    44
by (REPEAT (eresolve_tac [allE, impE] 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    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));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    47
by (fast_tac (!claset addSIs [ReplaceI] addEs [sym]) 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    48
qed "Ords_lepoll_set_lemma";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    49
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    50
goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==> P";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    51
by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    52
qed "Ords_lepoll_set";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    53
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    54
goal thy "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
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    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
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    58
by (Fast_tac 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    59
qed "ex_Ord_not_lepoll";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    60
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    61
goalw thy [Hartog_def] "~ Hartog(A) lepoll A";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    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
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    64
by (REPEAT (Fast_tac 1));
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    65
qed "HartogI";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    66
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    67
val HartogE = HartogI RS notE;
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    68
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    69
goalw thy [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
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    71
qed "Ord_Hartog";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    72
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    73
goalw thy [Hartog_def] "!!i. [| i < Hartog(A); ~ i lepoll A |] ==> P";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    74
by (fast_tac (!claset addEs [less_LeastE]) 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    75
qed "less_HartogE1";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    76
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    77
goal thy "!!i. [| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    78
by (fast_tac (!claset addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1299
diff changeset
    79
                RS lepoll_trans RS HartogE]) 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    80
qed "less_HartogE";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    81
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    82
goal thy "Card(Hartog(A))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    83
by (fast_tac (!claset addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1);
1299
e74f694ca1da added usage of qed
clasohm
parents: 1208
diff changeset
    84
qed "Card_Hartog";