| author | haftmann | 
| Mon, 17 Dec 2007 17:57:50 +0100 | |
| changeset 25667 | a089038c1893 | 
| parent 24893 | b8ef7afe3a6b | 
| child 27678 | 85ea2be46c71 | 
| permissions | -rw-r--r-- | 
| 1478 | 1  | 
(* Title: ZF/AC/AC18_AC19.thy  | 
| 1123 | 2  | 
ID: $Id$  | 
| 1478 | 3  | 
Author: Krzysztof Grabczewski  | 
| 1123 | 4  | 
|
| 12776 | 5  | 
The proof of AC1 ==> AC18 ==> AC19 ==> AC1  | 
| 1123 | 6  | 
*)  | 
7  | 
||
| 16417 | 8  | 
theory AC18_AC19 imports AC_Equiv begin  | 
| 12776 | 9  | 
|
| 24893 | 10  | 
definition  | 
11  | 
uu :: "i => i" where  | 
|
| 12776 | 12  | 
    "uu(a) == {c Un {0}. c \<in> a}"
 | 
13  | 
||
14  | 
||
15  | 
(* ********************************************************************** *)  | 
|
16  | 
(* AC1 ==> AC18 *)  | 
|
17  | 
(* ********************************************************************** *)  | 
|
18  | 
||
19  | 
lemma PROD_subsets:  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13421 
diff
changeset
 | 
20  | 
     "[| f \<in> (\<Pi> b \<in> {P(a). a \<in> A}. b);  \<forall>a \<in> A. P(a)<=Q(a) |]   
 | 
| 
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13421 
diff
changeset
 | 
21  | 
==> (\<lambda>a \<in> A. f`P(a)) \<in> (\<Pi> a \<in> A. Q(a))"  | 
| 12776 | 22  | 
by (rule lam_type, drule apply_type, auto)  | 
23  | 
||
24  | 
lemma lemma_AC18:  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13421 
diff
changeset
 | 
25  | 
"[| \<forall>A. 0 \<notin> A --> (\<exists>f. f \<in> (\<Pi> X \<in> A. X)); A \<noteq> 0 |]  | 
| 12776 | 26  | 
==> (\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a, b)) \<subseteq>  | 
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13421 
diff
changeset
 | 
27  | 
(\<Union>f \<in> \<Pi> a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a))"  | 
| 12776 | 28  | 
apply (rule subsetI)  | 
29  | 
apply (erule_tac x = "{{b \<in> B (a) . x \<in> X (a,b) }. a \<in> A}" in allE)
 | 
|
30  | 
apply (erule impE, fast)  | 
|
31  | 
apply (erule exE)  | 
|
32  | 
apply (rule UN_I)  | 
|
33  | 
apply (fast elim!: PROD_subsets)  | 
|
34  | 
apply (simp, fast elim!: not_emptyE dest: apply_type [OF _ RepFunI])  | 
|
35  | 
done  | 
|
36  | 
||
| 13416 | 37  | 
lemma AC1_AC18: "AC1 ==> PROP AC18"  | 
38  | 
apply (unfold AC1_def)  | 
|
| 13421 | 39  | 
apply (rule AC18.intro)  | 
| 12776 | 40  | 
apply (fast elim!: lemma_AC18 apply_type intro!: equalityI INT_I UN_I)  | 
41  | 
done  | 
|
42  | 
||
43  | 
(* ********************************************************************** *)  | 
|
44  | 
(* AC18 ==> AC19 *)  | 
|
45  | 
(* ********************************************************************** *)  | 
|
46  | 
||
| 13416 | 47  | 
theorem (in AC18) AC19  | 
48  | 
apply (unfold AC19_def)  | 
|
49  | 
apply (intro allI impI)  | 
|
50  | 
apply (rule AC18 [of _ "%x. x", THEN mp], blast)  | 
|
51  | 
done  | 
|
| 1123 | 52  | 
|
| 12776 | 53  | 
(* ********************************************************************** *)  | 
54  | 
(* AC19 ==> AC1 *)  | 
|
55  | 
(* ********************************************************************** *)  | 
|
56  | 
||
57  | 
lemma RepRep_conj:  | 
|
58  | 
        "[| A \<noteq> 0; 0 \<notin> A |] ==> {uu(a). a \<in> A} \<noteq> 0 & 0 \<notin> {uu(a). a \<in> A}"
 | 
|
59  | 
apply (unfold uu_def, auto)  | 
|
60  | 
apply (blast dest!: sym [THEN RepFun_eq_0_iff [THEN iffD1]])  | 
|
61  | 
done  | 
|
62  | 
||
63  | 
lemma lemma1_1: "[|c \<in> a; x = c Un {0}; x \<notin> a |] ==> x - {0} \<in> a"
 | 
|
64  | 
apply clarify  | 
|
| 12820 | 65  | 
apply (rule subst_elem, assumption)  | 
| 12776 | 66  | 
apply (fast elim: notE subst_elem)  | 
67  | 
done  | 
|
68  | 
||
69  | 
lemma lemma1_2:  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13421 
diff
changeset
 | 
70  | 
        "[| f`(uu(a)) \<notin> a; f \<in> (\<Pi> B \<in> {uu(a). a \<in> A}. B); a \<in> A |]   
 | 
| 12776 | 71  | 
                ==> f`(uu(a))-{0} \<in> a"
 | 
72  | 
apply (unfold uu_def, fast elim!: lemma1_1 dest!: apply_type)  | 
|
73  | 
done  | 
|
| 1123 | 74  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13421 
diff
changeset
 | 
75  | 
lemma lemma1: "\<exists>f. f \<in> (\<Pi> B \<in> {uu(a). a \<in> A}. B) ==> \<exists>f. f \<in> (\<Pi> B \<in> A. B)"
 | 
| 12776 | 76  | 
apply (erule exE)  | 
77  | 
apply (rule_tac x = "\<lambda>a\<in>A. if (f` (uu(a)) \<in> a, f` (uu(a)), f` (uu(a))-{0})" 
 | 
|
78  | 
in exI)  | 
|
79  | 
apply (rule lam_type)  | 
|
80  | 
apply (simp add: lemma1_2)  | 
|
81  | 
done  | 
|
82  | 
||
83  | 
lemma lemma2_1: "a\<noteq>0 ==> 0 \<in> (\<Union>b \<in> uu(a). b)"  | 
|
84  | 
by (unfold uu_def, auto)  | 
|
85  | 
||
86  | 
lemma lemma2: "[| A\<noteq>0; 0\<notin>A |] ==> (\<Inter>x \<in> {uu(a). a \<in> A}. \<Union>b \<in> x. b) \<noteq> 0"
 | 
|
87  | 
apply (erule not_emptyE)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12820 
diff
changeset
 | 
88  | 
apply (rule_tac a = 0 in not_emptyI)  | 
| 12776 | 89  | 
apply (fast intro!: lemma2_1)  | 
90  | 
done  | 
|
91  | 
||
92  | 
lemma AC19_AC1: "AC19 ==> AC1"  | 
|
93  | 
apply (unfold AC19_def AC1_def, clarify)  | 
|
94  | 
apply (case_tac "A=0", force)  | 
|
95  | 
apply (erule_tac x = "{uu (a) . a \<in> A}" in allE)
 | 
|
96  | 
apply (erule impE)  | 
|
| 12820 | 97  | 
apply (erule RepRep_conj, assumption)  | 
| 12776 | 98  | 
apply (rule lemma1)  | 
| 12820 | 99  | 
apply (drule lemma2, assumption, auto)  | 
| 12776 | 100  | 
done  | 
| 1123 | 101  | 
|
| 1203 | 102  | 
end  |