diff -r 54a3db2ed201 -r 903bb1495239 src/HOL/Library/Countable_Set.thy --- a/src/HOL/Library/Countable_Set.thy Wed Jun 17 10:57:11 2015 +0200 +++ b/src/HOL/Library/Countable_Set.thy Wed Jun 17 11:03:05 2015 +0200 @@ -3,13 +3,13 @@ Author: Andrei Popescu *) -section {* Countable sets *} +section \Countable sets\ theory Countable_Set imports Countable Infinite_Set begin -subsection {* Predicate for countable sets *} +subsection \Predicate for countable sets\ definition countable :: "'a set \ bool" where "countable S \ (\f::'a \ nat. inj_on f S)" @@ -49,18 +49,18 @@ lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)" using countableI[of to_nat A] by auto -subsection {* Enumerate a countable set *} +subsection \Enumerate a countable set\ lemma countableE_infinite: assumes "countable S" "infinite S" obtains e :: "'a \ nat" where "bij_betw e S UNIV" proof - obtain f :: "'a \ nat" where "inj_on f S" - using `countable S` by (rule countableE) + using \countable S\ by (rule countableE) then have "bij_betw f S (f`S)" unfolding bij_betw_def by simp moreover - from `inj_on f S` `infinite S` have inf_fS: "infinite (f`S)" + from \inj_on f S\ \infinite S\ have inf_fS: "infinite (f`S)" by (auto dest: finite_imageD) then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV" by (intro bij_betw_the_inv_into bij_enumerate) @@ -73,7 +73,7 @@ assumes "countable S" obtains (finite) f :: "'a \ nat" where "finite S" "bij_betw f S {.. nat" where "infinite S" "bij_betw f S UNIV" - using ex_bij_betw_finite_nat[of S] countableE_infinite `countable S` + using ex_bij_betw_finite_nat[of S] countableE_infinite \countable S\ by (cases "finite S") (auto simp add: atLeast0LessThan) definition to_nat_on :: "'a set \ 'a \ nat" where @@ -164,7 +164,7 @@ lemma inj_on_from_nat_into: "inj_on from_nat_into ({A. A \ {} \ countable A})" unfolding inj_on_def by auto -subsection {* Closure properties of countability *} +subsection \Closure properties of countability\ lemma countable_SIGMA[intro, simp]: "countable I \ (\i. i \ I \ countable (A i)) \ countable (SIGMA i : I. A i)" @@ -292,7 +292,7 @@ lemma countable_set_option [simp]: "countable (set_option x)" by(cases x) auto -subsection {* Misc lemmas *} +subsection \Misc lemmas\ lemma countable_all: assumes S: "countable S" @@ -306,13 +306,13 @@ apply (auto simp: image_iff Ball_def intro: from_nat_into split: split_if_asm) proof - fix x n assume "x \ X" "\i m. m \ i \ x \ from_nat_into X m" - with from_nat_into_surj[OF `countable X` `x \ X`] + with from_nat_into_surj[OF \countable X\ \x \ X\] show False by auto qed qed -subsection {* Uncountable *} +subsection \Uncountable\ abbreviation uncountable where "uncountable A \ \ countable A"