1(*  Title:      HOL/Library/Countable.thy
2    Author:     Alexander Krauss, TU Muenchen
3    Author:     Brian Huffman, Portland State University
4    Author:     Jasmin Blanchette, TU Muenchen
5*)
6
7section \<open>Encoding (almost) everything into natural numbers\<close>
8
9theory Countable
10imports Old_Datatype HOL.Rat Nat_Bijection
11begin
12
13subsection \<open>The class of countable types\<close>
14
15class countable =
16  assumes ex_inj: "\<exists>to_nat :: 'a \<Rightarrow> nat. inj to_nat"
17
18lemma countable_classI:
19  fixes f :: "'a \<Rightarrow> nat"
20  assumes "\<And>x y. f x = f y \<Longrightarrow> x = y"
21  shows "OFCLASS('a, countable_class)"
22proof (intro_classes, rule exI)
23  show "inj f"
24    by (rule injI [OF assms]) assumption
25qed
26
27
28subsection \<open>Conversion functions\<close>
29
30definition to_nat :: "'a::countable \<Rightarrow> nat" where
31  "to_nat = (SOME f. inj f)"
32
33definition from_nat :: "nat \<Rightarrow> 'a::countable" where
34  "from_nat = inv (to_nat :: 'a \<Rightarrow> nat)"
35
36lemma inj_to_nat [simp]: "inj to_nat"
37  by (rule exE_some [OF ex_inj]) (simp add: to_nat_def)
38
39lemma inj_on_to_nat[simp, intro]: "inj_on to_nat S"
40  using inj_to_nat by (auto simp: inj_on_def)
41
42lemma surj_from_nat [simp]: "surj from_nat"
43  unfolding from_nat_def by (simp add: inj_imp_surj_inv)
44
45lemma to_nat_split [simp]: "to_nat x = to_nat y \<longleftrightarrow> x = y"
46  using injD [OF inj_to_nat] by auto
47
48lemma from_nat_to_nat [simp]:
49  "from_nat (to_nat x) = x"
50  by (simp add: from_nat_def)
51
52
53subsection \<open>Finite types are countable\<close>
54
55subclass (in finite) countable
56proof
57  have "finite (UNIV::'a set)" by (rule finite_UNIV)
58  with finite_conv_nat_seg_image [of "UNIV::'a set"]
59  obtain n and f :: "nat \<Rightarrow> 'a"
60    where "UNIV = f ` {i. i < n}" by auto
61  then have "surj f" unfolding surj_def by auto
62  then have "inj (inv f)" by (rule surj_imp_inj_inv)
63  then show "\<exists>to_nat :: 'a \<Rightarrow> nat. inj to_nat" by (rule exI[of inj])
64qed
65
66
67subsection \<open>Automatically proving countability of old-style datatypes\<close>
68
69context
70begin
71
72qualified inductive finite_item :: "'a Old_Datatype.item \<Rightarrow> bool" where
73  undefined: "finite_item undefined"
74| In0: "finite_item x \<Longrightarrow> finite_item (Old_Datatype.In0 x)"
75| In1: "finite_item x \<Longrightarrow> finite_item (Old_Datatype.In1 x)"
76| Leaf: "finite_item (Old_Datatype.Leaf a)"
77| Scons: "\<lbrakk>finite_item x; finite_item y\<rbrakk> \<Longrightarrow> finite_item (Old_Datatype.Scons x y)"
78
79qualified function nth_item :: "nat \<Rightarrow> ('a::countable) Old_Datatype.item"
80where
81  "nth_item 0 = undefined"
82| "nth_item (Suc n) =
83  (case sum_decode n of
84    Inl i \<Rightarrow>
85    (case sum_decode i of
86      Inl j \<Rightarrow> Old_Datatype.In0 (nth_item j)
87    | Inr j \<Rightarrow> Old_Datatype.In1 (nth_item j))
88  | Inr i \<Rightarrow>
89    (case sum_decode i of
90      Inl j \<Rightarrow> Old_Datatype.Leaf (from_nat j)
91    | Inr j \<Rightarrow>
92      (case prod_decode j of
93        (a, b) \<Rightarrow> Old_Datatype.Scons (nth_item a) (nth_item b))))"
94by pat_completeness auto
95
96lemma le_sum_encode_Inl: "x \<le> y \<Longrightarrow> x \<le> sum_encode (Inl y)"
97unfolding sum_encode_def by simp
98
99lemma le_sum_encode_Inr: "x \<le> y \<Longrightarrow> x \<le> sum_encode (Inr y)"
100unfolding sum_encode_def by simp
101
102qualified termination
103by (relation "measure id")
104  (auto simp flip: sum_encode_eq prod_encode_eq
105    simp: le_imp_less_Suc le_sum_encode_Inl le_sum_encode_Inr
106    le_prod_encode_1 le_prod_encode_2)
107
108lemma nth_item_covers: "finite_item x \<Longrightarrow> \<exists>n. nth_item n = x"
109proof (induct set: finite_item)
110  case undefined
111  have "nth_item 0 = undefined" by simp
112  thus ?case ..
113next
114  case (In0 x)
115  then obtain n where "nth_item n = x" by fast
116  hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inl n))))) = Old_Datatype.In0 x" by simp
117  thus ?case ..
118next
119  case (In1 x)
120  then obtain n where "nth_item n = x" by fast
121  hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inr n))))) = Old_Datatype.In1 x" by simp
122  thus ?case ..
123next
124  case (Leaf a)
125  have "nth_item (Suc (sum_encode (Inr (sum_encode (Inl (to_nat a)))))) = Old_Datatype.Leaf a"
126    by simp
127  thus ?case ..
128next
129  case (Scons x y)
130  then obtain i j where "nth_item i = x" and "nth_item j = y" by fast
131  hence "nth_item
132    (Suc (sum_encode (Inr (sum_encode (Inr (prod_encode (i, j))))))) = Old_Datatype.Scons x y"
133    by simp
134  thus ?case ..
135qed
136
137theorem countable_datatype:
138  fixes Rep :: "'b \<Rightarrow> ('a::countable) Old_Datatype.item"
139  fixes Abs :: "('a::countable) Old_Datatype.item \<Rightarrow> 'b"
140  fixes rep_set :: "('a::countable) Old_Datatype.item \<Rightarrow> bool"
141  assumes type: "type_definition Rep Abs (Collect rep_set)"
142  assumes finite_item: "\<And>x. rep_set x \<Longrightarrow> finite_item x"
143  shows "OFCLASS('b, countable_class)"
144proof
145  define f where "f y = (LEAST n. nth_item n = Rep y)" for y
146  {
147    fix y :: 'b
148    have "rep_set (Rep y)"
149      using type_definition.Rep [OF type] by simp
150    hence "finite_item (Rep y)"
151      by (rule finite_item)
152    hence "\<exists>n. nth_item n = Rep y"
153      by (rule nth_item_covers)
154    hence "nth_item (f y) = Rep y"
155      unfolding f_def by (rule LeastI_ex)
156    hence "Abs (nth_item (f y)) = y"
157      using type_definition.Rep_inverse [OF type] by simp
158  }
159  hence "inj f"
160    by (rule inj_on_inverseI)
161  thus "\<exists>f::'b \<Rightarrow> nat. inj f"
162    by - (rule exI)
163qed
164
165ML \<open>
166  fun old_countable_datatype_tac ctxt =
167    SUBGOAL (fn (goal, _) =>
168      let
169        val ty_name =
170          (case goal of
171            (_ $ Const (@{const_name Pure.type}, Type (@{type_name itself}, [Type (n, _)]))) => n
172          | _ => raise Match)
173        val typedef_info = hd (Typedef.get_info ctxt ty_name)
174        val typedef_thm = #type_definition (snd typedef_info)
175        val pred_name =
176          (case HOLogic.dest_Trueprop (Thm.concl_of typedef_thm) of
177            (_ $ _ $ _ $ (_ $ Const (n, _))) => n
178          | _ => raise Match)
179        val induct_info = Inductive.the_inductive_global ctxt pred_name
180        val pred_names = #names (fst induct_info)
181        val induct_thms = #inducts (snd induct_info)
182        val alist = pred_names ~~ induct_thms
183        val induct_thm = the (AList.lookup (op =) alist pred_name)
184        val vars = rev (Term.add_vars (Thm.prop_of induct_thm) [])
185        val insts = vars |> map (fn (_, T) => try (Thm.cterm_of ctxt)
186          (Const (@{const_name Countable.finite_item}, T)))
187        val induct_thm' = Thm.instantiate' [] insts induct_thm
188        val rules = @{thms finite_item.intros}
189      in
190        SOLVED' (fn i => EVERY
191          [resolve_tac ctxt @{thms countable_datatype} i,
192           resolve_tac ctxt [typedef_thm] i,
193           eresolve_tac ctxt [induct_thm'] i,
194           REPEAT (resolve_tac ctxt rules i ORELSE assume_tac ctxt i)]) 1
195      end)
196\<close>
197
198end
199
200
201subsection \<open>Automatically proving countability of datatypes\<close>
202
203ML_file "../Tools/BNF/bnf_lfp_countable.ML"
204
205ML \<open>
206fun countable_datatype_tac ctxt st =
207  (case try (fn () => HEADGOAL (old_countable_datatype_tac ctxt) st) () of
208    SOME res => res
209  | NONE => BNF_LFP_Countable.countable_datatype_tac ctxt st);
210
211(* compatibility *)
212fun countable_tac ctxt =
213  SELECT_GOAL (countable_datatype_tac ctxt);
214\<close>
215
216method_setup countable_datatype = \<open>
217  Scan.succeed (SIMPLE_METHOD o countable_datatype_tac)
218\<close> "prove countable class instances for datatypes"
219
220
221subsection \<open>More Countable types\<close>
222
223text \<open>Naturals\<close>
224
225instance nat :: countable
226  by (rule countable_classI [of "id"]) simp
227
228text \<open>Pairs\<close>
229
230instance prod :: (countable, countable) countable
231  by (rule countable_classI [of "\<lambda>(x, y). prod_encode (to_nat x, to_nat y)"])
232    (auto simp add: prod_encode_eq)
233
234text \<open>Sums\<close>
235
236instance sum :: (countable, countable) countable
237  by (rule countable_classI [of "(\<lambda>x. case x of Inl a \<Rightarrow> to_nat (False, to_nat a)
238                                     | Inr b \<Rightarrow> to_nat (True, to_nat b))"])
239    (simp split: sum.split_asm)
240
241text \<open>Integers\<close>
242
243instance int :: countable
244  by (rule countable_classI [of int_encode]) (simp add: int_encode_eq)
245
246text \<open>Options\<close>
247
248instance option :: (countable) countable
249  by countable_datatype
250
251text \<open>Lists\<close>
252
253instance list :: (countable) countable
254  by countable_datatype
255
256text \<open>String literals\<close>
257
258instance String.literal :: countable
259  by (rule countable_classI [of "to_nat \<circ> String.explode"]) (simp add: String.explode_inject)
260
261text \<open>Functions\<close>
262
263instance "fun" :: (finite, countable) countable
264proof
265  obtain xs :: "'a list" where xs: "set xs = UNIV"
266    using finite_list [OF finite_UNIV] ..
267  show "\<exists>to_nat::('a \<Rightarrow> 'b) \<Rightarrow> nat. inj to_nat"
268  proof
269    show "inj (\<lambda>f. to_nat (map f xs))"
270      by (rule injI, simp add: xs fun_eq_iff)
271  qed
272qed
273
274text \<open>Typereps\<close>
275
276instance typerep :: countable
277  by countable_datatype
278
279
280subsection \<open>The rationals are countably infinite\<close>
281
282definition nat_to_rat_surj :: "nat \<Rightarrow> rat" where
283  "nat_to_rat_surj n = (let (a, b) = prod_decode n in Fract (int_decode a) (int_decode b))"
284
285lemma surj_nat_to_rat_surj: "surj nat_to_rat_surj"
286unfolding surj_def
287proof
288  fix r::rat
289  show "\<exists>n. r = nat_to_rat_surj n"
290  proof (cases r)
291    fix i j assume [simp]: "r = Fract i j" and "j > 0"
292    have "r = (let m = int_encode i; n = int_encode j in nat_to_rat_surj (prod_encode (m, n)))"
293      by (simp add: Let_def nat_to_rat_surj_def)
294    thus "\<exists>n. r = nat_to_rat_surj n" by(auto simp: Let_def)
295  qed
296qed
297
298lemma Rats_eq_range_nat_to_rat_surj: "\<rat> = range nat_to_rat_surj"
299  by (simp add: Rats_def surj_nat_to_rat_surj)
300
301context field_char_0
302begin
303
304lemma Rats_eq_range_of_rat_o_nat_to_rat_surj:
305  "\<rat> = range (of_rat \<circ> nat_to_rat_surj)"
306  using surj_nat_to_rat_surj
307  by (auto simp: Rats_def image_def surj_def) (blast intro: arg_cong[where f = of_rat])
308
309lemma surj_of_rat_nat_to_rat_surj:
310  "r \<in> \<rat> \<Longrightarrow> \<exists>n. r = of_rat (nat_to_rat_surj n)"
311  by (simp add: Rats_eq_range_of_rat_o_nat_to_rat_surj image_def)
312
313end
314
315instance rat :: countable
316proof
317  show "\<exists>to_nat::rat \<Rightarrow> nat. inj to_nat"
318  proof
319    have "surj nat_to_rat_surj"
320      by (rule surj_nat_to_rat_surj)
321    then show "inj (inv nat_to_rat_surj)"
322      by (rule surj_imp_inj_inv)
323  qed
324qed
325
326theorem rat_denum: "\<exists>f :: nat \<Rightarrow> rat. surj f"
327 using surj_nat_to_rat_surj by metis
328
329end
330