wenzelm@31795: (* Title: HOL/Hahn_Banach/Zorn_Lemma.thy wenzelm@7917: Author: Gertrud Bauer, TU Munich wenzelm@7917: *) wenzelm@7917: wenzelm@58744: header \Zorn's Lemma\ wenzelm@7917: wenzelm@31795: theory Zorn_Lemma blanchet@55018: imports Main wenzelm@27612: begin wenzelm@7917: wenzelm@58744: text \ wenzelm@10687: Zorn's Lemmas states: if every linear ordered subset of an ordered wenzelm@10687: set @{text S} has an upper bound in @{text S}, then there exists a wenzelm@10687: maximal element in @{text S}. In our application, @{text S} is a wenzelm@10687: set of sets ordered by set inclusion. Since the union of a chain of wenzelm@10687: sets is an upper bound for all elements of the chain, the conditions wenzelm@10687: of Zorn's lemma can be modified: if @{text S} is non-empty, it wenzelm@10687: suffices to show that for every non-empty chain @{text c} in @{text wenzelm@10687: S} the union of @{text c} also lies in @{text S}. wenzelm@58744: \ wenzelm@7917: wenzelm@10687: theorem Zorn's_Lemma: popescua@52183: assumes r: "\c. c \ chains S \ \x. x \ c \ \c \ S" wenzelm@13515: and aS: "a \ S" popescua@52183: shows "\y \ S. \z \ S. y \ z \ z = y" wenzelm@9035: proof (rule Zorn_Lemma2) popescua@52183: show "\c \ chains S. \y \ S. \z \ c. z \ y" wenzelm@9035: proof popescua@52183: fix c assume "c \ chains S" wenzelm@10687: show "\y \ S. \z \ c. z \ y" wenzelm@9035: proof cases wenzelm@13515: wenzelm@58744: txt \If @{text c} is an empty chain, then every element in wenzelm@58744: @{text S} is an upper bound of @{text c}.\ wenzelm@7917: wenzelm@13515: assume "c = {}" wenzelm@9035: with aS show ?thesis by fast wenzelm@7917: wenzelm@58744: txt \If @{text c} is non-empty, then @{text "\c"} is an upper wenzelm@58744: bound of @{text c}, lying in @{text S}.\ wenzelm@7917: wenzelm@9035: next wenzelm@27612: assume "c \ {}" wenzelm@13515: show ?thesis wenzelm@13515: proof wenzelm@10687: show "\z \ c. z \ \c" by fast wenzelm@13515: show "\c \ S" wenzelm@9035: proof (rule r) wenzelm@58744: from \c \ {}\ show "\x. x \ c" by fast popescua@52183: show "c \ chains S" by fact wenzelm@9035: qed wenzelm@9035: qed wenzelm@9035: qed wenzelm@9035: qed wenzelm@9035: qed wenzelm@7917: wenzelm@10687: end