src/ZF/AC.thy
author paulson
Fri, 10 May 2002 22:50:08 +0200
changeset 13134 bf37a3049251
parent 6053 8a1059aa01f0
child 13149 773657d466cb
permissions -rw-r--r--
converted the AC branch to Isar
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 484
diff changeset
     1
(*  Title:      ZF/AC.thy
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 484
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     5
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     6
The Axiom of Choice
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     7
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     8
This definition comes from Halmos (1960), page 59.
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
     9
*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    10
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    11
theory AC = Main:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    12
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    13
axioms AC: "[| a: A;  !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    14
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    15
(*The same as AC, but no premise a \<in> A*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    16
lemma AC_Pi: "[| !!x. x \<in> A ==> (\<exists>y. y \<in> B(x)) |] ==> \<exists>z. z \<in> Pi(A,B)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    17
apply (case_tac "A=0")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    18
apply (simp add: Pi_empty1, blast)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    19
(*The non-trivial case*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    20
apply (blast intro: AC)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    21
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    22
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    23
(*Using dtac, this has the advantage of DELETING the universal quantifier*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    24
lemma AC_ball_Pi: "\<forall>x \<in> A. \<exists>y. y \<in> B(x) ==> \<exists>y. y \<in> Pi(A,B)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    25
apply (rule AC_Pi)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    26
apply (erule bspec)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    27
apply assumption
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    28
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    29
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    30
lemma AC_Pi_Pow: "\<exists>f. f \<in> (\<Pi>X \<in> Pow(C)-{0}. X)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    31
apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    32
apply (erule_tac [2] exI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    33
apply blast
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    34
done
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 2469
diff changeset
    35
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    36
lemma AC_func:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    37
     "[| !!x. x \<in> A ==> (\<exists>y. y \<in> x) |] ==> \<exists>f \<in> A->Union(A). \<forall>x \<in> A. f`x \<in> x"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    38
apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    39
prefer 2 apply (blast dest: apply_type intro: Pi_type)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    40
apply (blast intro: elim:); 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    41
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    42
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    43
lemma non_empty_family: "[| 0 \<notin> A;  x \<in> A |] ==> \<exists>y. y \<in> x"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    44
apply (subgoal_tac "x \<noteq> 0")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    45
apply blast+
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    46
done
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 2469
diff changeset
    47
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    48
lemma AC_func0: "0 \<notin> A ==> \<exists>f \<in> A->Union(A). \<forall>x \<in> A. f`x \<in> x"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    49
apply (rule AC_func)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    50
apply (simp_all add: non_empty_family) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    51
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    52
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    53
lemma AC_func_Pow: "\<exists>f \<in> (Pow(C)-{0}) -> C. \<forall>x \<in> Pow(C)-{0}. f`x \<in> x"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    54
apply (rule AC_func0 [THEN bexE])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    55
apply (rule_tac [2] bexI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    56
prefer 2 apply (assumption)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    57
apply (erule_tac [2] fun_weaken_type)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    58
apply blast+
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    59
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    60
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    61
lemma AC_Pi0: "0 \<notin> A ==> \<exists>f. f \<in> (\<Pi>x \<in> A. x)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    62
apply (rule AC_Pi)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    63
apply (simp_all add: non_empty_family) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    64
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 6053
diff changeset
    65
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents:
diff changeset
    66
end