diff -r dd58f13a8eb4 -r eb85850d3eb7 src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy Fri Nov 17 02:19:55 2006 +0100 +++ b/src/HOL/Library/Infinite_Set.thy Fri Nov 17 02:20:03 2006 +0100 @@ -18,7 +18,7 @@ *} abbreviation - infinite :: "'a set \ bool" + infinite :: "'a set \ bool" where "infinite S == \ finite S" text {* @@ -349,17 +349,19 @@ *} definition - Inf_many :: "('a \ bool) \ bool" (binder "INF " 10) + Inf_many :: "('a \ bool) \ bool" (binder "INF " 10) where "Inf_many P = infinite {x. P x}" - Alm_all :: "('a \ bool) \ bool" (binder "MOST " 10) + +definition + Alm_all :: "('a \ bool) \ bool" (binder "MOST " 10) where "Alm_all P = (\ (INF x. \ P x))" notation (xsymbols) - Inf_many (binder "\\<^sub>\" 10) + Inf_many (binder "\\<^sub>\" 10) and Alm_all (binder "\\<^sub>\" 10) notation (HTML output) - Inf_many (binder "\\<^sub>\" 10) + Inf_many (binder "\\<^sub>\" 10) and Alm_all (binder "\\<^sub>\" 10) lemma INF_EX: @@ -453,7 +455,7 @@ *} definition - atmost_one :: "'a set \ bool" + atmost_one :: "'a set \ bool" where "atmost_one S = (\x y. x\S \ y\S \ x=y)" lemma atmost_one_empty: "S = {} \ atmost_one S"