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