src/ZF/AC/AC17_AC1.ML
author lcp
Fri, 28 Jul 1995 11:48:55 +0200
changeset 1206 30df104ceb91
parent 1196 d43c1f7a53fe
child 1461 6bcb44e4d6e5
permissions -rw-r--r--
Ran expandshort and changed spelling of Grabczewski
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/AC/AC17_AC1.ML
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     2
    ID:         $Id$
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
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
The proof of AC1 ==> AC17
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 AC17_AC1;
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
(* *********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    11
(* more properties of HH						   *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    12
(* *********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    13
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    14
goal thy "!!f. [| x - (UN j:LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x}. \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    15
\	HH(lam X:Pow(x)-{0}. {f`X}, x, j)) = 0;  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    16
\	f : Pow(x)-{0} -> x |]  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    17
\	==> EX r. well_ord(x,r)";
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    18
by (rtac exI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    19
by (eresolve_tac [[bij_Least_HH_x RS bij_converse_bij RS bij_is_inj,
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    20
		Ord_Least RS well_ord_Memrel] MRS well_ord_rvimage] 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    21
by (assume_tac 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    22
val UN_eq_imp_well_ord = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    23
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    24
(* *********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    25
(* theorems closer to the proof						   *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    26
(* *********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    27
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    28
goalw thy AC_defs "!!Z. ~AC1 ==>  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    29
\		EX A. ALL f:Pow(A)-{0} -> A. EX u:Pow(A)-{0}. f`u ~: u";
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    30
by (etac swap 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    31
by (rtac allI 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    32
by (etac swap 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    33
by (res_inst_tac [("x","Union(A)")] exI 1);
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    34
by (rtac ballI 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    35
by (etac swap 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    36
by (rtac impI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    37
by (fast_tac (AC_cs addSIs [restrict_type]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    38
val not_AC1_imp_ex = result();
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 f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u;  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    41
\	EX f: Pow(x)-{0}->x. \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    42
\	x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}).  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    43
\	HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0 |] \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    44
\	==> P";
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    45
by (etac bexE 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    46
by (eresolve_tac [UN_eq_imp_well_ord RS exE] 1 THEN (assume_tac 1));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    47
by (eresolve_tac [ex_choice_fun_Pow RS exE] 1);
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    48
by (etac ballE 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    49
by (fast_tac (FOL_cs addEs [bexE, notE, apply_type]) 1);
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    50
by (etac notE 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    51
by (rtac Pi_type 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    52
by (resolve_tac [apply_type RSN (2, subsetD)] 1 THEN TRYALL assume_tac);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    53
by (fast_tac AC_cs 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    54
val lemma1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    55
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    56
goal thy "!!x. ~ (EX f: Pow(x)-{0}->x. x - F(f) = 0)  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    57
\	==> (lam f: Pow(x)-{0}->x. x - F(f))  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    58
\		: (Pow(x) -{0} -> x) -> Pow(x) - {0}";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    59
by (fast_tac (AC_cs addSIs [lam_type] addIs [equalityI]
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    60
		addSDs [Diff_eq_0_iff RS iffD1]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    61
val lemma2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    62
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    63
goal thy "!!f. [| f`Z : Z; Z:Pow(x)-{0} |] ==>  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    64
\	(lam X:Pow(x)-{0}. {f`X})`Z : Pow(Z)-{0}";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    65
by (asm_full_simp_tac AC_ss 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    66
by (fast_tac (AC_cs addSDs [equals0D]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    67
val lemma3 = result();
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
goal thy "!!z. EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    70
\	==> EX f:F. f`Q(f) : Q(f)";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    71
by (asm_full_simp_tac AC_ss 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    72
val lemma4 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    73
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    74
goalw thy [AC17_def] "!!Z. AC17 ==> AC1";
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    75
by (rtac classical 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    76
by (eresolve_tac [not_AC1_imp_ex RS exE] 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    77
by (excluded_middle_tac
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    78
	"EX f: Pow(x)-{0}->x. \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    79
\	x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}).  \
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    80
\	HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0" 1);
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    81
by (etac lemma1 2 THEN (assume_tac 2));
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    82
by (dtac lemma2 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    83
by (etac allE 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    84
by (dtac bspec 1 THEN (assume_tac 1));
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    85
by (dtac lemma4 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    86
by (etac bexE 1);
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    87
by (dtac apply_type 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    88
by (dresolve_tac [beta RS sym RSN (2, subst_elem)] 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    89
by (assume_tac 1);
1206
30df104ceb91 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    90
by (dtac lemma3 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    91
by (fast_tac (AC_cs addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem),
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    92
		f_subset_imp_HH_subset] addSEs [mem_irrefl]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    93
qed "AC17_AC1";