diff -r fb6efe575c6d -r 4a2c9d32e7aa src/HOL/Library/FSet.thy --- a/src/HOL/Library/FSet.thy Mon Jul 10 16:38:42 2017 +0200 +++ b/src/HOL/Library/FSet.thy Mon Jul 10 18:53:38 2017 +0200 @@ -7,7 +7,7 @@ section \Type of finite sets defined as a subtype of sets\ theory FSet -imports Main +imports Main Countable begin subsection \Definition of the type\ @@ -1115,6 +1115,33 @@ qed +subsubsection \Countability\ + +lemma exists_fset_of_list: "\xs. fset_of_list xs = S" +including fset.lifting +by transfer (rule finite_list) + +lemma fset_of_list_surj[simp, intro]: "surj fset_of_list" +proof - + have "x \ range fset_of_list" for x :: "'a fset" + unfolding image_iff + using exists_fset_of_list by fastforce + thus ?thesis by auto +qed + +instance fset :: (countable) countable +proof + obtain to_nat :: "'a list \ nat" where "inj to_nat" + by (metis ex_inj) + moreover have "inj (inv fset_of_list)" + using fset_of_list_surj by (rule surj_imp_inj_inv) + ultimately have "inj (to_nat \ inv fset_of_list)" + by (rule inj_comp) + thus "\to_nat::'a fset \ nat. inj to_nat" + by auto +qed + + subsection \Quickcheck setup\ text \Setup adapted from sets.\