src/ZF/AC/AC_Equiv.thy
author wenzelm
Tue, 16 Jan 2018 09:30:00 +0100
changeset 67443 3abf6a722518
parent 65449 c82e63b11b8b
child 76213 e44d86131648
permissions -rw-r--r--
standardized towards new-style formal comments: isabelle update_comments;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/AC/AC_Equiv.thy
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     2
    Author:     Krzysztof Grabczewski
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     3
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     4
Axioms AC1 -- AC19 come from "Equivalents of the Axiom of Choice, II"
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     5
by H. Rubin and J.E. Rubin, 1985.
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     6
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     7
Axiom AC0 comes from "Axiomatic Set Theory" by P. Suppes, 1972.
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     8
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
     9
Some Isabelle proofs of equivalences of these axioms are formalizations of
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents: 1071
diff changeset
    10
proofs presented by the Rubins.  The others are based on the Rubins' proofs,
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents: 1071
diff changeset
    11
but slightly changed.
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    12
*)
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    13
27678
85ea2be46c71 dropped locale (open)
haftmann
parents: 24893
diff changeset
    14
theory AC_Equiv
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 61980
diff changeset
    15
imports ZF
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 61980
diff changeset
    16
begin (*obviously not ZFC*)
8123
a71686059be0 a bit of tidying
paulson
parents: 2469
diff changeset
    17
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    18
(* Well Ordering Theorems *)
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    19
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    20
definition  
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    21
    "WO1 == \<forall>A. \<exists>R. well_ord(A,R)"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    22
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    23
definition  
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    24
    "WO2 == \<forall>A. \<exists>a. Ord(a) & A\<approx>a"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    25
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    26
definition  
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    27
    "WO3 == \<forall>A. \<exists>a. Ord(a) & (\<exists>b. b \<subseteq> a & A\<approx>b)"
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents: 1071
diff changeset
    28
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    29
definition  
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    30
    "WO4(m) == \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &   
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
    31
                         (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    32
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    33
definition  
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    34
    "WO5 == \<exists>m \<in> nat. 1\<le>m & WO4(m)"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    35
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    36
definition  
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    37
    "WO6 == \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
    38
                               & (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    39
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    40
definition  
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    41
    "WO7 == \<forall>A. Finite(A) \<longleftrightarrow> (\<forall>R. well_ord(A,R) \<longrightarrow> well_ord(A,converse(R)))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    42
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    43
definition  
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    44
    "WO8 == \<forall>A. (\<exists>f. f \<in> (\<Prod>X \<in> A. X)) \<longrightarrow> (\<exists>R. well_ord(A,R))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    45
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    46
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    47
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    48
(* Auxiliary concepts needed below *)
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    49
  pairwise_disjoint :: "i => o"  where
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    50
    "pairwise_disjoint(A) == \<forall>A1 \<in> A. \<forall>A2 \<in> A. A1 \<inter> A2 \<noteq> 0 \<longrightarrow> A1=A2"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    51
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    52
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    53
  sets_of_size_between :: "[i, i, i] => o"  where
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    54
    "sets_of_size_between(A,m,n) == \<forall>B \<in> A. m \<lesssim> B & B \<lesssim> n"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    55
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    56
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    57
(* Axioms of Choice *)  
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    58
definition
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    59
    "AC0 == \<forall>A. \<exists>f. f \<in> (\<Prod>X \<in> Pow(A)-{0}. X)"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    60
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    61
definition
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    62
    "AC1 == \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. f \<in> (\<Prod>X \<in> A. X))"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    63
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    64
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    65
    "AC2 == \<forall>A. 0\<notin>A & pairwise_disjoint(A)   
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    66
                   \<longrightarrow> (\<exists>C. \<forall>B \<in> A. \<exists>y. B \<inter> C = {y})"
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    67
definition
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    68
    "AC3 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Prod>x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    69
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    70
definition
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    71
    "AC4 == \<forall>R A B. (R \<subseteq> A*B \<longrightarrow> (\<exists>f. f \<in> (\<Prod>x \<in> domain(R). R``{x})))"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    72
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    73
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    74
    "AC5 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g \<in> range(f)->A. \<forall>x \<in> domain(g). f`(g`x) = x"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    75
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    76
definition
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    77
    "AC6 == \<forall>A. 0\<notin>A \<longrightarrow> (\<Prod>B \<in> A. B)\<noteq>0"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    78
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    79
definition
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
    80
    "AC7 == \<forall>A. 0\<notin>A & (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow> (\<Prod>B \<in> A. B) \<noteq> 0"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    81
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    82
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    83
    "AC8 == \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1\<approx>B2)   
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    84
                   \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    85
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    86
definition
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    87
    "AC9 == \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow>   
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
    88
                   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    89
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    90
definition
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    91
    "AC10(n) ==  \<forall>A. (\<forall>B \<in> A. ~Finite(B)) \<longrightarrow>   
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
    92
                   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    93
                   sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
    94
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    95
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    96
    "AC11 == \<exists>n \<in> nat. 1\<le>n & AC10(n)"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
    97
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    98
definition
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
    99
    "AC12 == \<forall>A. (\<forall>B \<in> A. ~Finite(B)) \<longrightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
   100
                 (\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   101
                      sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B)))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   102
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   103
definition
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   104
    "AC13(m) == \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   105
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   106
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   107
    "AC14 == \<exists>m \<in> nat. 1\<le>m & AC13(m)"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   108
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   109
definition
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   110
    "AC15 == \<forall>A. 0\<notin>A \<longrightarrow> 
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   111
                 (\<exists>m \<in> nat. 1\<le>m & (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   112
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   113
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   114
    "AC16(n, k)  == 
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   115
       \<forall>A. ~Finite(A) \<longrightarrow>   
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
   116
           (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &   
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
   117
           (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   118
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   119
definition
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   120
    "AC17 == \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.   
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27678
diff changeset
   121
                   \<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   122
13416
5851987ab2e8 AC18: meta-level predicate via locale;
wenzelm
parents: 12820
diff changeset
   123
locale AC18 =
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   124
  assumes AC18: "A\<noteq>0 & (\<forall>a \<in> A. B(a) \<noteq> 0) \<longrightarrow>
13416
5851987ab2e8 AC18: meta-level predicate via locale;
wenzelm
parents: 12820
diff changeset
   125
    ((\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a,b)) =   
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
   126
      (\<Union>f \<in> \<Prod>a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a)))"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   127
  \<comment> \<open>AC18 cannot be expressed within the object-logic\<close>
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   128
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
   129
definition
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   130
    "AC19 == \<forall>A. A\<noteq>0 & 0\<notin>A \<longrightarrow> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =   
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
   131
                   (\<Union>f \<in> (\<Prod>B \<in> A. B). \<Inter>a \<in> A. f`a))"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   132
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   133
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   134
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   135
(* ********************************************************************** *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   136
(*                    Theorems concerning ordinals                        *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   137
(* ********************************************************************** *)
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   138
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   139
(* lemma for ordertype_Int *)
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   140
lemma rvimage_id: "rvimage(A,id(A),r) = r \<inter> A*A"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   141
apply (unfold rvimage_def)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   142
apply (rule equalityI, safe)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   143
apply (drule_tac P = "%a. <id (A) `xb,a>:r" in id_conv [THEN subst],
12820
02e2ff3e4d37 lexical tidying
paulson
parents: 12814
diff changeset
   144
       assumption)
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   145
apply (drule_tac P = "%a. <a,ya>:r" in id_conv [THEN subst], (assumption+))
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   146
apply (fast intro: id_conv [THEN ssubst])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   147
done
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   148
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   149
(* used only in Hartog.ML *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   150
lemma ordertype_Int:
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 32960
diff changeset
   151
     "well_ord(A,r) ==> ordertype(A, r \<inter> A*A) = ordertype(A,r)"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   152
apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   153
apply (erule id_bij [THEN bij_ordertype_vimage])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   154
done
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   155
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   156
lemma lam_sing_bij: "(\<lambda>x \<in> A. {x}) \<in> bij(A, {{x}. x \<in> A})"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   157
apply (rule_tac d = "%z. THE x. z={x}" in lam_bijective)
12814
paulson
parents: 12776
diff changeset
   158
apply (auto simp add: the_equality)
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   159
done
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   160
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   161
lemma inj_strengthen_type: 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   162
     "[| f \<in> inj(A, B);  !!a. a \<in> A ==> f`a \<in> C |] ==> f \<in> inj(A,C)"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   163
by (unfold inj_def, blast intro: Pi_type) 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   164
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   165
(* ********************************************************************** *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   166
(* Another elimination rule for \<exists>!                                       *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   167
(* ********************************************************************** *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   168
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   169
lemma ex1_two_eq: "[| \<exists>! x. P(x); P(x); P(y) |] ==> x=y"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   170
by blast
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   171
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   172
(* ********************************************************************** *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   173
(* Lemmas used in the proofs like WO? ==> AC?                             *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   174
(* ********************************************************************** *)
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   175
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   176
lemma first_in_B:
47101
ded5cc757bc9 proof tidying
paulson
parents: 46822
diff changeset
   177
     "[| well_ord(\<Union>(A),r); 0 \<notin> A; B \<in> A |] ==> (THE b. first(b,B,r)) \<in> B"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   178
by (blast dest!: well_ord_imp_ex1_first
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   179
                    [THEN theI, THEN first_def [THEN def_imp_iff, THEN iffD1]])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   180
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
   181
lemma ex_choice_fun: "[| well_ord(\<Union>(A), R); 0 \<notin> A |] ==> \<exists>f. f \<in> (\<Prod>X \<in> A. X)"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   182
by (fast elim!: first_in_B intro!: lam_type)
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   183
61980
6b780867d426 clarified syntax;
wenzelm
parents: 61798
diff changeset
   184
lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \<exists>f. f \<in> (\<Prod>X \<in> Pow(A)-{0}. X)"
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   185
by (fast elim!: well_ord_subset [THEN ex_choice_fun])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   186
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   187
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   188
(* ********************************************************************** *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   189
(* Lemmas needed to state when a finite relation is a function.           *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   190
(*     The criteria are cardinalities of the relation and its domain.     *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   191
(*     Used in WO6WO1.ML                                                  *)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   192
(* ********************************************************************** *)
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   193
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   194
(*Using AC we could trivially prove, for all u, domain(u) \<lesssim> u*)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   195
lemma lepoll_m_imp_domain_lepoll_m: 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   196
     "[| m \<in> nat; u \<lesssim> m |] ==> domain(u) \<lesssim> m"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   197
apply (unfold lepoll_def)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   198
apply (erule exE)
61394
6142b282b164 tuned syntax -- more symbols;
wenzelm
parents: 47101
diff changeset
   199
apply (rule_tac x = "\<lambda>x \<in> domain(u). \<mu> i. \<exists>y. <x,y> \<in> u & f`<x,y> = i" 
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   200
       in exI)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   201
apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   202
apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   203
                           inj_is_fun [THEN apply_type])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   204
apply (erule domainE)
12820
02e2ff3e4d37 lexical tidying
paulson
parents: 12814
diff changeset
   205
apply (frule inj_is_fun [THEN apply_type], assumption)
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   206
apply (rule LeastI2)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   207
apply (auto elim!: nat_into_Ord [THEN Ord_in_Ord])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   208
done
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   209
12776
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   210
lemma rel_domain_ex1: 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   211
    "[| succ(m) \<lesssim> domain(r); r \<lesssim> succ(m); m \<in> nat |] ==> function(r)"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   212
apply (unfold function_def, safe)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   213
apply (rule ccontr) 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   214
apply (fast elim!: lepoll_trans [THEN succ_lepoll_natE] 
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   215
                   lepoll_m_imp_domain_lepoll_m [OF _ Diff_sing_lepoll]
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   216
            elim: domain_Diff_eq [OF _ not_sym, THEN subst])
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   217
done
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   218
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   219
lemma rel_is_fun:
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   220
     "[| succ(m) \<lesssim> domain(r);  r \<lesssim> succ(m);  m \<in> nat;   
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   221
         r \<subseteq> A*B; A=domain(r) |] ==> r \<in> A->B"
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   222
by (simp add: Pi_iff rel_domain_ex1)
249600a63ba9 Isar version of AC
paulson
parents: 12536
diff changeset
   223
991
547931cbbf08 New example of AC Equivalences by Krzysztof Grabczewski
lcp
parents:
diff changeset
   224
end