diff -r 2bc2a8599369 -r a03a63b81f44 src/HOL/Library/Countable_Set.thy --- a/src/HOL/Library/Countable_Set.thy Mon Jan 14 18:33:53 2019 +0000 +++ b/src/HOL/Library/Countable_Set.thy Mon Jan 14 18:35:03 2019 +0000 @@ -334,16 +334,13 @@ using S[THEN subset_range_from_nat_into] by auto lemma finite_sequence_to_countable_set: - assumes "countable X" obtains F where "\i. F i \ X" "\i. F i \ F (Suc i)" "\i. finite (F i)" "(\i. F i) = X" -proof - show thesis + assumes "countable X" + obtains F where "\i. F i \ X" "\i. F i \ F (Suc i)" "\i. finite (F i)" "(\i. F i) = X" +proof - + show thesis apply (rule that[of "\i. if X = {} then {} else from_nat_into X ` {..i}"]) - apply (auto simp: image_iff Ball_def intro: from_nat_into split: if_split_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\] - show False - by auto - qed + apply (auto simp add: image_iff intro: from_nat_into split: if_splits) + using assms from_nat_into_surj by (fastforce cong: image_cong) qed lemma transfer_countable[transfer_rule]: