src/HOL/Library/Zorn.thy
author haftmann
Sun, 22 Jul 2012 09:56:34 +0200
changeset 49442 571cb1df0768
parent 47846 6bc213e90401
child 49765 a151db85a62b
permissions -rw-r--r--
library theories for debugging and parallel computing using code generation towards Isabelle/ML
wenzelm@32962
     1
(*  Title:      HOL/Library/Zorn.thy
wenzelm@32962
     2
    Author:     Jacques D. Fleuriot, Tobias Nipkow
wenzelm@32962
     3
wenzelm@32962
     4
Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF).
wenzelm@32962
     5
The well-ordering theorem.
wenzelm@14706
     6
*)
paulson@13551
     7
wenzelm@14706
     8
header {* Zorn's Lemma *}
paulson@13551
     9
nipkow@15131
    10
theory Zorn
haftmann@30663
    11
imports Order_Relation Main
nipkow@15131
    12
begin
paulson@13551
    13
nipkow@26272
    14
(* Define globally? In Set.thy? *)
haftmann@47846
    15
definition chain_subset :: "'a set set \<Rightarrow> bool" ("chain\<^bsub>\<subseteq>\<^esub>")
haftmann@47846
    16
where
haftmann@47846
    17
  "chain\<^bsub>\<subseteq>\<^esub> C \<equiv> \<forall>A\<in>C.\<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A"
nipkow@26272
    18
wenzelm@14706
    19
text{*
wenzelm@14706
    20
  The lemma and section numbers refer to an unpublished article
wenzelm@14706
    21
  \cite{Abrial-Laffitte}.
wenzelm@14706
    22
*}
paulson@13551
    23
haftmann@47846
    24
definition chain :: "'a set set \<Rightarrow> 'a set set set"
haftmann@47846
    25
where
haftmann@47846
    26
  "chain S = {F. F \<subseteq> S \<and> chain\<^bsub>\<subseteq>\<^esub> F}"
paulson@13551
    27
haftmann@47846
    28
definition super :: "'a set set \<Rightarrow> 'a set set \<Rightarrow> 'a set set set"
haftmann@47846
    29
where
haftmann@47846
    30
  "super S c = {d. d \<in> chain S \<and> c \<subset> d}"
paulson@13551
    31
haftmann@47846
    32
definition maxchain  ::  "'a set set \<Rightarrow> 'a set set set"
haftmann@47846
    33
where
haftmann@47846
    34
  "maxchain S = {c. c \<in> chain S \<and> super S c = {}}"
paulson@13551
    35
haftmann@47846
    36
definition succ :: "'a set set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
haftmann@47846
    37
where
haftmann@47846
    38
  "succ S c = (if c \<notin> chain S \<or> c \<in> maxchain S then c else SOME c'. c' \<in> super S c)"
paulson@13551
    39
haftmann@47846
    40
inductive_set TFin :: "'a set set \<Rightarrow> 'a set set set"
haftmann@47846
    41
for S :: "'a set set"
haftmann@47846
    42
where
haftmann@47846
    43
  succI:      "x \<in> TFin S \<Longrightarrow> succ S x \<in> TFin S"
haftmann@47846
    44
| Pow_UnionI: "Y \<in> Pow (TFin S) \<Longrightarrow> \<Union>Y \<in> TFin S"
paulson@13551
    45
paulson@13551
    46
paulson@13551
    47
subsection{*Mathematical Preamble*}
paulson@13551
    48
wenzelm@17200
    49
lemma Union_lemma0:
paulson@18143
    50
    "(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C) \<subseteq> A | B \<subseteq> Union(C)"
wenzelm@17200
    51
  by blast
paulson@13551
    52
paulson@13551
    53
paulson@13551
    54
text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*}
wenzelm@17200
    55
paulson@13551
    56
lemma Abrial_axiom1: "x \<subseteq> succ S x"
berghofe@26806
    57
  apply (auto simp add: succ_def super_def maxchain_def)
wenzelm@18585
    58
  apply (rule contrapos_np, assumption)
wenzelm@46895
    59
  apply (rule someI2)
wenzelm@46895
    60
  apply blast+
wenzelm@17200
    61
  done
paulson@13551
    62
paulson@13551
    63
lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI]
paulson@13551
    64
wenzelm@14706
    65
lemma TFin_induct:
wenzelm@46895
    66
  assumes H: "n \<in> TFin S" and
wenzelm@46895
    67
    I: "!!x. x \<in> TFin S ==> P x ==> P (succ S x)"
wenzelm@46895
    68
      "!!Y. Y \<subseteq> TFin S ==> Ball Y P ==> P (Union Y)"
wenzelm@46895
    69
  shows "P n"
wenzelm@46895
    70
  using H by induct (blast intro: I)+
paulson@13551
    71
paulson@13551
    72
lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y"
wenzelm@17200
    73
  apply (erule subset_trans)
wenzelm@17200
    74
  apply (rule Abrial_axiom1)
wenzelm@17200
    75
  done
paulson@13551
    76
paulson@13551
    77
text{*Lemma 1 of section 3.1*}
paulson@13551
    78
lemma TFin_linear_lemma1:
wenzelm@14706
    79
     "[| n \<in> TFin S;  m \<in> TFin S;
wenzelm@14706
    80
         \<forall>x \<in> TFin S. x \<subseteq> m --> x = m | succ S x \<subseteq> m
paulson@13551
    81
      |] ==> n \<subseteq> m | succ S m \<subseteq> n"
wenzelm@17200
    82
  apply (erule TFin_induct)
wenzelm@17200
    83
   apply (erule_tac [2] Union_lemma0)
wenzelm@17200
    84
  apply (blast del: subsetI intro: succ_trans)
wenzelm@17200
    85
  done
paulson@13551
    86
paulson@13551
    87
text{* Lemma 2 of section 3.2 *}
paulson@13551
    88
lemma TFin_linear_lemma2:
paulson@13551
    89
     "m \<in> TFin S ==> \<forall>n \<in> TFin S. n \<subseteq> m --> n=m | succ S n \<subseteq> m"
wenzelm@17200
    90
  apply (erule TFin_induct)
wenzelm@17200
    91
   apply (rule impI [THEN ballI])
wenzelm@17200
    92
   txt{*case split using @{text TFin_linear_lemma1}*}
wenzelm@17200
    93
   apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
wenzelm@17200
    94
     assumption+)
wenzelm@17200
    95
    apply (drule_tac x = n in bspec, assumption)
wenzelm@17200
    96
    apply (blast del: subsetI intro: succ_trans, blast)
wenzelm@17200
    97
  txt{*second induction step*}
wenzelm@17200
    98
  apply (rule impI [THEN ballI])
wenzelm@17200
    99
  apply (rule Union_lemma0 [THEN disjE])
wenzelm@17200
   100
    apply (rule_tac [3] disjI2)
wenzelm@17200
   101
    prefer 2 apply blast
wenzelm@17200
   102
   apply (rule ballI)
wenzelm@17200
   103
   apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
wenzelm@17200
   104
     assumption+, auto)
wenzelm@17200
   105
  apply (blast intro!: Abrial_axiom1 [THEN subsetD])
wenzelm@17200
   106
  done
paulson@13551
   107
paulson@13551
   108
text{*Re-ordering the premises of Lemma 2*}
paulson@13551
   109
lemma TFin_subsetD:
paulson@13551
   110
     "[| n \<subseteq> m;  m \<in> TFin S;  n \<in> TFin S |] ==> n=m | succ S n \<subseteq> m"
wenzelm@17200
   111
  by (rule TFin_linear_lemma2 [rule_format])
paulson@13551
   112
paulson@13551
   113
text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*}
paulson@13551
   114
lemma TFin_subset_linear: "[| m \<in> TFin S;  n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n"
wenzelm@17200
   115
  apply (rule disjE)
wenzelm@17200
   116
    apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2])
wenzelm@17200
   117
      apply (assumption+, erule disjI2)
wenzelm@17200
   118
  apply (blast del: subsetI
wenzelm@17200
   119
    intro: subsetI Abrial_axiom1 [THEN subset_trans])
wenzelm@17200
   120
  done
paulson@13551
   121
paulson@13551
   122
text{*Lemma 3 of section 3.3*}
paulson@13551
   123
lemma eq_succ_upper: "[| n \<in> TFin S;  m \<in> TFin S;  m = succ S m |] ==> n \<subseteq> m"
wenzelm@17200
   124
  apply (erule TFin_induct)
wenzelm@17200
   125
   apply (drule TFin_subsetD)
wenzelm@17200
   126
     apply (assumption+, force, blast)
wenzelm@17200
   127
  done
paulson@13551
   128
paulson@13551
   129
text{*Property 3.3 of section 3.3*}
paulson@13551
   130
lemma equal_succ_Union: "m \<in> TFin S ==> (m = succ S m) = (m = Union(TFin S))"
wenzelm@17200
   131
  apply (rule iffI)
wenzelm@17200
   132
   apply (rule Union_upper [THEN equalityI])
paulson@18143
   133
    apply assumption
paulson@18143
   134
   apply (rule eq_succ_upper [THEN Union_least], assumption+)
wenzelm@17200
   135
  apply (erule ssubst)
wenzelm@17200
   136
  apply (rule Abrial_axiom1 [THEN equalityI])
wenzelm@17200
   137
  apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI)
wenzelm@17200
   138
  done
paulson@13551
   139
paulson@13551
   140
subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*}
paulson@13551
   141
wenzelm@14706
   142
text{*NB: We assume the partial ordering is @{text "\<subseteq>"},
paulson@13551
   143
 the subset relation!*}
paulson@13551
   144
paulson@13551
   145
lemma empty_set_mem_chain: "({} :: 'a set set) \<in> chain S"
nipkow@26272
   146
by (unfold chain_def chain_subset_def) auto
paulson@13551
   147
paulson@13551
   148
lemma super_subset_chain: "super S c \<subseteq> chain S"
wenzelm@17200
   149
  by (unfold super_def) blast
paulson@13551
   150
paulson@13551
   151
lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S"
wenzelm@17200
   152
  by (unfold maxchain_def) blast
paulson@13551
   153
nipkow@26191
   154
lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> EX d. d \<in> super S c"
wenzelm@17200
   155
  by (unfold super_def maxchain_def) auto
paulson@13551
   156
paulson@18143
   157
lemma select_super:
paulson@18143
   158
     "c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c): super S c"
wenzelm@17200
   159
  apply (erule mem_super_Ex [THEN exE])
wenzelm@46895
   160
  apply (rule someI2)
wenzelm@46895
   161
  apply auto
wenzelm@17200
   162
  done
paulson@13551
   163
paulson@18143
   164
lemma select_not_equals:
paulson@18143
   165
     "c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c) \<noteq> c"
wenzelm@17200
   166
  apply (rule notI)
wenzelm@17200
   167
  apply (drule select_super)
berghofe@26806
   168
  apply (simp add: super_def less_le)
wenzelm@17200
   169
  done
paulson@13551
   170
wenzelm@17200
   171
lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (\<some>c'. c': super S c)"
wenzelm@17200
   172
  by (unfold succ_def) (blast intro!: if_not_P)
paulson@13551
   173
paulson@13551
   174
lemma succ_not_equals: "c \<in> chain S - maxchain S ==> succ S c \<noteq> c"
wenzelm@17200
   175
  apply (frule succI3)
wenzelm@17200
   176
  apply (simp (no_asm_simp))
wenzelm@17200
   177
  apply (rule select_not_equals, assumption)
wenzelm@17200
   178
  done
paulson@13551
   179
paulson@13551
   180
lemma TFin_chain_lemma4: "c \<in> TFin S ==> (c :: 'a set set): chain S"
wenzelm@17200
   181
  apply (erule TFin_induct)
wenzelm@17200
   182
   apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]])
nipkow@26272
   183
  apply (unfold chain_def chain_subset_def)
wenzelm@17200
   184
  apply (rule CollectI, safe)
wenzelm@17200
   185
   apply (drule bspec, assumption)
wenzelm@46895
   186
   apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE])
wenzelm@46895
   187
      apply blast+
wenzelm@17200
   188
  done
wenzelm@14706
   189
paulson@13551
   190
theorem Hausdorff: "\<exists>c. (c :: 'a set set): maxchain S"
paulson@18143
   191
  apply (rule_tac x = "Union (TFin S)" in exI)
wenzelm@17200
   192
  apply (rule classical)
wenzelm@17200
   193
  apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ")
wenzelm@17200
   194
   prefer 2
wenzelm@17200
   195
   apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric])
wenzelm@17200
   196
  apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4])
wenzelm@17200
   197
  apply (drule DiffI [THEN succ_not_equals], blast+)
wenzelm@17200
   198
  done
paulson@13551
   199
paulson@13551
   200
wenzelm@14706
   201
subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then
paulson@13551
   202
                               There Is  a Maximal Element*}
paulson@13551
   203
wenzelm@14706
   204
lemma chain_extend:
nipkow@26272
   205
  "[| c \<in> chain S; z \<in> S; \<forall>x \<in> c. x \<subseteq> (z:: 'a set) |] ==> {z} Un c \<in> chain S"
nipkow@26272
   206
by (unfold chain_def chain_subset_def) blast
paulson@13551
   207
paulson@13551
   208
lemma chain_Union_upper: "[| c \<in> chain S; x \<in> c |] ==> x \<subseteq> Union(c)"
nipkow@26272
   209
by auto
paulson@13551
   210
paulson@13551
   211
lemma chain_ball_Union_upper: "c \<in> chain S ==> \<forall>x \<in> c. x \<subseteq> Union(c)"
nipkow@26272
   212
by auto
paulson@13551
   213
paulson@13551
   214
lemma maxchain_Zorn:
nipkow@26272
   215
  "[| c \<in> maxchain S; u \<in> S; Union(c) \<subseteq> u |] ==> Union(c) = u"
nipkow@26272
   216
apply (rule ccontr)
nipkow@26272
   217
apply (simp add: maxchain_def)
nipkow@26272
   218
apply (erule conjE)
nipkow@26272
   219
apply (subgoal_tac "({u} Un c) \<in> super S c")
nipkow@26272
   220
 apply simp
berghofe@26806
   221
apply (unfold super_def less_le)
nipkow@26272
   222
apply (blast intro: chain_extend dest: chain_Union_upper)
nipkow@26272
   223
done
paulson@13551
   224
paulson@13551
   225
theorem Zorn_Lemma:
nipkow@26272
   226
  "\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z"
nipkow@26272
   227
apply (cut_tac Hausdorff maxchain_subset_chain)
nipkow@26272
   228
apply (erule exE)
nipkow@26272
   229
apply (drule subsetD, assumption)
nipkow@26272
   230
apply (drule bspec, assumption)
nipkow@26272
   231
apply (rule_tac x = "Union(c)" in bexI)
nipkow@26272
   232
 apply (rule ballI, rule impI)
nipkow@26272
   233
 apply (blast dest!: maxchain_Zorn, assumption)
nipkow@26272
   234
done
paulson@13551
   235
paulson@13551
   236
subsection{*Alternative version of Zorn's Lemma*}
paulson@13551
   237
paulson@13551
   238
lemma Zorn_Lemma2:
wenzelm@17200
   239
  "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y
wenzelm@17200
   240
    ==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x"
nipkow@26272
   241
apply (cut_tac Hausdorff maxchain_subset_chain)
nipkow@26272
   242
apply (erule exE)
nipkow@26272
   243
apply (drule subsetD, assumption)
nipkow@26272
   244
apply (drule bspec, assumption, erule bexE)
nipkow@26272
   245
apply (rule_tac x = y in bexI)
nipkow@26272
   246
 prefer 2 apply assumption
nipkow@26272
   247
apply clarify
nipkow@26272
   248
apply (rule ccontr)
nipkow@26272
   249
apply (frule_tac z = x in chain_extend)
nipkow@26272
   250
  apply (assumption, blast)
berghofe@26806
   251
apply (unfold maxchain_def super_def less_le)
nipkow@26272
   252
apply (blast elim!: equalityCE)
nipkow@26272
   253
done
paulson@13551
   254
paulson@13551
   255
text{*Various other lemmas*}
paulson@13551
   256
paulson@13551
   257
lemma chainD: "[| c \<in> chain S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x"
nipkow@26272
   258
by (unfold chain_def chain_subset_def) blast
paulson@13551
   259
paulson@13551
   260
lemma chainD2: "!!(c :: 'a set set). c \<in> chain S ==> c \<subseteq> S"
nipkow@26272
   261
by (unfold chain_def) blast
paulson@13551
   262
nipkow@26191
   263
nipkow@26191
   264
(* Define globally? In Relation.thy? *)
nipkow@26191
   265
definition Chain :: "('a*'a)set \<Rightarrow> 'a set set" where
nipkow@26191
   266
"Chain r \<equiv> {A. \<forall>a\<in>A.\<forall>b\<in>A. (a,b) : r \<or> (b,a) \<in> r}"
nipkow@26191
   267
nipkow@26191
   268
lemma mono_Chain: "r \<subseteq> s \<Longrightarrow> Chain r \<subseteq> Chain s"
nipkow@26191
   269
unfolding Chain_def by blast
nipkow@26191
   270
nipkow@26191
   271
text{* Zorn's lemma for partial orders: *}
nipkow@26191
   272
nipkow@26191
   273
lemma Zorns_po_lemma:
nipkow@26191
   274
assumes po: "Partial_order r" and u: "\<forall>C\<in>Chain r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a,u):r"
nipkow@26191
   275
shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m,a):r \<longrightarrow> a=m"
nipkow@26191
   276
proof-
nipkow@26295
   277
  have "Preorder r" using po by(simp add:partial_order_on_def)
nipkow@26191
   278
--{* Mirror r in the set of subsets below (wrt r) elements of A*}
nipkow@26191
   279
  let ?B = "%x. r^-1 `` {x}" let ?S = "?B ` Field r"
nipkow@26191
   280
  have "\<forall>C \<in> chain ?S. EX U:?S. ALL A:C. A\<subseteq>U"
nipkow@26272
   281
  proof (auto simp:chain_def chain_subset_def)
nipkow@26191
   282
    fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C.\<forall>B\<in>C. A\<subseteq>B | B\<subseteq>A"
nipkow@26191
   283
    let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}"
nipkow@26191
   284
    have "C = ?B ` ?A" using 1 by(auto simp: image_def)
nipkow@26191
   285
    have "?A\<in>Chain r"
nipkow@26191
   286
    proof (simp add:Chain_def, intro allI impI, elim conjE)
nipkow@26191
   287
      fix a b
nipkow@26191
   288
      assume "a \<in> Field r" "?B a \<in> C" "b \<in> Field r" "?B b \<in> C"
nipkow@26191
   289
      hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
nipkow@26191
   290
      thus "(a, b) \<in> r \<or> (b, a) \<in> r" using `Preorder r` `a:Field r` `b:Field r`
wenzelm@32962
   291
        by (simp add:subset_Image1_Image1_iff)
nipkow@26191
   292
    qed
nipkow@26191
   293
    then obtain u where uA: "u:Field r" "\<forall>a\<in>?A. (a,u) : r" using u by auto
nipkow@26191
   294
    have "\<forall>A\<in>C. A \<subseteq> r^-1 `` {u}" (is "?P u")
nipkow@26191
   295
    proof auto
nipkow@26191
   296
      fix a B assume aB: "B:C" "a:B"
nipkow@26191
   297
      with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto
nipkow@26191
   298
      thus "(a,u) : r" using uA aB `Preorder r`
wenzelm@32962
   299
        by (auto simp add: preorder_on_def refl_on_def) (metis transD)
nipkow@26191
   300
    qed
nipkow@26191
   301
    thus "EX u:Field r. ?P u" using `u:Field r` by blast
nipkow@26191
   302
  qed
nipkow@26191
   303
  from Zorn_Lemma2[OF this]
nipkow@26191
   304
  obtain m B where "m:Field r" "B = r^-1 `` {m}"
nipkow@26191
   305
    "\<forall>x\<in>Field r. B \<subseteq> r^-1 `` {x} \<longrightarrow> B = r^-1 `` {x}"
ballarin@27064
   306
    by auto
nipkow@26191
   307
  hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m" using po `Preorder r` `m:Field r`
nipkow@26191
   308
    by(auto simp:subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
nipkow@26191
   309
  thus ?thesis using `m:Field r` by blast
nipkow@26191
   310
qed
nipkow@26191
   311
nipkow@26191
   312
(* The initial segment of a relation appears generally useful.
nipkow@26191
   313
   Move to Relation.thy?
nipkow@26191
   314
   Definition correct/most general?
nipkow@26191
   315
   Naming?
nipkow@26191
   316
*)
nipkow@26191
   317
definition init_seg_of :: "(('a*'a)set * ('a*'a)set)set" where
nipkow@26191
   318
"init_seg_of == {(r,s). r \<subseteq> s \<and> (\<forall>a b c. (a,b):s \<and> (b,c):r \<longrightarrow> (a,b):r)}"
nipkow@26191
   319
nipkow@26191
   320
abbreviation initialSegmentOf :: "('a*'a)set \<Rightarrow> ('a*'a)set \<Rightarrow> bool"
nipkow@26191
   321
             (infix "initial'_segment'_of" 55) where
nipkow@26191
   322
"r initial_segment_of s == (r,s):init_seg_of"
nipkow@26191
   323
nipkow@30198
   324
lemma refl_on_init_seg_of[simp]: "r initial_segment_of r"
nipkow@26191
   325
by(simp add:init_seg_of_def)
nipkow@26191
   326
nipkow@26191
   327
lemma trans_init_seg_of:
nipkow@26191
   328
  "r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
nipkow@26191
   329
by(simp (no_asm_use) add: init_seg_of_def)
nipkow@26191
   330
  (metis Domain_iff UnCI Un_absorb2 subset_trans)
nipkow@26191
   331
nipkow@26191
   332
lemma antisym_init_seg_of:
nipkow@26191
   333
  "r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r=s"
huffman@35175
   334
unfolding init_seg_of_def by safe
nipkow@26191
   335
nipkow@26191
   336
lemma Chain_init_seg_of_Union:
nipkow@26191
   337
  "R \<in> Chain init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R"
nipkow@26191
   338
by(auto simp add:init_seg_of_def Chain_def Ball_def) blast
nipkow@26191
   339
nipkow@26272
   340
lemma chain_subset_trans_Union:
nipkow@26272
   341
  "chain\<^bsub>\<subseteq>\<^esub> R \<Longrightarrow> \<forall>r\<in>R. trans r \<Longrightarrow> trans(\<Union>R)"
nipkow@26272
   342
apply(auto simp add:chain_subset_def)
nipkow@26191
   343
apply(simp (no_asm_use) add:trans_def)
nipkow@26191
   344
apply (metis subsetD)
nipkow@26191
   345
done
nipkow@26191
   346
nipkow@26272
   347
lemma chain_subset_antisym_Union:
nipkow@26272
   348
  "chain\<^bsub>\<subseteq>\<^esub> R \<Longrightarrow> \<forall>r\<in>R. antisym r \<Longrightarrow> antisym(\<Union>R)"
nipkow@26272
   349
apply(auto simp add:chain_subset_def antisym_def)
nipkow@26191
   350
apply (metis subsetD)
nipkow@26191
   351
done
nipkow@26191
   352
nipkow@26272
   353
lemma chain_subset_Total_Union:
nipkow@26272
   354
assumes "chain\<^bsub>\<subseteq>\<^esub> R" "\<forall>r\<in>R. Total r"
nipkow@26191
   355
shows "Total (\<Union>R)"
nipkow@26295
   356
proof (simp add: total_on_def Ball_def, auto del:disjCI)
nipkow@26191
   357
  fix r s a b assume A: "r:R" "s:R" "a:Field r" "b:Field s" "a\<noteq>b"
nipkow@26272
   358
  from `chain\<^bsub>\<subseteq>\<^esub> R` `r:R` `s:R` have "r\<subseteq>s \<or> s\<subseteq>r"
nipkow@26272
   359
    by(simp add:chain_subset_def)
nipkow@26191
   360
  thus "(\<exists>r\<in>R. (a,b) \<in> r) \<or> (\<exists>r\<in>R. (b,a) \<in> r)"
nipkow@26191
   361
  proof
nipkow@26191
   362
    assume "r\<subseteq>s" hence "(a,b):s \<or> (b,a):s" using assms(2) A
nipkow@26295
   363
      by(simp add:total_on_def)(metis mono_Field subsetD)
nipkow@26191
   364
    thus ?thesis using `s:R` by blast
nipkow@26191
   365
  next
nipkow@26191
   366
    assume "s\<subseteq>r" hence "(a,b):r \<or> (b,a):r" using assms(2) A
nipkow@26295
   367
      by(simp add:total_on_def)(metis mono_Field subsetD)
nipkow@26191
   368
    thus ?thesis using `r:R` by blast
nipkow@26191
   369
  qed
nipkow@26191
   370
qed
nipkow@26191
   371
nipkow@26191
   372
lemma wf_Union_wf_init_segs:
nipkow@26191
   373
assumes "R \<in> Chain init_seg_of" and "\<forall>r\<in>R. wf r" shows "wf(\<Union>R)"
nipkow@26191
   374
proof(simp add:wf_iff_no_infinite_down_chain, rule ccontr, auto)
nipkow@26191
   375
  fix f assume 1: "\<forall>i. \<exists>r\<in>R. (f(Suc i), f i) \<in> r"
nipkow@26191
   376
  then obtain r where "r:R" and "(f(Suc 0), f 0) : r" by auto
nipkow@26191
   377
  { fix i have "(f(Suc i), f i) \<in> r"
nipkow@26191
   378
    proof(induct i)
nipkow@26191
   379
      case 0 show ?case by fact
nipkow@26191
   380
    next
nipkow@26191
   381
      case (Suc i)
nipkow@26191
   382
      moreover obtain s where "s\<in>R" and "(f(Suc(Suc i)), f(Suc i)) \<in> s"
wenzelm@32962
   383
        using 1 by auto
nipkow@26191
   384
      moreover hence "s initial_segment_of r \<or> r initial_segment_of s"
wenzelm@32962
   385
        using assms(1) `r:R` by(simp add: Chain_def)
nipkow@26191
   386
      ultimately show ?case by(simp add:init_seg_of_def) blast
nipkow@26191
   387
    qed
nipkow@26191
   388
  }
nipkow@26191
   389
  thus False using assms(2) `r:R`
nipkow@26191
   390
    by(simp add:wf_iff_no_infinite_down_chain) blast
nipkow@26191
   391
qed
nipkow@26191
   392
huffman@27476
   393
lemma initial_segment_of_Diff:
huffman@27476
   394
  "p initial_segment_of q \<Longrightarrow> p - s initial_segment_of q - s"
huffman@27476
   395
unfolding init_seg_of_def by blast
huffman@27476
   396
nipkow@26191
   397
lemma Chain_inits_DiffI:
nipkow@26191
   398
  "R \<in> Chain init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chain init_seg_of"
huffman@27476
   399
unfolding Chain_def by (blast intro: initial_segment_of_Diff)
nipkow@26191
   400
nipkow@26272
   401
theorem well_ordering: "\<exists>r::('a*'a)set. Well_order r \<and> Field r = UNIV"
nipkow@26191
   402
proof-
nipkow@26191
   403
-- {*The initial segment relation on well-orders: *}
nipkow@26191
   404
  let ?WO = "{r::('a*'a)set. Well_order r}"
nipkow@26191
   405
  def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
nipkow@26191
   406
  have I_init: "I \<subseteq> init_seg_of" by(auto simp:I_def)
nipkow@26272
   407
  hence subch: "!!R. R : Chain I \<Longrightarrow> chain\<^bsub>\<subseteq>\<^esub> R"
nipkow@26272
   408
    by(auto simp:init_seg_of_def chain_subset_def Chain_def)
nipkow@26191
   409
  have Chain_wo: "!!R r. R \<in> Chain I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
nipkow@26191
   410
    by(simp add:Chain_def I_def) blast
nipkow@26191
   411
  have FI: "Field I = ?WO" by(auto simp add:I_def init_seg_of_def Field_def)
nipkow@26191
   412
  hence 0: "Partial_order I"
nipkow@30198
   413
    by(auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def trans_def I_def elim!: trans_init_seg_of)
nipkow@26191
   414
-- {*I-chains have upper bounds in ?WO wrt I: their Union*}
nipkow@26191
   415
  { fix R assume "R \<in> Chain I"
nipkow@26191
   416
    hence Ris: "R \<in> Chain init_seg_of" using mono_Chain[OF I_init] by blast
nipkow@26272
   417
    have subch: "chain\<^bsub>\<subseteq>\<^esub> R" using `R : Chain I` I_init
nipkow@26272
   418
      by(auto simp:init_seg_of_def chain_subset_def Chain_def)
nipkow@26191
   419
    have "\<forall>r\<in>R. Refl r" "\<forall>r\<in>R. trans r" "\<forall>r\<in>R. antisym r" "\<forall>r\<in>R. Total r"
nipkow@26191
   420
         "\<forall>r\<in>R. wf(r-Id)"
nipkow@26295
   421
      using Chain_wo[OF `R \<in> Chain I`] by(simp_all add:order_on_defs)
nipkow@30198
   422
    have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by(auto simp:refl_on_def)
nipkow@26191
   423
    moreover have "trans (\<Union>R)"
nipkow@26272
   424
      by(rule chain_subset_trans_Union[OF subch `\<forall>r\<in>R. trans r`])
nipkow@26191
   425
    moreover have "antisym(\<Union>R)"
nipkow@26272
   426
      by(rule chain_subset_antisym_Union[OF subch `\<forall>r\<in>R. antisym r`])
nipkow@26191
   427
    moreover have "Total (\<Union>R)"
nipkow@26272
   428
      by(rule chain_subset_Total_Union[OF subch `\<forall>r\<in>R. Total r`])
nipkow@26191
   429
    moreover have "wf((\<Union>R)-Id)"
nipkow@26191
   430
    proof-
nipkow@26191
   431
      have "(\<Union>R)-Id = \<Union>{r-Id|r. r \<in> R}" by blast
nipkow@26191
   432
      with `\<forall>r\<in>R. wf(r-Id)` wf_Union_wf_init_segs[OF Chain_inits_DiffI[OF Ris]]
nipkow@26191
   433
      show ?thesis by (simp (no_asm_simp)) blast
nipkow@26191
   434
    qed
nipkow@26295
   435
    ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs)
nipkow@26191
   436
    moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
nipkow@26191
   437
      by(simp add: Chain_init_seg_of_Union)
nipkow@26191
   438
    ultimately have "\<Union>R : ?WO \<and> (\<forall>r\<in>R. (r,\<Union>R) : I)"
nipkow@26191
   439
      using mono_Chain[OF I_init] `R \<in> Chain I`
nipkow@26191
   440
      by(simp (no_asm) add:I_def del:Field_Union)(metis Chain_wo subsetD)
nipkow@26191
   441
  }
nipkow@26191
   442
  hence 1: "\<forall>R \<in> Chain I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r,u) : I" by (subst FI) blast
nipkow@26191
   443
--{*Zorn's Lemma yields a maximal well-order m:*}
nipkow@26191
   444
  then obtain m::"('a*'a)set" where "Well_order m" and
nipkow@26191
   445
    max: "\<forall>r. Well_order r \<and> (m,r):I \<longrightarrow> r=m"
nipkow@26191
   446
    using Zorns_po_lemma[OF 0 1] by (auto simp:FI)
nipkow@26191
   447
--{*Now show by contradiction that m covers the whole type:*}
nipkow@26191
   448
  { fix x::'a assume "x \<notin> Field m"
nipkow@26191
   449
--{*We assume that x is not covered and extend m at the top with x*}
nipkow@26191
   450
    have "m \<noteq> {}"
nipkow@26191
   451
    proof
nipkow@26191
   452
      assume "m={}"
nipkow@26191
   453
      moreover have "Well_order {(x,x)}"
haftmann@47620
   454
        by(simp add:order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def Domain_unfold Domain_converse [symmetric])
nipkow@26191
   455
      ultimately show False using max
wenzelm@32962
   456
        by (auto simp:I_def init_seg_of_def simp del:Field_insert)
nipkow@26191
   457
    qed
nipkow@26191
   458
    hence "Field m \<noteq> {}" by(auto simp:Field_def)
nipkow@26295
   459
    moreover have "wf(m-Id)" using `Well_order m`
nipkow@26295
   460
      by(simp add:well_order_on_def)
nipkow@26191
   461
--{*The extension of m by x:*}
nipkow@26191
   462
    let ?s = "{(a,x)|a. a : Field m}" let ?m = "insert (x,x) m Un ?s"
nipkow@26191
   463
    have Fm: "Field ?m = insert x (Field m)"
nipkow@26191
   464
      apply(simp add:Field_insert Field_Un)
nipkow@26191
   465
      unfolding Field_def by auto
nipkow@26191
   466
    have "Refl m" "trans m" "antisym m" "Total m" "wf(m-Id)"
nipkow@26295
   467
      using `Well_order m` by(simp_all add:order_on_defs)
nipkow@26191
   468
--{*We show that the extension is a well-order*}
nipkow@30198
   469
    have "Refl ?m" using `Refl m` Fm by(auto simp:refl_on_def)
nipkow@26191
   470
    moreover have "trans ?m" using `trans m` `x \<notin> Field m`
haftmann@47620
   471
      unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast
nipkow@26191
   472
    moreover have "antisym ?m" using `antisym m` `x \<notin> Field m`
haftmann@47620
   473
      unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast
nipkow@26295
   474
    moreover have "Total ?m" using `Total m` Fm by(auto simp: total_on_def)
nipkow@26191
   475
    moreover have "wf(?m-Id)"
nipkow@26191
   476
    proof-
nipkow@26191
   477
      have "wf ?s" using `x \<notin> Field m`
haftmann@47620
   478
        by(auto simp add:wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis
nipkow@26191
   479
      thus ?thesis using `wf(m-Id)` `x \<notin> Field m`
wenzelm@32962
   480
        wf_subset[OF `wf ?s` Diff_subset]
nipkow@45761
   481
        by (fastforce intro!: wf_Un simp add: Un_Diff Field_def)
nipkow@26191
   482
    qed
nipkow@26295
   483
    ultimately have "Well_order ?m" by(simp add:order_on_defs)
nipkow@26191
   484
--{*We show that the extension is above m*}
nipkow@26191
   485
    moreover hence "(m,?m) : I" using `Well_order m` `x \<notin> Field m`
haftmann@47620
   486
      by(fastforce simp:I_def init_seg_of_def Field_def Domain_unfold Domain_converse [symmetric])
nipkow@26191
   487
    ultimately
nipkow@26191
   488
--{*This contradicts maximality of m:*}
nipkow@26191
   489
    have False using max `x \<notin> Field m` unfolding Field_def by blast
nipkow@26191
   490
  }
nipkow@26191
   491
  hence "Field m = UNIV" by auto
nipkow@26272
   492
  moreover with `Well_order m` have "Well_order m" by simp
nipkow@26272
   493
  ultimately show ?thesis by blast
nipkow@26272
   494
qed
nipkow@26272
   495
nipkow@26295
   496
corollary well_order_on: "\<exists>r::('a*'a)set. well_order_on A r"
nipkow@26272
   497
proof -
nipkow@26272
   498
  obtain r::"('a*'a)set" where wo: "Well_order r" and univ: "Field r = UNIV"
nipkow@26272
   499
    using well_ordering[where 'a = "'a"] by blast
nipkow@26272
   500
  let ?r = "{(x,y). x:A & y:A & (x,y):r}"
nipkow@26272
   501
  have 1: "Field ?r = A" using wo univ
haftmann@47620
   502
    by(fastforce simp: Field_def Domain_unfold Domain_converse [symmetric] order_on_defs refl_on_def)
nipkow@26272
   503
  have "Refl r" "trans r" "antisym r" "Total r" "wf(r-Id)"
nipkow@26295
   504
    using `Well_order r` by(simp_all add:order_on_defs)
nipkow@30198
   505
  have "Refl ?r" using `Refl r` by(auto simp:refl_on_def 1 univ)
nipkow@26272
   506
  moreover have "trans ?r" using `trans r`
nipkow@26272
   507
    unfolding trans_def by blast
nipkow@26272
   508
  moreover have "antisym ?r" using `antisym r`
nipkow@26272
   509
    unfolding antisym_def by blast
nipkow@26295
   510
  moreover have "Total ?r" using `Total r` by(simp add:total_on_def 1 univ)
nipkow@26272
   511
  moreover have "wf(?r - Id)" by(rule wf_subset[OF `wf(r-Id)`]) blast
nipkow@26295
   512
  ultimately have "Well_order ?r" by(simp add:order_on_defs)
nipkow@26295
   513
  with 1 show ?thesis by metis
nipkow@26191
   514
qed
nipkow@26191
   515
paulson@13551
   516
end