src/HOL/Hahn_Banach/Zorn_Lemma.thy
changeset 61879 e4f9d8f094fe
parent 61539 a29295dac1ca
--- a/src/HOL/Hahn_Banach/Zorn_Lemma.thy	Sun Dec 20 13:11:47 2015 +0100
+++ b/src/HOL/Hahn_Banach/Zorn_Lemma.thy	Sun Dec 20 13:56:02 2015 +0100
@@ -9,13 +9,13 @@
 begin
 
 text \<open>
-  Zorn's Lemmas states: if every linear ordered subset of an ordered
-  set \<open>S\<close> has an upper bound in \<open>S\<close>, then there exists a
-  maximal element in \<open>S\<close>.  In our application, \<open>S\<close> is a
-  set of sets ordered by set inclusion. Since the union of a chain of
-  sets is an upper bound for all elements of the chain, the conditions
-  of Zorn's lemma can be modified: if \<open>S\<close> is non-empty, it
-  suffices to show that for every non-empty chain \<open>c\<close> in \<open>S\<close> the union of \<open>c\<close> also lies in \<open>S\<close>.
+  Zorn's Lemmas states: if every linear ordered subset of an ordered set \<open>S\<close>
+  has an upper bound in \<open>S\<close>, then there exists a maximal element in \<open>S\<close>. In
+  our application, \<open>S\<close> is a set of sets ordered by set inclusion. Since the
+  union of a chain of sets is an upper bound for all elements of the chain,
+  the conditions of Zorn's lemma can be modified: if \<open>S\<close> is non-empty, it
+  suffices to show that for every non-empty chain \<open>c\<close> in \<open>S\<close> the union of \<open>c\<close>
+  also lies in \<open>S\<close>.
 \<close>
 
 theorem Zorn's_Lemma:
@@ -28,16 +28,14 @@
     fix c assume "c \<in> chains S"
     show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
     proof cases
-
-      txt \<open>If \<open>c\<close> is an empty chain, then every element in
-        \<open>S\<close> is an upper bound of \<open>c\<close>.\<close>
+      txt \<open>If \<open>c\<close> is an empty chain, then every element in \<open>S\<close> is an upper
+        bound of \<open>c\<close>.\<close>
 
       assume "c = {}"
       with aS show ?thesis by fast
 
-      txt \<open>If \<open>c\<close> is non-empty, then \<open>\<Union>c\<close> is an upper
-        bound of \<open>c\<close>, lying in \<open>S\<close>.\<close>
-
+      txt \<open>If \<open>c\<close> is non-empty, then \<open>\<Union>c\<close> is an upper bound of \<open>c\<close>, lying in
+        \<open>S\<close>.\<close>
     next
       assume "c \<noteq> {}"
       show ?thesis