src/HOL/Library/Zorn.thy
changeset 17200 3a4d03d1a31b
parent 15140 322485b816ac
child 18143 fe14f0282c60
--- a/src/HOL/Library/Zorn.thy	Wed Aug 31 15:46:36 2005 +0200
+++ b/src/HOL/Library/Zorn.thy	Wed Aug 31 15:46:37 2005 +0200
@@ -42,18 +42,20 @@
 
 subsection{*Mathematical Preamble*}
 
-lemma Union_lemma0: "(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C)<=A | B \<subseteq> Union(C)"
-by blast
+lemma Union_lemma0:
+    "(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C)<=A | B \<subseteq> Union(C)"
+  by blast
 
 
 text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*}
+
 lemma Abrial_axiom1: "x \<subseteq> succ S x"
-apply (unfold succ_def)
-apply (rule split_if [THEN iffD2])
-apply (auto simp add: super_def maxchain_def psubset_def)
-apply (rule swap, assumption)
-apply (rule someI2, blast+)
-done
+  apply (unfold succ_def)
+  apply (rule split_if [THEN iffD2])
+  apply (auto simp add: super_def maxchain_def psubset_def)
+  apply (rule swap, assumption)
+  apply (rule someI2, blast+)
+  done
 
 lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI]
 
@@ -62,79 +64,77 @@
              !!x. [| x \<in> TFin S; P(x) |] ==> P(succ S x);
              !!Y. [| Y \<subseteq> TFin S; Ball Y P |] ==> P(Union Y) |]
           ==> P(n)"
-apply (erule TFin.induct, blast+)
-done
+  apply (erule TFin.induct)
+   apply blast+
+  done
 
 lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y"
-apply (erule subset_trans)
-apply (rule Abrial_axiom1)
-done
+  apply (erule subset_trans)
+  apply (rule Abrial_axiom1)
+  done
 
 text{*Lemma 1 of section 3.1*}
 lemma TFin_linear_lemma1:
      "[| n \<in> TFin S;  m \<in> TFin S;
          \<forall>x \<in> TFin S. x \<subseteq> m --> x = m | succ S x \<subseteq> m
       |] ==> n \<subseteq> m | succ S m \<subseteq> n"
-apply (erule TFin_induct)
-apply (erule_tac [2] Union_lemma0) (*or just blast*)
-apply (blast del: subsetI intro: succ_trans)
-done
+  apply (erule TFin_induct)
+   apply (erule_tac [2] Union_lemma0)
+  apply (blast del: subsetI intro: succ_trans)
+  done
 
 text{* Lemma 2 of section 3.2 *}
 lemma TFin_linear_lemma2:
      "m \<in> TFin S ==> \<forall>n \<in> TFin S. n \<subseteq> m --> n=m | succ S n \<subseteq> m"
-apply (erule TFin_induct)
-apply (rule impI [THEN ballI])
-txt{*case split using @{text TFin_linear_lemma1}*}
-apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
-       assumption+)
-apply (drule_tac x = n in bspec, assumption)
-apply (blast del: subsetI intro: succ_trans, blast)
-txt{*second induction step*}
-apply (rule impI [THEN ballI])
-apply (rule Union_lemma0 [THEN disjE])
-apply (rule_tac [3] disjI2)
- prefer 2 apply blast
-apply (rule ballI)
-apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
-       assumption+, auto)
-apply (blast intro!: Abrial_axiom1 [THEN subsetD])
-done
+  apply (erule TFin_induct)
+   apply (rule impI [THEN ballI])
+   txt{*case split using @{text TFin_linear_lemma1}*}
+   apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
+     assumption+)
+    apply (drule_tac x = n in bspec, assumption)
+    apply (blast del: subsetI intro: succ_trans, blast)
+  txt{*second induction step*}
+  apply (rule impI [THEN ballI])
+  apply (rule Union_lemma0 [THEN disjE])
+    apply (rule_tac [3] disjI2)
+    prefer 2 apply blast
+   apply (rule ballI)
+   apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
+     assumption+, auto)
+  apply (blast intro!: Abrial_axiom1 [THEN subsetD])
+  done
 
 text{*Re-ordering the premises of Lemma 2*}
 lemma TFin_subsetD:
      "[| n \<subseteq> m;  m \<in> TFin S;  n \<in> TFin S |] ==> n=m | succ S n \<subseteq> m"
-apply (rule TFin_linear_lemma2 [rule_format])
-apply (assumption+)
-done
+  by (rule TFin_linear_lemma2 [rule_format])
 
 text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*}
 lemma TFin_subset_linear: "[| m \<in> TFin S;  n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n"
-apply (rule disjE)
-apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2])
-apply (assumption+, erule disjI2)
-apply (blast del: subsetI
-             intro: subsetI Abrial_axiom1 [THEN subset_trans])
-done
+  apply (rule disjE)
+    apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2])
+      apply (assumption+, erule disjI2)
+  apply (blast del: subsetI
+    intro: subsetI Abrial_axiom1 [THEN subset_trans])
+  done
 
 text{*Lemma 3 of section 3.3*}
 lemma eq_succ_upper: "[| n \<in> TFin S;  m \<in> TFin S;  m = succ S m |] ==> n \<subseteq> m"
-apply (erule TFin_induct)
-apply (drule TFin_subsetD)
-apply (assumption+, force, blast)
-done
+  apply (erule TFin_induct)
+   apply (drule TFin_subsetD)
+     apply (assumption+, force, blast)
+  done
 
 text{*Property 3.3 of section 3.3*}
 lemma equal_succ_Union: "m \<in> TFin S ==> (m = succ S m) = (m = Union(TFin S))"
-apply (rule iffI)
-apply (rule Union_upper [THEN equalityI])
-apply (rule_tac [2] eq_succ_upper [THEN Union_least])
-apply (assumption+)
-apply (erule ssubst)
-apply (rule Abrial_axiom1 [THEN equalityI])
-apply (blast del: subsetI
-             intro: subsetI TFin_UnionI TFin.succI)
-done
+  apply (rule iffI)
+   apply (rule Union_upper [THEN equalityI])
+    apply (rule_tac [2] eq_succ_upper [THEN Union_least])
+      apply (assumption+)
+  apply (erule ssubst)
+  apply (rule Abrial_axiom1 [THEN equalityI])
+  apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI)
+  done
 
 subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*}
 
@@ -142,60 +142,58 @@
  the subset relation!*}
 
 lemma empty_set_mem_chain: "({} :: 'a set set) \<in> chain S"
-by (unfold chain_def, auto)
+  by (unfold chain_def) auto
 
 lemma super_subset_chain: "super S c \<subseteq> chain S"
-by (unfold super_def, fast)
+  by (unfold super_def) blast
 
 lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S"
-by (unfold maxchain_def, fast)
+  by (unfold maxchain_def) blast
 
 lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> ? d. d \<in> super S c"
-by (unfold super_def maxchain_def, auto)
+  by (unfold super_def maxchain_def) auto
 
 lemma select_super: "c \<in> chain S - maxchain S ==>
-                          (@c'. c': super S c): super S c"
-apply (erule mem_super_Ex [THEN exE])
-apply (rule someI2, auto)
-done
+                          (\<some>c'. c': super S c): super S c"
+  apply (erule mem_super_Ex [THEN exE])
+  apply (rule someI2, auto)
+  done
 
 lemma select_not_equals: "c \<in> chain S - maxchain S ==>
-                          (@c'. c': super S c) \<noteq> c"
-apply (rule notI)
-apply (drule select_super)
-apply (simp add: super_def psubset_def)
-done
+                          (\<some>c'. c': super S c) \<noteq> c"
+  apply (rule notI)
+  apply (drule select_super)
+  apply (simp add: super_def psubset_def)
+  done
 
-lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (@c'. c': super S c)"
-apply (unfold succ_def)
-apply (fast intro!: if_not_P)
-done
+lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (\<some>c'. c': super S c)"
+  by (unfold succ_def) (blast intro!: if_not_P)
 
 lemma succ_not_equals: "c \<in> chain S - maxchain S ==> succ S c \<noteq> c"
-apply (frule succI3)
-apply (simp (no_asm_simp))
-apply (rule select_not_equals, assumption)
-done
+  apply (frule succI3)
+  apply (simp (no_asm_simp))
+  apply (rule select_not_equals, assumption)
+  done
 
 lemma TFin_chain_lemma4: "c \<in> TFin S ==> (c :: 'a set set): chain S"
-apply (erule TFin_induct)
-apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]])
-apply (unfold chain_def)
-apply (rule CollectI, safe)
-apply (drule bspec, assumption)
-apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE],
-       blast+)
-done
+  apply (erule TFin_induct)
+   apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]])
+  apply (unfold chain_def)
+  apply (rule CollectI, safe)
+   apply (drule bspec, assumption)
+   apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE],
+     blast+)
+  done
 
 theorem Hausdorff: "\<exists>c. (c :: 'a set set): maxchain S"
-apply (rule_tac x = "Union (TFin S) " in exI)
-apply (rule classical)
-apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ")
- prefer 2
- apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric])
-apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4])
-apply (drule DiffI [THEN succ_not_equals], blast+)
-done
+  apply (rule_tac x = "Union (TFin S) " in exI)
+  apply (rule classical)
+  apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ")
+   prefer 2
+   apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric])
+  apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4])
+  apply (drule DiffI [THEN succ_not_equals], blast+)
+  done
 
 
 subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then
@@ -204,61 +202,61 @@
 lemma chain_extend:
     "[| c \<in> chain S; z \<in> S;
         \<forall>x \<in> c. x<=(z:: 'a set) |] ==> {z} Un c \<in> chain S"
-by (unfold chain_def, blast)
+  by (unfold chain_def) blast
 
 lemma chain_Union_upper: "[| c \<in> chain S; x \<in> c |] ==> x \<subseteq> Union(c)"
-by (unfold chain_def, auto)
+  by (unfold chain_def) auto
 
 lemma chain_ball_Union_upper: "c \<in> chain S ==> \<forall>x \<in> c. x \<subseteq> Union(c)"
-by (unfold chain_def, auto)
+  by (unfold chain_def) auto
 
 lemma maxchain_Zorn:
      "[| c \<in> maxchain S; u \<in> S; Union(c) \<subseteq> u |] ==> Union(c) = u"
-apply (rule ccontr)
-apply (simp add: maxchain_def)
-apply (erule conjE)
-apply (subgoal_tac " ({u} Un c) \<in> super S c")
-apply simp
-apply (unfold super_def psubset_def)
-apply (blast intro: chain_extend dest: chain_Union_upper)
-done
+  apply (rule ccontr)
+  apply (simp add: maxchain_def)
+  apply (erule conjE)
+  apply (subgoal_tac " ({u} Un c) \<in> super S c")
+   apply simp
+  apply (unfold super_def psubset_def)
+  apply (blast intro: chain_extend dest: chain_Union_upper)
+  done
 
 theorem Zorn_Lemma:
-     "\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z"
-apply (cut_tac Hausdorff maxchain_subset_chain)
-apply (erule exE)
-apply (drule subsetD, assumption)
-apply (drule bspec, assumption)
-apply (rule_tac x = "Union (c) " in bexI)
-apply (rule ballI, rule impI)
-apply (blast dest!: maxchain_Zorn, assumption)
-done
+    "\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z"
+  apply (cut_tac Hausdorff maxchain_subset_chain)
+  apply (erule exE)
+  apply (drule subsetD, assumption)
+  apply (drule bspec, assumption)
+  apply (rule_tac x = "Union (c) " in bexI)
+   apply (rule ballI, rule impI)
+   apply (blast dest!: maxchain_Zorn, assumption)
+  done
 
 subsection{*Alternative version of Zorn's Lemma*}
 
 lemma Zorn_Lemma2:
-     "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y
-      ==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x"
-apply (cut_tac Hausdorff maxchain_subset_chain)
-apply (erule exE)
-apply (drule subsetD, assumption)
-apply (drule bspec, assumption, erule bexE)
-apply (rule_tac x = y in bexI)
- prefer 2 apply assumption
-apply clarify
-apply (rule ccontr)
-apply (frule_tac z = x in chain_extend)
-apply (assumption, blast)
-apply (unfold maxchain_def super_def psubset_def)
-apply (blast elim!: equalityCE)
-done
+  "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y
+    ==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x"
+  apply (cut_tac Hausdorff maxchain_subset_chain)
+  apply (erule exE)
+  apply (drule subsetD, assumption)
+  apply (drule bspec, assumption, erule bexE)
+  apply (rule_tac x = y in bexI)
+   prefer 2 apply assumption
+  apply clarify
+  apply (rule ccontr)
+  apply (frule_tac z = x in chain_extend)
+    apply (assumption, blast)
+  apply (unfold maxchain_def super_def psubset_def)
+  apply (blast elim!: equalityCE)
+  done
 
 text{*Various other lemmas*}
 
 lemma chainD: "[| c \<in> chain S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x"
-by (unfold chain_def, blast)
+  by (unfold chain_def) blast
 
 lemma chainD2: "!!(c :: 'a set set). c \<in> chain S ==> c \<subseteq> S"
-by (unfold chain_def, blast)
+  by (unfold chain_def) blast
 
 end