diff -r 71af1fd6a5e4 -r be3e1cc5005c src/HOL/Hahn_Banach/Zorn_Lemma.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hahn_Banach/Zorn_Lemma.thy Wed Jun 24 21:46:54 2009 +0200 @@ -0,0 +1,57 @@ +(* Title: HOL/Hahn_Banach/Zorn_Lemma.thy + Author: Gertrud Bauer, TU Munich +*) + +header {* Zorn's Lemma *} + +theory Zorn_Lemma +imports Zorn +begin + +text {* + Zorn's Lemmas states: if every linear ordered subset of an ordered + set @{text S} has an upper bound in @{text S}, then there exists a + maximal element in @{text S}. In our application, @{text S} 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 @{text S} is non-empty, it + suffices to show that for every non-empty chain @{text c} in @{text + S} the union of @{text c} also lies in @{text S}. +*} + +theorem Zorn's_Lemma: + assumes r: "\c. c \ chain S \ \x. x \ c \ \c \ S" + and aS: "a \ S" + shows "\y \ S. \z \ S. y \ z \ y = z" +proof (rule Zorn_Lemma2) + show "\c \ chain S. \y \ S. \z \ c. z \ y" + proof + fix c assume "c \ chain S" + show "\y \ S. \z \ c. z \ y" + proof cases + + txt {* If @{text c} is an empty chain, then every element in + @{text S} is an upper bound of @{text c}. *} + + assume "c = {}" + with aS show ?thesis by fast + + txt {* If @{text c} is non-empty, then @{text "\c"} is an upper + bound of @{text c}, lying in @{text S}. *} + + next + assume "c \ {}" + show ?thesis + proof + show "\z \ c. z \ \c" by fast + show "\c \ S" + proof (rule r) + from `c \ {}` show "\x. x \ c" by fast + show "c \ chain S" by fact + qed + qed + qed + qed +qed + +end