diff --git a/thys/Inductive_Inference/CONS_LIM.thy b/thys/Inductive_Inference/CONS_LIM.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/CONS_LIM.thy @@ -0,0 +1,544 @@ +section \CONS is a proper subset of LIM\label{s:cons_lim}\ + +theory CONS_LIM + imports Inductive_Inference_Basics +begin + +text \That there are classes in @{term "LIM - CONS"} was noted by +Barzdin~\cite{b-iiafp-74,b-iiafp-77} and Blum and Blum~\cite{bb-tmtii-75}. It +was proven by Wiehagen~\cite{w-lerfss-76} (see also Wiehagen and +Zeugmann~\cite{wz-idmowle-94}). The proof uses this class:\ + +definition U_LIMCONS :: "partial1 set" ("U\<^bsub>LIM-CONS\<^esub>") where + "U\<^bsub>LIM-CONS\<^esub> \ {vs @ [j] \ p| vs j p. j \ 2 \ p \ \\<^sub>0\<^sub>1 \ \ j = vs @ [j] \ p}" + +text \Every function in @{term "U\<^bsub>LIM-CONS\<^esub>"} carries a Gödel number +greater or equal two of itself, after which only zeros and ones occur. +Thus, a strategy that always outputs the rightmost value greater or equal two +in the given prefix will converge to this Gödel number. + +The next function searches an encoded list for the rightmost element +greater or equal two.\ + +definition rmge2 :: partial1 where + "rmge2 e \ + if \i e_nth e i \ 2))" + +lemma rmge2: + assumes "xs = list_decode e" + shows "rmge2 e = + (if \i xs ! i \ 2)))" +proof - + have "(i < e_length e \ e_nth e i \ 2) = (i < length xs \ xs ! i \ 2)" for i + using assms by simp + then have "(GREATEST i. i < e_length e \ e_nth e i \ 2) = + (GREATEST i. i < length xs \ xs ! i \ 2)" + by simp + moreover have "(\ii xs ! i \ 2) < length xs" (is "Greatest ?P < _") + if "\ (\i n) = + (if \i the (f i) \ 2))))" +proof - + let ?xs = "prefix f n" + have "f \ n = list_encode ?xs" by (simp add: init_def) + moreover have "(\ii the (f i) \ 2) = + (GREATEST i. i < length ?xs \ ?xs ! i \ 2)" + using length_prefix[of f n] prefix_nth[of _ n f] by metis + moreover have "(GREATEST i. i < Suc n \ the (f i) \ 2) < Suc n" + if "\ (\ii. i the (f i) \ 2" n] by fastforce + ultimately show ?thesis using rmge2 by auto +qed + +corollary rmge2_init_total: + assumes "total1 f" + shows "rmge2 (f \ n) = + (if \i the (f i) \ 2))" + using assms total1_def rmge2_init by auto + +lemma rmge2_in_R1: "rmge2 \ \" +proof - + define g where + "g = Cn 3 r_ifle [r_constn 2 2, Cn 3 r_nth [Id 3 2, Id 3 0], Cn 3 r_nth [Id 3 2, Id 3 0], Id 3 1]" + then have "recfn 3 g" by simp + then have g: "eval g [j, r, e] \= (if 2 \ e_nth e j then e_nth e j else r)" for j r e + using g_def by simp + + let ?h = "Pr 1 Z g" + have "recfn 2 ?h" + by (simp add: \recfn 3 g\) + have h: "eval ?h [j, e] = + (if \i e_nth e i \ 2)))" for j e + proof (induction j) + case 0 + then show ?case using \recfn 2 ?h\ by auto + next + case (Suc j) + then have "eval ?h [Suc j, e] = eval g [j, the (eval ?h [j, e]), e]" + using \recfn 2 ?h\ by auto + then have *: "eval ?h [Suc j, e] \= + (if 2 \ e_nth e j then e_nth e j + else if \i e_nth e i \ 2)))" + using g Suc by auto + show ?case + proof (cases "\i e_nth e j") + case True + then have "eval ?h [Suc j, e] \= e_nth e j" + using * by simp + moreover have "(GREATEST i. i < Suc j \ e_nth e i \ 2) = j" + using ex True Greatest_equality[of "\i. i < Suc j \ e_nth e i \ 2"] + by simp + ultimately show ?thesis using ex by auto + next + case False + then have "\i 2" + using ex leI less_Suc_eq by blast + with * have "eval ?h [Suc j, e] \= e_nth e (GREATEST i. i < j \ e_nth e i \ 2)" + using False by (smt leD) + moreover have "(GREATEST i. i < Suc j \ e_nth e i \ 2) = + (GREATEST i. i < j \ e_nth e i \ 2)" + using False ex by (metis less_SucI less_Suc_eq less_antisym numeral_2_eq_2) + ultimately show ?thesis using ex by metis + qed + qed + qed + + let ?hh = "Cn 1 ?h [Cn 1 r_length [Id 1 0], Id 1 0]" + have "recfn 1 ?hh" + using `recfn 2 ?h` by simp + with h have hh: "eval ?hh [e] \= + (if \i e_nth e i \ 2))" for e + by auto + then have "eval ?hh [e] = rmge2 e" for e + unfolding rmge2_def by auto + moreover have "total ?hh" + using hh totalI1 `recfn 1 ?hh` by simp + ultimately show ?thesis using `recfn 1 ?hh` by blast +qed + +text \The first part of the main result is that @{term "U\<^bsub>LIM-CONS\<^esub> \ LIM"}.\ + +lemma U_LIMCONS_in_Lim: "U\<^bsub>LIM-CONS\<^esub> \ LIM" +proof - + have "U\<^bsub>LIM-CONS\<^esub> \ \" + unfolding U_LIMCONS_def using prepend_in_R1 RPred1_subseteq_R1 by blast + have "learn_lim \ U\<^bsub>LIM-CONS\<^esub> rmge2" + proof (rule learn_limI) + show "environment \ U\<^bsub>LIM-CONS\<^esub> rmge2" + using \U_LIMCONS \ \\ phi_in_P2 rmge2_def rmge2_in_R1 by simp + show "\i. \ i = f \ (\\<^sup>\n. rmge2 (f \ n) \= i)" if "f \ U\<^bsub>LIM-CONS\<^esub>" for f + proof - + from that obtain vs j p where + j: "j \ 2" + and p: "p \ \\<^sub>0\<^sub>1" + and s: "\ j = vs @ [j] \ p" + and f: "f = vs @ [j] \ p" + unfolding U_LIMCONS_def by auto + then have "\ j = f" by simp + from that have "total1 f" + using `U\<^bsub>LIM-CONS\<^esub> \ \` R1_imp_total1 total1_def by auto + define n\<^sub>0 where "n\<^sub>0 = length vs" + have f_gr_n0: "f n \= 0 \ f n \= 1" if "n > n\<^sub>0" for n + proof - + have "f n = p (n - n\<^sub>0 - 1)" + using that n\<^sub>0_def f by simp + with RPred1_def p show ?thesis by auto + qed + have "rmge2 (f \ n) \= j" if "n \ n\<^sub>0" for n + proof - + have n0_greatest: "(GREATEST i. i < Suc n \ the (f i) \ 2) = n\<^sub>0" + proof (rule Greatest_equality) + show "n\<^sub>0 < Suc n \ the (f n\<^sub>0) \ 2" + using n\<^sub>0_def f that j by simp + show "\y. y < Suc n \ the (f y) \ 2 \ y \ n\<^sub>0" + proof - + fix y assume "y < Suc n \ 2 \ the (f y)" + moreover have "p \ \ \ (\n. p n \= 0 \ p n \= 1)" + using RPred1_def p by blast + ultimately show "y \ n\<^sub>0" + using f_gr_n0 + by (metis Suc_1 Suc_n_not_le_n Zero_neq_Suc le_less_linear le_zero_eq option.sel) + qed + qed + have "f n\<^sub>0 \= j" + using n\<^sub>0_def f by simp + then have "\ (\i n) = f (GREATEST i. i < Suc n \ the (f i) \ 2)" + using rmge2_init_total `total1 f` by auto + with n0_greatest `f n\<^sub>0 \= j` show ?thesis by simp + qed + with `\ j = f` show ?thesis by auto + qed + qed + then show ?thesis using Lim_def by auto +qed + +text \The class @{term "U_LIMCONS"} is \emph{prefix-complete}, which +means that every non-empty list is the prefix of some function in @{term +"U_LIMCONS"}. To show this we use an auxiliary lemma: For every $f \in +\mathcal{R}$ and $k \in \mathbb{N}$ the value of $f$ at $k$ can be replaced +by a Gödel number of the function resulting from the replacement.\ + +lemma goedel_at: + fixes m :: nat and k :: nat + assumes "f \ \" + shows "\n\m. \ n = (\x. if x = k then Some n else f x)" +proof - + define psi :: "partial1 \ nat \ partial2" where + "psi = (\f k i x. (if x = k then Some i else f x))" + have "psi f k \ \\<^sup>2" + proof - + obtain r where r: "recfn 1 r" "total r" "eval r [x] = f x" for x + using assms by auto + define r_psi where + "r_psi = Cn 2 r_ifeq [Id 2 1, r_dummy 1 (r_const k), Id 2 0, Cn 2 r [Id 2 1]]" + show ?thesis + proof (rule R2I[of r_psi]) + from r_psi_def show "recfn 2 r_psi" + using r(1) by simp + have "eval r_psi [i, x] = (if x = k then Some i else f x)" for i x + proof - + have "eval (Cn 2 r [Id 2 1]) [i, x] = f x" + using r by simp + then have "eval r_psi [i, x] = eval r_ifeq [x, k, i, the (f x)]" + unfolding r_psi_def using \recfn 2 r_psi\ r R1_imp_total1[OF assms] + by simp + then show ?thesis using assms by simp + qed + then show "\x y. eval r_psi [x, y] = psi f k x y" + unfolding psi_def by simp + then show "total r_psi" + using totalI2[of r_psi] \recfn 2 r_psi\ assms psi_def by fastforce + qed + qed + then obtain n where "n \ m" "\ n = psi f k n" + using assms kleene_fixed_point[of "psi f k" m] by auto + then show ?thesis unfolding psi_def by auto +qed + +lemma U_LIMCONS_prefix_complete: + assumes "length vs > 0" + shows "\f\U\<^bsub>LIM-CONS\<^esub>. prefix f (length vs - 1) = vs" +proof - + let ?p = "\_. Some 0" + let ?f = "vs @ [0] \ ?p" + have "?f \ \" + using prepend_in_R1 RPred1_subseteq_R1 const0_in_RPred1 by blast + with goedel_at[of ?f 2 "length vs"] obtain j where + j: "j \ 2" "\ j = (\x. if x = length vs then Some j else ?f x)" (is "_ = ?g") + by auto + moreover have g: "?g x = (vs @ [j] \ ?p) x" for x + by (simp add: nth_append) + ultimately have "?g \ U\<^bsub>LIM-CONS\<^esub>" + unfolding U_LIMCONS_def using const0_in_RPred1 by fastforce + moreover have "prefix ?g (length vs - 1) = vs" + using g assms prefixI prepend_associative by auto + ultimately show ?thesis by auto +qed + +text \Roughly speaking, a strategy learning a prefix-complete class +must be total because it must be defined for every prefix in +the class. Technically, however, the empty list is not a prefix, and thus a +strategy may diverge on input 0. We can work around this by +showing that if there is a strategy learning a prefix-complete class then +there is also a total strategy learning this class. We need the result only +for consistent learning.\ + +lemma U_prefix_complete_imp_total_strategy: + assumes "\vs. length vs > 0 \ \f\U. prefix f (length vs - 1) = vs" + and "learn_cons \ U s" + shows "\t. total1 t \ learn_cons \ U t" +proof - + define t where "t = (\e. if e = 0 then Some 0 else s e)" + have "s e \" if "e > 0" for e + proof - + from that have "list_decode e \ []" (is "?vs \ _") + using list_encode_0 list_encode_decode by (metis less_imp_neq) + then have "length ?vs > 0" by simp + with assms(1) obtain f where f: "f \ U" "prefix f (length ?vs - 1) = ?vs" + by auto + with learn_cons_def learn_limE have "s (f \ (length ?vs - 1)) \" + using assms(2) by auto + then show "s e \" + using f(2) init_def by auto + qed + then have "total1 t" + using t_def by auto + have "t \ \

" + proof - + from assms(2) have "s \ \

" + using learn_consE by simp + then obtain rs where rs: "recfn 1 rs" "eval rs [x] = s x" for x + by auto + define rt where "rt = Cn 1 (r_lifz Z rs) [Id 1 0, Id 1 0]" + then have "recfn 1 rt" + using rs by auto + moreover have "eval rt [x] = t x" for x + using rs rt_def t_def by simp + ultimately show ?thesis by blast + qed + have "s (f \ n) = t (f \ n)" if "f \ U" for f n + unfolding t_def by (simp add: init_neq_zero) + then have "learn_cons \ U t" + using `t \ \

` assms(2) learn_consE[of \ U s] learn_consI[of \ U t] by simp + with `total1 t` show ?thesis by auto +qed + +text \The proof of @{prop "U\<^bsub>LIM-CONS\<^esub> \ CONS"} is by contradiction. +Assume there is a consistent learning strategy $S$. By the previous +lemma $S$ can be assumed to be total. Moreover it outputs a consistent +hypothesis for every prefix. Thus for every $e \in \mathbb{N}^+$, $S(e) \neq +S(e0)$ or $S(e) \neq S(e1)$ because $S(e)$ cannot be consistent with both +$e0$ and $e1$. We use this property of $S$ to construct a function in @{term +"U\<^bsub>LIM-CONS\<^esub>"} for which $S$ fails as a learning strategy. To +this end we define a numbering $\psi \in \mathcal{R}^2$ with $\psi_i(0) = i$ +and +\[ +\psi_i(x + 1) = \left\{\begin{array}{ll} + 0 & \mbox{if } S(\psi_i^x0) \neq S(\psi_i^x),\\ + 1 & \mbox{otherwise}. +\end{array}\right. +\] +This numbering is recursive because $S$ is total. The ``otherwise'' case is +equivalent to $S(\psi_i^x1) \neq S(\psi_i^x)$ because $S(\psi_i^x)$ cannot be +consistent with both $\psi_i^x0$ and $\psi_i^x1$. Therefore every prefix +$\psi_i^x$ is extended in such a way that $S$ changes its hypothesis. Hence +$S$ does not learn $\psi_i$ in the limit. Kleene's fixed-point theorem +ensures that for some $j \geq 2$, $\varphi_j = \psi_j$. This $\psi_j$ is the +sought function in @{term "U\<^bsub>LIM-CONS\<^esub>"}. + +The following locale formalizes the construction of $\psi$ for a total +strategy $S$.\ + +locale cons_lim = + fixes s :: partial1 + assumes s_in_R1: "s \ \" +begin + +text \A @{typ recf} computing the strategy:\ + +definition r_s :: recf where + "r_s \ SOME r_s. recfn 1 r_s \ total r_s \ s = (\x. eval r_s [x])" + +lemma r_s_recfn [simp]: "recfn 1 r_s" + and r_s_total [simp]: "\x. eval r_s [x] \" + and eval_r_s: "s = (\x. eval r_s [x])" + using r_s_def R1_SOME[OF s_in_R1, of r_s] by simp_all + +text \The next function represents the prefixes of $\psi_i$.\ + +fun prefixes :: "nat \ nat \ nat list" where + "prefixes i 0 = [i]" +| "prefixes i (Suc x) = (prefixes i x) @ + [if s (e_snoc (list_encode (prefixes i x)) 0) = s (list_encode (prefixes i x)) + then 1 else 0]" + +definition "r_prefixes_aux \ + Cn 3 r_ifeq + [Cn 3 r_s [Cn 3 r_snoc [Id 3 1, r_constn 2 0]], + Cn 3 r_s [Id 3 1], + Cn 3 r_snoc [Id 3 1, r_constn 2 1], + Cn 3 r_snoc [Id 3 1, r_constn 2 0]]" + +lemma r_prefixes_aux_recfn: "recfn 3 r_prefixes_aux" + unfolding r_prefixes_aux_def by simp + +lemma r_prefixes_aux: + "eval r_prefixes_aux [j, v, i] \= + e_snoc v (if eval r_s [e_snoc v 0] = eval r_s [v] then 1 else 0)" + unfolding r_prefixes_aux_def by auto + +definition "r_prefixes \ r_swap (Pr 1 r_singleton_encode r_prefixes_aux)" + +lemma r_prefixes_recfn: "recfn 2 r_prefixes" + unfolding r_prefixes_def r_prefixes_aux_def by simp + +lemma r_prefixes: "eval r_prefixes [i, n] \= list_encode (prefixes i n)" +proof - + let ?h = "Pr 1 r_singleton_encode r_prefixes_aux" + have "eval ?h [n, i] \= list_encode (prefixes i n)" + proof (induction n) + case 0 + then show ?case + using r_prefixes_def r_prefixes_aux_recfn r_singleton_encode by simp + next + case (Suc n) + then show ?case + using r_prefixes_aux_recfn r_prefixes_aux eval_r_s + by auto metis+ + qed + moreover have "eval ?h [n, i] = eval r_prefixes [i, n]" for i n + unfolding r_prefixes_def by (simp add: r_prefixes_aux_recfn) + ultimately show ?thesis by simp +qed + +lemma prefixes_neq_nil: "length (prefixes i x) > 0" + by (induction x) auto + +text \The actual numbering can then be defined via @{term prefixes}.\ + +definition psi :: "partial2" ("\") where + "\ i x \ Some (last (prefixes i x))" + +lemma psi_in_R2: "\ \ \\<^sup>2" +proof + define r_psi where "r_psi \ Cn 2 r_last [r_prefixes]" + have "recfn 2 r_psi" + unfolding r_psi_def by (simp add: r_prefixes_recfn) + then have "eval r_psi [i, n] \= last (prefixes i n)" for n i + unfolding r_psi_def using r_prefixes r_prefixes_recfn prefixes_neq_nil by simp + then have "(\i x. Some (last (prefixes i x))) \ \

\<^sup>2" + using `recfn 2 r_psi` P2I[of "r_psi"] by simp + with psi_def show "\ \ \

\<^sup>2" by presburger + moreover show "total2 psi" + unfolding psi_def by auto +qed + +lemma psi_0_or_1: + assumes "n > 0" + shows "\ i n \= 0 \ \ i n \= 1" +proof - + from assms obtain m where "n = Suc m" + using gr0_implies_Suc by blast + then have "last (prefixes i (Suc m)) = 0 \ last (prefixes i (Suc m)) = 1" + by simp + then show ?thesis using \n = Suc m\ psi_def by simp +qed + +text \The function @{term "prefixes"} does indeed provide the prefixes +for @{term "\"}.\ + +lemma psi_init: "(\ i) \ x = list_encode (prefixes i x)" +proof - + have "prefix (\ i) x = prefixes i x" + unfolding psi_def + by (induction x) (simp_all add: prefix_0 prefix_Suc) + with init_def show ?thesis by simp +qed + +text \One of the functions $\psi_i$ is in @{term "U\<^bsub>LIM-CONS\<^esub>"}.\ + +lemma ex_psi_in_U: "\j. \ j \ U\<^bsub>LIM-CONS\<^esub>" +proof - + obtain j where j: "j \ 2" "\ j = \ j" + using kleene_fixed_point[of \] psi_in_R2 R2_imp_P2 by metis + then have "\ j \ \

" by (simp add: phi_in_P2) + define p where "p = (\x. \ j (x + 1))" + have "p \ \\<^sub>0\<^sub>1" + proof - + from p_def `\ j \ \

` skip_P1 have "p \ \

" by blast + from psi_in_R2 have "total1 (\ j)" by simp + with p_def have "total1 p" + by (simp add: total1_def) + with psi_0_or_1 have "p n \= 0 \ p n \= 1" for n + using psi_def p_def by simp + then show ?thesis + by (simp add: RPred1_def P1_total_imp_R1 \p \ \

\ \total1 p\) + qed + moreover have "\ j = [j] \ p" + proof + fix x + show "\ j x = ([j] \ p) x" + proof (cases "x = 0") + case True + then show ?thesis using psi_def psi_def prepend_at_less by simp + next + case False + then show ?thesis using p_def by simp + qed + qed + ultimately have "\ j \ U\<^bsub>LIM-CONS\<^esub>" + using j U_LIMCONS_def by (metis (mono_tags, lifting) append_Nil mem_Collect_eq) + then show ?thesis by auto +qed + +text \The strategy fails to learn @{term U_LIMCONS} because it changes +its hypothesis all the time on functions $\psi_j \in V_0$.\ + +lemma U_LIMCONS_not_learn_cons: "\ learn_cons \ U\<^bsub>LIM-CONS\<^esub> s" +proof + assume learn: "learn_cons \ U\<^bsub>LIM-CONS\<^esub> s" + have "s (list_encode (vs @ [0])) \ s (list_encode (vs @ [1]))" for vs + proof - + obtain f\<^sub>0 where f0: "f\<^sub>0 \ U\<^bsub>LIM-CONS\<^esub>" "prefix f\<^sub>0 (length vs) = vs @ [0]" + using U_LIMCONS_prefix_complete[of "vs @ [0]"] by auto + obtain f\<^sub>1 where f1: "f\<^sub>1 \ U\<^bsub>LIM-CONS\<^esub>" "prefix f\<^sub>1 (length vs) = vs @ [1]" + using U_LIMCONS_prefix_complete[of "vs @ [1]"] by auto + have "f\<^sub>0 (length vs) \ f\<^sub>1 (length vs)" + using f0 f1 by (metis lessI nth_append_length prefix_nth zero_neq_one) + moreover have "\ (the (s (f\<^sub>0 \ length vs))) (length vs) = f\<^sub>0 (length vs)" + using learn_consE(3)[of \ U_LIMCONS s, OF learn, of f\<^sub>0 "length vs", OF f0(1)] + by simp + moreover have "\ (the (s (f\<^sub>1 \ length vs))) (length vs) = f\<^sub>1 (length vs)" + using learn_consE(3)[of \ U_LIMCONS s, OF learn, of f\<^sub>1 "length vs", OF f1(1)] + by simp + ultimately have "the (s (f\<^sub>0 \ length vs)) \ the (s (f\<^sub>1 \ length vs))" + by auto + then have "s (f\<^sub>0 \ length vs) \ s (f\<^sub>1 \ length vs)" + by auto + with f0(2) f1(2) show ?thesis by (simp add: init_def) + qed + then have "s (list_encode (vs @ [0])) \ s (list_encode vs) \ + s (list_encode (vs @ [1])) \ s (list_encode vs)" + for vs + by metis + then have "s (list_encode (prefixes i (Suc x))) \ s (list_encode (prefixes i x))" for i x + by simp + then have "\ learn_lim \ {\ i} s" for i + using psi_def psi_init always_hyp_change_not_Lim by simp + then have "\ learn_lim \ U_LIMCONS s" + using ex_psi_in_U learn_lim_closed_subseteq by blast + then show False + using learn learn_cons_def by simp +qed + +end + +text \With the locale we can now show the second part of the main +result:\ + +lemma U_LIMCONS_not_in_CONS: "U\<^bsub>LIM-CONS\<^esub> \ CONS" +proof + assume "U\<^bsub>LIM-CONS\<^esub> \ CONS" + then have "U\<^bsub>LIM-CONS\<^esub> \ CONS_wrt \" + by (simp add: CONS_wrt_phi_eq_CONS) + then obtain almost_s where "learn_cons \ U\<^bsub>LIM-CONS\<^esub> almost_s" + using CONS_wrt_def by auto + then obtain s where s: "total1 s" "learn_cons \ U\<^bsub>LIM-CONS\<^esub> s" + using U_LIMCONS_prefix_complete U_prefix_complete_imp_total_strategy by blast + then have "s \ \" + using learn_consE(1) P1_total_imp_R1 by blast + with cons_lim_def interpret cons_lim s by simp + show False + using s(2) U_LIMCONS_not_learn_cons by simp +qed + +text \The main result of this section:\ + +theorem CONS_subset_Lim: "CONS \ LIM" + using U_LIMCONS_in_Lim U_LIMCONS_not_in_CONS CONS_subseteq_Lim by auto + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/CP_FIN_NUM.thy b/thys/Inductive_Inference/CP_FIN_NUM.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/CP_FIN_NUM.thy @@ -0,0 +1,1032 @@ +section \FIN is a proper subset of CP\label{s:fin_cp}\ + +theory CP_FIN_NUM + imports Inductive_Inference_Basics +begin + +text \Let $S$ be a FIN strategy for a non-empty class $U$. Let $T$ be a +strategy that hypothesizes an arbitrary function from $U$ while $S$ outputs +``don't know'' and the hypothesis of $S$ otherwise. Then $T$ is a CP strategy +for $U$.\ + +lemma nonempty_FIN_wrt_impl_CP: + assumes "U \ {}" and "U \ FIN_wrt \" + shows "U \ CP_wrt \" +proof - + obtain s where "learn_fin \ U s" + using assms(2) FIN_wrt_def by auto + then have env: "environment \ U s" and + fin: "\f. f \ U \ + \i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + using learn_finE by auto + from assms(1) obtain f\<^sub>0 where "f\<^sub>0 \ U" + by auto + with fin obtain i\<^sub>0 where "\ i\<^sub>0 = f\<^sub>0" + by blast + define t where "t x \ + (if s x \ then None else if s x \= 0 then Some i\<^sub>0 else Some (the (s x) - 1))" + for x + have "t \ \

" + proof - + from env obtain rs where rs: "recfn 1 rs" "\x. eval rs [x] = s x" + by auto + define rt where "rt = Cn 1 r_ifz [rs, r_const i\<^sub>0, Cn 1 r_dec [rs]]" + then have "recfn 1 rt" + using rs(1) by simp + then have "eval rt [x] \= (if s x \= 0 then i\<^sub>0 else (the (s x)) - 1)" if "s x \" for x + using rs rt_def that by auto + moreover have "eval rt [x] \" if "eval rs [x] \" for x + using rs rt_def that by simp + ultimately have "eval rt [x] = t x" for x + using rs(2) t_def by simp + with \recfn 1 rt\ show ?thesis by auto + qed + have "learn_cp \ U t" + proof (rule learn_cpI) + show "environment \ U t" + using env t_def \t \ \

\ by simp + show "\i. \ i = f \ (\\<^sup>\n. t (f \ n) \= i)" if "f \ U" for f + proof - + from that fin obtain i n\<^sub>0 where + i: "\ i = f" "\n0. s (f \ n) \= 0" "\n\n\<^sub>0. s (f \ n) \= Suc i" + by blast + moreover have "\n\n\<^sub>0. t (f \ n) \= i" + using that t_def i(3) by simp + ultimately show ?thesis by auto + qed + show "\ (the (t (f \ n))) \ U" if "f \ U" for f n + using \\ i\<^sub>0 = f\<^sub>0\ \f\<^sub>0 \ U\ t_def fin env that + by (metis (no_types, lifting) diff_Suc_1 not_less option.sel) + qed + then show ?thesis using CP_wrt_def env by auto +qed + +lemma FIN_wrt_impl_CP: + assumes "U \ FIN_wrt \" + shows "U \ CP_wrt \" +proof (cases "U = {}") + case True + then have "\ \ \

\<^sup>2 \ U \ CP_wrt \" + using CP_wrt_def learn_cpI[of \ "{}" "\x. Some 0"] const_in_Prim1 by auto + moreover have "\ \ \

\<^sup>2" + using assms FIN_wrt_def learn_finE by auto + ultimately show "U \ CP_wrt \" by simp +next + case False + with nonempty_FIN_wrt_impl_CP assms show ?thesis + by simp +qed + +corollary FIN_subseteq_CP: "FIN \ CP" +proof + fix U + assume "U \ FIN" + then have "\\. U \ FIN_wrt \" + using FIN_def FIN_wrt_def by auto + then have "\\. U \ CP_wrt \" + using FIN_wrt_impl_CP by auto + then show "U \ CP" + by (simp add: CP_def CP_wrt_def) +qed + +text \In order to show the \emph{proper} inclusion, we show @{term +"U\<^sub>0 \ CP - FIN"}. A CP strategy for @{term "U\<^sub>0"} simply +hypothesizes the function in @{term U0} with the longest prefix of $f^n$ not +ending in zero. For that we define a function computing the index of the +rightmost non-zero value in a list, returning the length of the list if there +is no such value.\ + +definition findr :: partial1 where + "findr e \ + if \i 0 + then Some (GREATEST i. i < e_length e \ e_nth e i \ 0) + else Some (e_length e)" + +lemma findr_total: "findr e \" + unfolding findr_def by simp + +lemma findr_ex: + assumes "\i 0" + shows "the (findr e) < e_length e" + and "e_nth e (the (findr e)) \ 0" + and "\i. the (findr e) < i \ i < e_length e \ e_nth e i = 0" +proof - + let ?P = "\i. i < e_length e \ e_nth e i \ 0" + from assms have "\i. ?P i" by simp + then have "?P (Greatest ?P)" + using GreatestI_ex_nat[of ?P "e_length e"] by fastforce + moreover have *: "findr e = Some (Greatest ?P)" + using assms findr_def by simp + ultimately show "the (findr e) < e_length e" and "e_nth e (the (findr e)) \ 0" + by fastforce+ + show "\i. the (findr e) < i \ i < e_length e \ e_nth e i = 0" + using * Greatest_le_nat[of ?P _ "e_length e"] by fastforce +qed + +definition "r_findr \ + let g = + Cn 3 r_ifz + [Cn 3 r_nth [Id 3 2, Id 3 0], + Cn 3 r_ifeq [Id 3 0, Id 3 1, Cn 3 S [Id 3 0], Id 3 1], + Id 3 0] + in Cn 1 (Pr 1 Z g) [Cn 1 r_length [Id 1 0], Id 1 0]" + +lemma r_findr_prim [simp]: "prim_recfn 1 r_findr" + unfolding r_findr_def by simp + +lemma r_findr [simp]: "eval r_findr [e] = findr e" +proof - + define g where "g = + Cn 3 r_ifz + [Cn 3 r_nth [Id 3 2, Id 3 0], + Cn 3 r_ifeq [Id 3 0, Id 3 1, Cn 3 S [Id 3 0], Id 3 1], + Id 3 0]" + then have "recfn 3 g" + by simp + with g_def have g: "eval g [j, r, e] \= + (if e_nth e j \ 0 then j else if j = r then Suc j else r)" for j r e + by simp + let ?h = "Pr 1 Z g" + have "recfn 2 ?h" + by (simp add: \recfn 3 g\) + let ?P = "\e j i. i < j \ e_nth e i \ 0" + let ?G = "\e j. Greatest (?P e j)" + have h: "eval ?h [j, e] = + (if \irecfn 2 ?h\ by auto + next + case (Suc j) + then have "eval ?h [Suc j, e] = eval g [j, the (eval ?h [j, e]), e]" + using \recfn 2 ?h\ by auto + then have "eval ?h [Suc j, e] = + eval g [j, if \i= + (if e_nth e j \ 0 then j + else if j = (if \iiii 0" + by auto + show ?thesis + proof (cases "e_nth e j = 0") + case True + then have ex': "\i 0" + using ex less_Suc_eq by fastforce + then have "(if \i= ?G e j" + using * True by simp + moreover have "?G e j = ?G e (Suc j)" + using True by (metis less_SucI less_Suc_eq) + ultimately show ?thesis using ex by metis + next + case False + then have "eval ?h [Suc j, e] \= j" + using * by simp + moreover have "?G e (Suc j) = j" + using ex False Greatest_equality[of "?P e (Suc j)"] by simp + ultimately show ?thesis using ex by simp + qed + qed + qed + let ?hh = "Cn 1 ?h [Cn 1 r_length [Id 1 0], Id 1 0]" + have "recfn 1 ?hh" + using `recfn 2 ?h` by simp + with h have hh: "eval ?hh [e] \= + (if \i0 \ CP" +proof - + define s where + "s \ \x. if findr x \= e_length x then Some 0 else Some (e_take (Suc (the (findr x))) x)" + have "s \ \

" + proof - + define r where + "r \ Cn 1 r_ifeq [r_findr, r_length, Z, Cn 1 r_take [Cn 1 S [r_findr], Id 1 0]]" + then have "\x. eval r [x] = s x" + using s_def findr_total by fastforce + moreover have "recfn 1 r" + using r_def by simp + ultimately show ?thesis by auto + qed + moreover have "learn_cp prenum U\<^sub>0 s" + proof (rule learn_cpI) + show "environment prenum U\<^sub>0 s" + using \s \ \

\ s_def prenum_in_R2 U0_in_NUM by auto + show "\i. prenum i = f \ (\\<^sup>\n. s (f \ n) \= i)" if "f \ U\<^sub>0" for f + proof (cases "f = (\_. Some 0)") + case True + then have "s (f \ n) \= 0" for n + using findr_def s_def by simp + then have "\n\0. s (f \ n) \= 0" by simp + moreover have "prenum 0 = f" + using True by auto + ultimately show ?thesis by auto + next + case False + then obtain ws where ws: "length ws > 0" "last ws \ 0" "f = ws \ 0\<^sup>\" + using U0_def \f \ U\<^sub>0\ almost0_canonical by blast + let ?m = "length ws - 1" + let ?i = "list_encode ws" + have "prenum ?i = f" + using ws by auto + moreover have "s (f \ n) \= ?i" if "n \ ?m" for n + proof - + have "e_nth (f \ n) ?m \ 0" + using ws that by (simp add: last_conv_nth) + then have "\k n) k \ 0" + using le_imp_less_Suc that by blast + moreover have + "(GREATEST k. k < e_length (f \ n) \ e_nth (f \ n) k \ 0) = ?m" + proof (rule Greatest_equality) + show "?m < e_length (f \ n) \ e_nth (f \ n) ?m \ 0" + using \e_nth (f \ n) ?m \ 0\ that by auto + show "\y. y < e_length (f \ n) \ e_nth (f \ n) y \ 0 \ y \ ?m" + using ws less_Suc_eq_le by fastforce + qed + ultimately have "findr (f \ n) \= ?m" + using that findr_def by simp + moreover have "?m < e_length (f \ n)" + using that by simp + ultimately have "s (f \ n) \= e_take (Suc ?m) (f \ n)" + using s_def by simp + moreover have "e_take (Suc ?m) (f \ n) = list_encode ws" + proof - + have "take (Suc ?m) (prefix f n) = prefix f ?m" + using take_prefix[of f ?m n] ws that by (simp add: almost0_in_R1) + then have "take (Suc ?m) (prefix f n) = ws" + using ws prefixI by auto + then show ?thesis by simp + qed + ultimately show ?thesis by simp + qed + ultimately show ?thesis by auto + qed + show "\f n. f \ U\<^sub>0 \ prenum (the (s (f \ n))) \ U\<^sub>0" + using U0_def by fastforce + qed + ultimately show ?thesis using CP_def by blast +qed + +text \As a bit of an interlude, we can now show that CP is not +closed under the subset relation. This works by removing functions from +@{term "U\<^sub>0"} in a ``noncomputable'' way such that a strategy cannot ensure +that every intermediate hypothesis is in that new class.\ + +lemma CP_not_closed_subseteq: "\V U. V \ U \ U \ CP \ V \ CP" +proof - + \ \The numbering $g\in\mathcal{R}^2$ enumerates all + functions $i0^\infty \in U_0$.\ + define g where "g \ \i. [i] \ 0\<^sup>\" + have g_inj: "i = j" if "g i = g j" for i j + proof - + have "g i 0 \= i" and "g j 0 \= j" + by (simp_all add: g_def) + with that show "i = j" + by (metis option.inject) + qed + + \ \Define a class $V$. If the strategy $\varphi_i$ learns + $g_i$, it outputs a hypothesis for $g_i$ on some shortest prefix $g_i^m$. + Then the function $g_i^m10^\infty$ is included in the class $V$; otherwise + $g_i$ is included.\ + define V where "V \ + {if learn_lim \ {g i} (\ i) + then (prefix (g i) (LEAST n. \ (the (\ i ((g i) \ n))) = g i)) @ [1] \ 0\<^sup>\ + else g i | + i. i \ UNIV}" + have "V \ CP_wrt \" + proof + \ \Assuming $V \in CP_\varphi$, there is a CP strategy + $\varphi_i$ for $V$.\ + assume "V \ CP_wrt \" + then obtain s where s: "s \ \

" "learn_cp \ V s" + using CP_wrt_def learn_cpE(1) by auto + then obtain i where i: "\ i = s" + using phi_universal by auto + + show False + proof (cases "learn_lim \ {g i} (\ i)") + case learn: True + \ \If $\varphi_i$ learns $g_i$, it hypothesizes $g_i$ on + some shortest prefix $g_i^m$. Thus it hypothesizes $g_i$ on some prefix + of $g_i^m10^\infty \in V$, too. But $g_i$ is not a class-preserving + hypothesis because $g_i \notin V$.\ + let ?P = "\n. \ (the (\ i ((g i) \ n))) = g i" + let ?m = "Least ?P" + have "\n. ?P n" + using i s by (meson learn infinite_hyp_wrong_not_Lim insertI1 lessI) + then have "?P ?m" + using LeastI_ex[of ?P] by simp + define h where "h = (prefix (g i) ?m) @ [1] \ 0\<^sup>\" + then have "h \ V" + using V_def learn by auto + have "(g i) \ ?m = h \ ?m" + proof - + have "prefix (g i) ?m = prefix h ?m" + unfolding h_def by (simp add: prefix_prepend_less) + then show ?thesis by auto + qed + then have "\ (the (\ i (h \ ?m))) = g i" + using `?P ?m` by simp + moreover have "g i \ V" + proof + assume "g i \ V" + then obtain j where j: "g i = + (if learn_lim \ {g j} (\ j) + then (prefix (g j) (LEAST n. \ (the (\ j ((g j) \ n))) = g j)) @ [1] \ 0\<^sup>\ + else g j)" + using V_def by auto + show False + proof (cases "learn_lim \ {g j} (\ j)") + case True + then have "g i = + (prefix (g j) (LEAST n. \ (the (\ j ((g j) \ n))) = g j)) @ [1] \ 0\<^sup>\" + (is "g i = ?vs @ [1] \ 0\<^sup>\") + using j by simp + moreover have len: "length ?vs > 0" by simp + ultimately have "g i (length ?vs) \= 1" + by (simp add: prepend_associative) + moreover have "g i (length ?vs) \= 0" + using g_def len by simp + ultimately show ?thesis by simp + next + case False + then show ?thesis + using j g_inj learn by auto + qed + qed + ultimately have "\ (the (\ i (h \ ?m))) \ V" by simp + then have "\ learn_cp \ V (\ i)" + using `h \ V` learn_cpE(3) by auto + then show ?thesis by (simp add: i s(2)) + next + \ \If $\varphi_i$ does not learn $g_i$, then $g_i\in V$. + Hence $\varphi_i$ does not learn $V$.\ + case False + then have "g i \ V" + using V_def by auto + with False have "\ learn_lim \ V (\ i)" + using learn_lim_closed_subseteq by auto + then show ?thesis + using s(2) i by (simp add: learn_cp_def) + qed + qed + then have "V \ CP" + using CP_wrt_phi by simp + moreover have "V \ U\<^sub>0" + using V_def g_def U0_def by auto + ultimately show ?thesis using U0_in_CP by auto +qed + +text \Continuing with the main result of this section, we show that +@{term "U\<^sub>0"} cannot be learned finitely. Any FIN strategy would have +to output a hypothesis for the constant zero function on some prefix. But +@{term "U\<^sub>0"} contains infinitely many other functions starting with +the same prefix, which the strategy then would not learn finitely.\ + +lemma U0_not_in_FIN: "U\<^sub>0 \ FIN" +proof + assume "U\<^sub>0 \ FIN" + then obtain \ s where "learn_fin \ U\<^sub>0 s" + using FIN_def by blast + with learn_finE have cp: "\f. f \ U\<^sub>0 \ + \i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + by simp_all + + define z where "z = [] \ 0\<^sup>\" + then have "z \ U\<^sub>0" + using U0_def by auto + with cp obtain i n\<^sub>0 where i: "\ i = z" and n0: "\n\n\<^sub>0. s (z \ n) \= Suc i" + by blast + + define w where "w = replicate (Suc n\<^sub>0) 0 @ [1] \ 0\<^sup>\" + then have "prefix w n\<^sub>0 = replicate (Suc n\<^sub>0) 0" + by (simp add: prefix_prepend_less) + moreover have "prefix z n\<^sub>0 = replicate (Suc n\<^sub>0) 0" + using prefixI[of "replicate (Suc n\<^sub>0) 0" z] less_Suc_eq_0_disj unfolding z_def + by fastforce + ultimately have "z \ n\<^sub>0 = w \ n\<^sub>0" + by (simp add: init_prefixE) + with n0 have *: "s (w \ n\<^sub>0) \= Suc i" by auto + + have "w \ U\<^sub>0" using w_def U0_def by auto + with cp obtain i' n\<^sub>0' where i': "\ i' = w" + and n0': "\n0'. s (w \ n) \= 0" "\n\n\<^sub>0'. s (w \ n) \= Suc i'" + by blast + + have "i \ i'" + proof + assume "i = i'" + then have "w = z" + using i i' by simp + have "w (Suc n\<^sub>0) \= 1" + using w_def prepend[of "replicate (Suc n\<^sub>0) 0 @ [1]" "0\<^sup>\" "Suc n\<^sub>0"] + by (metis length_append_singleton length_replicate lessI nth_append_length) + moreover have "z (Suc n\<^sub>0) \= 0" + using z_def by simp + ultimately show False + using \w = z\ by simp + qed + then have "s (w \ n\<^sub>0) \\ Suc i" + using n0' by (cases "n\<^sub>0 < n\<^sub>0'") simp_all + with * show False by simp +qed + +theorem FIN_subset_CP: "FIN \ CP" + using U0_in_CP U0_not_in_FIN FIN_subseteq_CP by auto + + +section \NUM and FIN are incomparable\label{s:num_fin}\ + +text \The class $V_0$ of all total recursive functions $f$ where $f(0)$ +is a Gödel number of $f$ can be learned finitely by always hypothesizing +$f(0)$. The class is not in NUM and therefore serves to separate NUM and +FIN.\ + +definition V0 :: "partial1 set" ("V\<^sub>0") where + "V\<^sub>0 = {f. f \ \ \ \ (the (f 0)) = f}" + +lemma V0_altdef: "V\<^sub>0 = {[i] \ f| i f. f \ \ \ \ i = [i] \ f}" + (is "V\<^sub>0 = ?W") +proof + show "V\<^sub>0 \ ?W" + proof + fix f + assume "f \ V\<^sub>0" + then have "f \ \" + unfolding V0_def by simp + then obtain i where i: "f 0 \= i" by fastforce + define g where "g = (\x. f (x + 1))" + then have "g \ \" + using skip_R1[OF `f \ \`] by blast + moreover have "[i] \ g = f" + using g_def i by auto + moreover have "\ i = f" + using `f \ V\<^sub>0` V0_def i by force + ultimately show "f \ ?W" by auto + qed + show "?W \ V\<^sub>0" + proof + fix g + assume "g \ ?W" + then have "\ (the (g 0)) = g" by auto + moreover have "g \ \" + using prepend_in_R1 `g \ ?W` by auto + ultimately show "g \ V\<^sub>0" + by (simp add: V0_def) + qed +qed + +lemma V0_in_FIN: "V\<^sub>0 \ FIN" +proof - + define s where "s = (\x. Some (Suc (e_hd x)))" + have "s \ \

" + proof - + define r where "r = Cn 1 S [r_hd]" + then have "recfn 1 r" by simp + moreover have "eval r [x] \= Suc (e_hd x)" for x + unfolding r_def by simp + ultimately show ?thesis + using s_def by blast + qed + have s: "s (f \ n) \= Suc (the (f 0))" for f n + unfolding s_def by simp + have "learn_fin \ V\<^sub>0 s" + proof (rule learn_finI) + show "environment \ V\<^sub>0 s" + using s_def \s \ \

\ phi_in_P2 V0_def by auto + show "\i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + if "f \ V\<^sub>0" for f + using that V0_def s by auto + qed + then show ?thesis using FIN_def by auto +qed + +text \To every @{term "f \ \"} a number can be prepended that is +a Gödel number of the resulting function. Such a function is then in $V_0$. + +If $V_0$ was in NUM, it would be embedded in a total numbering. Shifting this +numbering to the left, essentially discarding the values at point $0$, would +yield a total numbering for @{term "\"}, which contradicts @{thm[source] +R1_not_in_NUM}. This proves @{prop "V\<^sub>0 \ NUM"}.\ + +lemma prepend_goedel: + assumes "f \ \" + shows "\i. \ i = [i] \ f" +proof - + obtain r where r: "recfn 1 r" "total r" "\x. eval r [x] = f x" + using assms by auto + define r_psi where "r_psi = Cn 2 r_ifz [Id 2 1, Id 2 0, Cn 2 r [Cn 2 r_dec [Id 2 1]]]" + then have "recfn 2 r_psi" + using r(1) by simp + have "eval r_psi [i, x] = (if x = 0 then Some i else f (x - 1))" for i x + proof - + have "eval (Cn 2 r [Cn 2 r_dec [Id 2 1]]) [i, x] = f (x - 1)" + using r by simp + then have "eval r_psi [i, x] = eval r_ifz [x, i, the (f (x - 1))]" + unfolding r_psi_def using \recfn 2 r_psi\ r R1_imp_total1[OF assms] by auto + then show ?thesis + using assms by simp + qed + with \recfn 2 r_psi\ have "(\i x. if x = 0 then Some i else f (x - 1)) \ \

\<^sup>2" + by auto + with kleene_fixed_point obtain i where + "\ i = (\x. if x = 0 then Some i else f (x - 1))" + by blast + then have "\ i = [i] \ f" by auto + then show ?thesis by auto +qed + +lemma V0_in_FIN_minus_NUM: "V\<^sub>0 \ FIN - NUM" +proof - + have "V\<^sub>0 \ NUM" + proof + assume "V\<^sub>0 \ NUM" + then obtain \ where \: "\ \ \\<^sup>2" "\f. f \ V\<^sub>0 \ \i. \ i = f" + by auto + define \' where "\' i x = \ i (Suc x)" for i x + have "\' \ \\<^sup>2" + proof + from \(1) obtain r_psi where + r_psi: "recfn 2 r_psi" "total r_psi" "\i x. eval r_psi [i, x] = \ i x" + by blast + define r_psi' where "r_psi' = Cn 2 r_psi [Id 2 0, Cn 2 S [Id 2 1]]" + then have "recfn 2 r_psi'" and "\i x. eval r_psi' [i, x] = \' i x" + unfolding r_psi'_def \'_def using r_psi by simp_all + then show "\' \ \

\<^sup>2" by blast + show "total2 \'" + using \'_def \(1) by (simp add: total2I) + qed + have "\i. \' i = f" if "f \ \" for f + proof - + from that obtain j where j: "\ j = [j] \ f" + using prepend_goedel by auto + then have "\ j \ V\<^sub>0" + using that V0_altdef by auto + with \ obtain i where "\ i = \ j" by auto + then have "\' i = f" + using \'_def j by (auto simp add: prepend_at_ge) + then show ?thesis by auto + qed + with \\' \ \\<^sup>2\ have "\ \ NUM" by auto + with R1_not_in_NUM show False by simp + qed + then show ?thesis + using V0_in_FIN by auto +qed + +corollary FIN_not_subseteq_NUM: "\ FIN \ NUM" + using V0_in_FIN_minus_NUM by auto + + +section \NUM and CP are incomparable\label{s:num_cp}\ + +text \There are FIN classes outside of NUM, and CP encompasses FIN. +Hence there are CP classes outside of NUM, too.\ + +theorem CP_not_subseteq_NUM: "\ CP \ NUM" + using FIN_subseteq_CP FIN_not_subseteq_NUM by blast + +text \Conversely there is a subclass of @{term "U\<^sub>0"} that +is in NUM but cannot be learned in a class-preserving way. The following +proof is due to Jantke and Beick~\cite{jb-cpnii-81}. The idea is to +diagonalize against all strategies, that is, all partial recursive +functions.\ + +theorem NUM_not_subseteq_CP: "\ NUM \ CP" +proof- + \ \Define a family of functions $f_k$.\ + define f where "f \ \k. [k] \ 0\<^sup>\" + then have "f k \ \" for k + using almost0_in_R1 by auto + + \ \If the strategy $\varphi_k$ learns $f_k$ it hypothesizes + $f_k$ for some shortest prefix $f_k^{a_k}$. Define functions $f'_k = + k0^{a_k}10^\infty$.\ + define a where + "a \ \k. LEAST x. (\ (the ((\ k) ((f k) \ x)))) = f k" + define f' where "f' \ \k. (k # (replicate (a k) 0) @ [1]) \ 0\<^sup>\" + then have "f' k \ \" for k + using almost0_in_R1 by auto + + \ \Although $f_k$ and $f'_k$ differ, they share the prefix of length $a_k + 1$.\ + have init_eq: "(f' k) \ (a k) = (f k) \ (a k)" for k + proof (rule init_eqI) + fix x assume "x \ a k" + then show "f' k x = f k x" + by (cases "x = 0") (simp_all add: nth_append f'_def f_def) + qed + have "f k \ f' k" for k + proof - + have "f k (Suc (a k)) \= 0" using f_def by auto + moreover have "f' k (Suc (a k)) \= 1" + using f'_def prepend[of "(k # (replicate (a k) 0) @ [1])" "0\<^sup>\" "Suc (a k)"] + by (metis length_Cons length_append_singleton length_replicate lessI nth_Cons_Suc + nth_append_length) + ultimately show ?thesis by auto + qed + + \ \The separating class $U$ contains $f'_k$ if $\varphi_k$ + learns $f_k$; otherwise it contains $f_k$.\ + define U where + "U \ {if learn_lim \ {f k} (\ k) then f' k else f k |k. k \ UNIV}" + have "U \ CP" + proof + assume "U \ CP" + have "\k. learn_cp \ U (\ k)" + proof - + have "\\ s. learn_cp \ U s" + using CP_def `U \ CP` by auto + then obtain s where s: "learn_cp \ U s" + using learn_cp_wrt_goedel[OF goedel_numbering_phi] by blast + then obtain k where "\ k = s" + using phi_universal learn_cp_def learn_lim_def by auto + then show ?thesis using s by auto + qed + then obtain k where k: "learn_cp \ U (\ k)" by auto + then have learn: "learn_lim \ U (\ k)" + using learn_cp_def by simp + \ \If $f_k$ was in $U$, $\varphi_k$ would learn it. But then, + by definition of $U$, $f_k$ would not be in $U$. Hence $f_k \notin U$.\ + have "f k \ U" + proof + assume "f k \ U" + then obtain m where m: "f k = (if learn_lim \ {f m} (\ m) then f' m else f m)" + using U_def by auto + have "f k 0 \= m" + using f_def f'_def m by simp + moreover have "f k 0 \= k" by (simp add: f_def) + ultimately have "m = k" by simp + with m have "f k = (if learn_lim \ {f k} (\ k) then f' k else f k)" + by auto + moreover have "learn_lim \ {f k} (\ k)" + using \f k \ U\ learn_lim_closed_subseteq[OF learn] by simp + ultimately have "f k = f' k" + by simp + then show False + using \f k \ f' k\ by simp + qed + then have "f' k \ U" using U_def by fastforce + then have in_U: "\n. \ (the ((\ k) ((f' k) \ n))) \ U" + using learn_cpE(3)[OF k] by simp + + \ \Since $f'_k \in U$, the strategy $\varphi_k$ learns $f_k$. + Then $a_k$ is well-defined, $f'^{a_k} = f^{a_k}$, and $\varphi_k$ + hypothesizes $f_k$ on $f'^{a_k}$, which is not a class-preserving + hypothesis.\ + have "learn_lim \ {f k} (\ k)" using U_def \f k \ U\ by fastforce + then have "\i n\<^sub>0. \ i = f k \ (\n\n\<^sub>0. \ k ((f k) \ n) \= i)" + using learn_limE(2) by simp + then obtain i n\<^sub>0 where "\ i = f k \ (\n\n\<^sub>0. \ k ((f k) \ n) \= i)" + by auto + then have "\ (the (\ k ((f k) \ (a k)))) = f k" + using a_def LeastI[of "\x. (\ (the ((\ k) ((f k) \ x)))) = f k" n\<^sub>0] + by simp + then have "\ (the ((\ k) ((f' k) \ (a k)))) = f k" + using init_eq by simp + then show False + using \f k \ U\ in_U by metis + qed + moreover have "U \ NUM" + using NUM_closed_subseteq[OF U0_in_NUM, of U] f_def f'_def U0_def U_def + by fastforce + ultimately show ?thesis by auto +qed + + +section \NUM is a proper subset of TOTAL\label{s:num_total}\ + +text \A NUM class $U$ is embedded in a total numbering @{term \}. +The strategy $S$ with $S(f^n) = \min \{i \mid \forall k \le n: \psi_i(k) = +f(k)\}$ for $f \in U$ converges to the least index of $f$ in @{term \}, +and thus learns $f$ in the limit. Moreover it will be a TOTAL strategy +because @{term \} contains only total functions. This shows @{prop "NUM +\ TOTAL"}.\ + +text \First we define, for every hypothesis space $\psi$, a +function that tries to determine for a given list $e$ and index $i$ whether +$e$ is a prefix of $\psi_i$. In other words it tries to decide whether $i$ is +a consistent hypothesis for $e$. ``Tries'' refers to the fact that the +function will diverge if $\psi_i(x)\uparrow$ for any $x \le |e|$. We start +with a version that checks the list only up to a given length.\ + +definition r_consist_upto :: "recf \ recf" where + "r_consist_upto r_psi \ + let g = Cn 4 r_ifeq + [Cn 4 r_psi [Id 4 2, Id 4 0], Cn 4 r_nth [Id 4 3, Id 4 0], Id 4 1, r_constn 3 1] + in Pr 2 (r_constn 1 0) g" + +lemma r_consist_upto_recfn: "recfn 2 r_psi \ recfn 3 (r_consist_upto r_psi)" + using r_consist_upto_def by simp + +lemma r_consist_upto: + assumes "recfn 2 r_psi" + shows "\k \ + eval (r_consist_upto r_psi) [j, i, e] = + (if \k= e_nth e k then Some 0 else Some 1)" + and "\ (\k) \ eval (r_consist_upto r_psi) [j, i, e] \" +proof - + define g where "g = + Cn 4 r_ifeq + [Cn 4 r_psi [Id 4 2, Id 4 0], Cn 4 r_nth [Id 4 3, Id 4 0], Id 4 1, r_constn 3 1]" + then have "recfn 4 g" + using assms by simp + moreover have "eval (Cn 4 r_nth [Id 4 3, Id 4 0]) [j, r, i, e] \= e_nth e j" for j r i e + by simp + moreover have "eval (r_constn 3 1) [j, r, i, e] \= 1" for j r i e + by simp + moreover have "eval (Cn 4 r_psi [Id 4 2, Id 4 0]) [j, r, i, e] = eval r_psi [i, j]" for j r i e + using assms(1) by simp + ultimately have g: "eval g [j, r, i, e] = + (if eval r_psi [i, j] \ then None + else if eval r_psi [i, j] \= e_nth e j then Some r else Some 1)" + for j r i e + using `recfn 4 g` g_def assms by auto + have goal1: "\k \ + eval (r_consist_upto r_psi) [j, i, e] = + (if \k= e_nth e k then Some 0 else Some 1)" + for j i e + proof (induction j) + case 0 + then show ?case + using r_consist_upto_def r_consist_upto_recfn assms eval_Pr_0 by simp + next + case (Suc j) + then have "eval (r_consist_upto r_psi) [Suc j, i, e] = + eval g [j, the (eval (r_consist_upto r_psi) [j, i, e]), i, e]" + using assms eval_Pr_converg_Suc g_def r_consist_upto_def r_consist_upto_recfn + by simp + also have "... = eval g [j, if \k= e_nth e k then 0 else 1, i, e]" + using Suc by auto + also have "... \= (if eval r_psi [i, j] \= e_nth e j + then if \k= e_nth e k then 0 else 1 else 1)" + using g by (simp add: Suc.prems) + also have "... \= (if \k= e_nth e k then 0 else 1)" + by (simp add: less_Suc_eq) + finally show ?case by simp + qed + then show "\k \ + eval (r_consist_upto r_psi) [j, i, e] = + (if \k= e_nth e k then Some 0 else Some 1)" + by simp + show "\ (\k) \ eval (r_consist_upto r_psi) [j, i, e] \" + proof - + assume "\ (\k)" + then have "\k" by simp + let ?P = "\k. k < j \ eval r_psi [i, k] \" + define kmin where "kmin = Least ?P" + then have "?P kmin" + using LeastI_ex[of ?P] \\k\ by auto + from kmin_def have "\k. k < kmin \ \ ?P k" + using kmin_def not_less_Least[of _ ?P] by blast + then have "\k < kmin. eval r_psi [i, k] \" + using `?P kmin` by simp + then have "eval (r_consist_upto r_psi) [kmin, i, e] = + (if \k= e_nth e k then Some 0 else Some 1)" + using goal1 by simp + moreover have "eval r_psi [i, kmin] \" + using `?P kmin` by simp + ultimately have "eval (r_consist_upto r_psi) [Suc kmin, i, e] \" + using r_consist_upto_def g assms by simp + moreover have "j \ kmin" + using `?P kmin` by simp + ultimately show "eval (r_consist_upto r_psi) [j, i, e] \" + using r_consist_upto_def r_consist_upto_recfn `?P kmin` eval_Pr_converg_le assms + by (metis (full_types) Suc_leI length_Cons list.size(3) numeral_2_eq_2 numeral_3_eq_3) + qed +qed + +text \The next function provides the consistency decision functions we +need.\ + +definition consistent :: "partial2 \ partial2" where + "consistent \ i e \ + if \k i k \ + then if \k i k \= e_nth e k + then Some 0 else Some 1 + else None" + +text \Given $i$ and $e$, @{term "consistent \"} decides whether $e$ +is a prefix of $\psi_i$, provided $\psi_i$ is defined for the length of +$e$.\ + +definition r_consistent :: "recf \ recf" where + "r_consistent r_psi \ + Cn 2 (r_consist_upto r_psi) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]" + +lemma r_consistent_recfn [simp]: "recfn 2 r_psi \ recfn 2 (r_consistent r_psi)" + using r_consistent_def r_consist_upto_recfn by simp + +lemma r_consistent_converg: + assumes "recfn 2 r_psi" and "\k" + shows "eval (r_consistent r_psi) [i, e] \= + (if \k= e_nth e k then 0 else 1)" +proof - + have "eval (r_consistent r_psi) [i, e] = eval (r_consist_upto r_psi) [e_length e, i, e]" + using r_consistent_def r_consist_upto_recfn assms(1) by simp + then show ?thesis using assms r_consist_upto(1) by simp +qed + +lemma r_consistent_diverg: + assumes "recfn 2 r_psi" and "\k" + shows "eval (r_consistent r_psi) [i, e] \" + unfolding r_consistent_def + using r_consist_upto_recfn[OF assms(1)] r_consist_upto[OF assms(1)] assms(2) + by simp + +lemma r_consistent: + assumes "recfn 2 r_psi" and "\x y. eval r_psi [x, y] = \ x y" + shows "eval (r_consistent r_psi) [i, e] = consistent \ i e" +proof (cases "\k i k \") + case True + then have "\k" + using assms by simp + then show ?thesis + unfolding consistent_def using True by (simp add: assms r_consistent_converg) +next + case False + then have "consistent \ i e \" + unfolding consistent_def by auto + moreover have "eval (r_consistent r_psi) [i, e] \" + using r_consistent_diverg[OF assms(1)] assms False by simp + ultimately show ?thesis by simp +qed + +lemma consistent_in_P2: + assumes "\ \ \

\<^sup>2" + shows "consistent \ \ \

\<^sup>2" + using assms r_consistent P2E[OF assms(1)] P2I r_consistent_recfn by metis + +lemma consistent_for_R2: + assumes "\ \ \\<^sup>2" + shows "consistent \ i e = + (if \j i j \= e_nth e j then Some 0 else Some 1)" + using assms by (simp add: consistent_def) + +lemma consistent_init: + assumes "\ \ \\<^sup>2" and "f \ \" + shows "consistent \ i (f \ n) = (if \ i \ n = f \ n then Some 0 else Some 1)" + using consistent_def[of _ _ "init f n"] assms init_eq_iff_eq_upto by simp + +lemma consistent_in_R2: + assumes "\ \ \\<^sup>2" + shows "consistent \ \ \\<^sup>2" + using total2I consistent_in_P2 consistent_for_R2[OF assms] P2_total_imp_R2 R2_imp_P2 assms + by (metis option.simps(3)) + +text \For total hypothesis spaces the next function computes the +minimum hypothesis consistent with a given prefix. It diverges if no such +hypothesis exists.\ + +definition min_cons_hyp :: "partial2 \ partial1" where + "min_cons_hyp \ e \ + if \i. consistent \ i e \= 0 then Some (LEAST i. consistent \ i e \= 0) else None" + +lemma min_cons_hyp_in_P1: + assumes "\ \ \\<^sup>2" + shows "min_cons_hyp \ \ \

" +proof - + from assms consistent_in_R2 obtain rc where + rc: "recfn 2 rc" "total rc" "\i e. eval rc [i, e] = consistent \ i e" + using R2E[of "consistent \"] by metis + define r where "r = Mn 1 rc" + then have "recfn 1 r" + using rc(1) by simp + moreover from this have "eval r [e] = min_cons_hyp \ e" for e + using r_def eval_Mn'[of 1 rc "[e]"] rc min_cons_hyp_def assms + by (auto simp add: consistent_in_R2) + ultimately show ?thesis by auto +qed + +text \The function @{term "min_cons_hyp \"} is a strategy for +learning all NUM classes embedded in @{term \}. It is an example of an +``identification-by-enumeration'' strategy.\ + +lemma NUM_imp_learn_total: + assumes "\ \ \\<^sup>2" and "U \ NUM_wrt \" + shows "learn_total \ U (min_cons_hyp \)" +proof (rule learn_totalI) + have ex_psi_i_f: "\i. \ i = f" if "f \ U" for f + using assms that NUM_wrt_def by simp + moreover have consistent_eq_0: "consistent \ i ((\ i) \ n) \= 0" for i n + using assms by (simp add: consistent_init) + ultimately have "\f n. f \ U \ min_cons_hyp \ (f \ n) \" + using min_cons_hyp_def assms(1) by fastforce + then show env: "environment \ U (min_cons_hyp \)" + using assms NUM_wrt_def min_cons_hyp_in_P1 NUM_E(1) NUM_I by auto + + show "\f n. f \ U \ \ (the (min_cons_hyp \ (f \ n))) \ \" + using assms by (simp) + + show "\i. \ i = f \ (\\<^sup>\n. min_cons_hyp \ (f \ n) \= i)" if "f \ U" for f + proof - + from that env have "f \ \" by auto + + let ?P = "\i. \ i = f" + define imin where "imin \ Least ?P" + with ex_psi_i_f that have imin: "?P imin" "\j. ?P j \ j \ imin" + using LeastI_ex[of ?P] Least_le[of ?P] by simp_all + then have f_neq: "\ i \ f" if "i < imin" for i + using leD that by auto + + let ?Q = "\i n. \ i \ n \ f \ n" + define nu :: "nat \ nat" where "nu = (\i. SOME n. ?Q i n)" + have nu_neq: "\ i \ (nu i) \ f \ (nu i)" if "i < imin" for i + proof - + from assms have "\ i \ \" by simp + moreover from assms imin(1) have "f \ \" by auto + moreover have "f \ \ i" + using that f_neq by auto + ultimately have "\n. f \ n \ (\ i) \ n" + using neq_fun_neq_init by simp + then show "?Q i (nu i)" + unfolding nu_def using someI_ex[of "\n. ?Q i n"] by metis + qed + + have "\n\<^sub>0. \n\n\<^sub>0. min_cons_hyp \ (f \ n) \= imin" + proof (cases "imin = 0") + case True + then have "\n. min_cons_hyp \ (f \ n) \= imin" + using consistent_eq_0 assms(1) imin(1) min_cons_hyp_def by auto + then show ?thesis by simp + next + case False + define n\<^sub>0 where "n\<^sub>0 = Max (set (map nu [0.. n\<^sub>0" if "i < imin" for i + proof - + have "finite ?N" + using n\<^sub>0_def by simp + moreover have "?N \ {}" + using False n\<^sub>0_def by simp + moreover have "nu i \ ?N" + using that by simp + ultimately show ?thesis + using that Max_ge n\<^sub>0_def by blast + qed + then have "\ i \ n\<^sub>0 \ f \ n\<^sub>0" if "i < imin" for i + using nu_neq neq_init_forall_ge that by blast + then have *: "\ i \ n \ f \ n" if "i < imin" and "n \ n\<^sub>0" for i n + using nu_neq neq_init_forall_ge that by blast + + have "\ imin \ n = f \ n" for n + using imin(1) by simp + moreover have "(consistent \ i (f \ n) \= 0) = (\ i \ n = f \ n)" for i n + by (simp add: \f \ \\ assms(1) consistent_init) + ultimately have "min_cons_hyp \ (f \ n) \= (LEAST i. \ i \ n = f \ n)" for n + using min_cons_hyp_def[of \ "f \ n"] by auto + moreover have "(LEAST i. \ i \ n = f \ n) = imin" if "n \ n\<^sub>0" for n + proof (rule Least_equality) + show "\ imin \ n = f \ n" + using imin(1) by simp + show "\y. \ y \ n = f \ n \ imin \ y" + using imin * leI that by blast + qed + ultimately have "min_cons_hyp \ (f \ n) \= imin" if "n \ n\<^sub>0" for n + using that by blast + then show ?thesis by auto + qed + with imin(1) show ?thesis by auto + qed +qed + +corollary NUM_subseteq_TOTAL: "NUM \ TOTAL" +proof + fix U + assume "U \ NUM" + then have "\\\\\<^sup>2. \f\U. \i. \ i = f" by auto + then have "\\\\\<^sup>2. U \ NUM_wrt \" + using NUM_wrt_def by simp + then have "\\ s. learn_total \ U s" + using NUM_imp_learn_total by auto + then show "U \ TOTAL" + using TOTAL_def by auto +qed + +text \The class @{term V0} is in @{term "TOTAL - NUM"}. \ + +theorem NUM_subset_TOTAL: "NUM \ TOTAL" + using CP_subseteq_TOTAL FIN_not_subseteq_NUM FIN_subseteq_CP NUM_subseteq_TOTAL + by auto + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/Inductive_Inference_Basics.thy b/thys/Inductive_Inference/Inductive_Inference_Basics.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/Inductive_Inference_Basics.thy @@ -0,0 +1,1222 @@ +chapter \Inductive inference of recursive functions\label{c:iirf}\ + +theory Inductive_Inference_Basics + imports Standard_Results +begin + +text \Inductive inference originates from work by +Solomonoff~\cite{s-ftiip1-64,s-ftiip2-64} and Gold~\cite{g-lil-67,g-lr-65} +and comes in many variations. The common theme is to infer additional +information about objects, such as formal languages or functions, from incomplete +data, such as finitely many words contained in the language or argument-value +pairs of the function. Oftentimes ``additional information'' means complete +information, such that the task becomes identification of the object. + +The basic setting in inductive inference of recursive functions is as follows. +Let us denote, for a total function $f$, by $f^n$ the code of the list +$[f(0), ..., f(n)]$. Let $U$ be a set (called \emph{class}) of total +recursive functions, and $\psi$ a binary partial recursive function +(called \emph{hypothesis space}). +A partial recursive function $S$ (called \emph{strategy}) +is said to \emph{learn $U$ in the limit with respect to $\psi$} if +for all $f \in U$, +\begin{itemize} + \item the value $S(f^n)$ is defined for all $n\in\mathbb{N}$, + \item the sequence $S(f^0), S(f^1), \ldots$ converges to an + $i\in\mathbb{N}$ with $\psi_i = f$. +\end{itemize} + +Both the output $S(f^n)$ of the strategy and its interpretation +as a function $\psi_{S(f^n)}$ are called \emph{hypothesis}. The set +of all classes learnable in the limit by $S$ with respect to $\psi$ is +denoted by $\mathrm{LIM}_\psi(S)$. Moreover we set $\mathrm{LIM}_\psi = +\bigcup_{S\in\mathcal{P}} \mathrm{LIM}_\psi(S)$ and $\mathrm{LIM} = +\bigcup_{\psi\in\mathcal{P}^2} \mathrm{LIM}_\psi$. We call the latter set the +\emph{inference type} $\mathrm{LIM}$. + +Many aspects of this setting can be varied. We shall consider: +\begin{itemize} + \item Intermediate hypotheses: $\psi_{S(f^n)}$ can be required to be total or + to be in the class $U$, or to coincide with $f$ on arguments up to $n$, or + a myriad of other conditions or combinations thereof. + \item Convergence of hypotheses: + \begin{itemize} + \item The strategy can be required to output not a sequence but a single + hypothesis, which must be correct. + \item The strategy can be required to converge to a \emph{function} rather + than an index. + \end{itemize} +\end{itemize} + +We formalize five kinds of results (\\\ and \\'\ stand for +inference types): +\begin{itemize} + \item Comparison of learning power: results of the form @{prop "\ + \ \'"}, in particular showing that the inclusion is proper + (Sections~\ref{s:fin_cp}, \ref{s:num_fin}, \ref{s:num_cp}, + \ref{s:num_total}, \ref{s:cons_lim}, \ref{s:lim_bc}, \ref{s:total_cons}, + \ref{s:r1_bc}). + \item Whether \\\ is closed under the subset relation: @{prop "U + \ \ \ V \ U \ V \ \"}. + \item Whether \\\ is closed under union: @{prop "U \ \ \ + V \ \ \ U \ V \ \"} (Section~\ref{s:union}). + \item Whether every class in \\\ can be learned with respect to a + Gödel numbering as hypothesis space (Section~\ref{s:inference_types}). + \item Whether every class in \\\ can be learned by a \emph{total} + recursive strategy (Section~\ref{s:lemma_r}). +\end{itemize} + +The bulk of this chapter is devoted to the first category of results. Most +results that we are going to formalize have been called ``classical'' by +Jantke and Beick~\cite{jb-cpnii-81}, who compare a large number of inference +types. Another comparison is by Case and Smith~\cite{cs-cicmii-83}. Angluin +and Smith~\cite{as-ii-87} give an overview of various forms of inductive +inference. + +All (interesting) proofs herein are based on my lecture notes of the +\emph{Induktive Inferenz} lectures by Rolf Wiehagen from 1999/2000 and +2000/2001 at the University of Kaiserslautern. I have given references to the +original proofs whenever I was able to find them. For the other proofs, as +well as for those that I had to contort beyond recognition, I provide proof +sketches.\ + + +section \Preliminaries\ + +text \Throughout the chapter, in particular in proof sketches, we use +the following notation. + +Let $b\in\mathbb{N}^*$ be a list of numbers. We write $|b|$ for its length +and $b_i$ for the $i$-th element ($i=0,\dots, |b| - 1$). Concatenation of +numbers and lists works in the obvious way; for instance, $jbk$ with +$j,k\in\mathbb{N}$, $b\in\mathbb{N}^*$ refers to the list $jb_0\dots +b_{|b|-1}k$. For $0 \leq i < |b|$, the term $b_{i:=v}$ denotes the list +$b_0\dots b_{i-1}vb_{i+1}\dots b_{|b|-1}$. The notation $b_{ + + +subsection \The prefixes of a function\ + +text \A \emph{prefix}, also called \emph{initial segment}, is a list of +initial values of a function.\ + +definition prefix :: "partial1 \ nat \ nat list" where + "prefix f n \ map (\x. the (f x)) [0..x. the (f x)"] by simp + +lemma prefixI: + assumes "length vs > 0" and "\x. x < length vs \ f x \= vs ! x" + shows "prefix f (length vs - 1) = vs" + using assms nth_equalityI[of "prefix f (length vs - 1)" vs] by simp + +lemma prefixI': + assumes "length vs = Suc n" and "\x. x < Suc n \ f x \= vs ! x" + shows "prefix f n = vs" + using assms nth_equalityI[of "prefix f (length vs - 1)" vs] by simp + +lemma prefixE: + assumes "prefix f (length vs - 1) = vs" + and "f \ \" + and "length vs > 0" + and "x < length vs" + shows "f x \= vs ! x" + using assms length_prefix prefix_nth[of x "length vs - 1" f] by simp + +lemma prefix_eqI: + assumes "\x. x \ n \ f x = g x" + shows "prefix f n = prefix g n" + using assms prefix_def by simp + +lemma prefix_0: "prefix f 0 = [the (f 0)]" + using prefix_def by simp + +lemma prefix_Suc: "prefix f (Suc n) = prefix f n @ [the (f (Suc n))]" + unfolding prefix_def by simp + +lemma take_prefix: + assumes "f \ \" and "k \ n" + shows "prefix f k = take (Suc k) (prefix f n)" +proof - + let ?vs = "take (Suc k) (prefix f n)" + have "length ?vs = Suc k" + using assms(2) by simp + then have "\x. x < length ?vs \ f x \= ?vs ! x" + using assms by auto + then show ?thesis + using prefixI[where ?vs="?vs"] `length ?vs = Suc k` by simp +qed + +text \Strategies receive prefixes in the form of encoded lists. The +term ``prefix'' refers to both encoded and unencoded lists. We use the +notation @{text "f \ n"} for the prefix $f^n$.\ + +definition init :: "partial1 \ nat \ nat" (infix "\" 110) where + "f \ n \ list_encode (prefix f n)" + +lemma init_neq_zero: "f \ n \ 0" + unfolding init_def prefix_def using list_encode_0 by fastforce + +lemma init_prefixE [elim]: "prefix f n = prefix g n \ f \ n = g \ n" + unfolding init_def by simp + +lemma init_eqI: + assumes "\x. x \ n \ f x = g x" + shows "f \ n = g \ n" + unfolding init_def using prefix_eqI[OF assms] by simp + +lemma initI: + assumes "e_length e > 0" and "\x. x < e_length e \ f x \= e_nth e x" + shows "f \ (e_length e - 1) = e" + unfolding init_def using assms prefixI by simp + +lemma initI': + assumes "e_length e = Suc n" and "\x. x < Suc n \ f x \= e_nth e x" + shows "f \ n = e" + unfolding init_def using assms prefixI' by simp + +lemma init_iff_list_eq_upto: + assumes "f \ \" and "e_length vs > 0" + shows "(\x= e_nth vs x) \ prefix f (e_length vs - 1) = list_decode vs" + using prefixI[OF assms(2)] prefixE[OF _ assms] by auto + +lemma length_init [simp]: "e_length (f \ n) = Suc n" + unfolding init_def by simp + +lemma init_Suc_snoc: "f \ (Suc n) = e_snoc (f \ n) (the (f (Suc n)))" + unfolding init_def by (simp add: prefix_Suc) + +lemma nth_init: "i < Suc n \ e_nth (f \ n) i = the (f i)" + unfolding init_def using prefix_nth by auto + +lemma hd_init [simp]: "e_hd (f \ n) = the (f 0)" + unfolding init_def using init_neq_zero by (simp add: e_hd_nth0) + +lemma list_decode_init [simp]: "list_decode (f \ n) = prefix f n" + unfolding init_def by simp + +lemma init_eq_iff_eq_upto: + assumes "g \ \" and "f \ \" + shows "(\j g \ n = f \ n" + using assms initI' init_iff_list_eq_upto length_init list_decode_init + by (metis diff_Suc_1 zero_less_Suc) + +definition is_init_of :: "nat \ partial1 \ bool" where + "is_init_of t f \ \i= e_nth t i" + +lemma not_initial_imp_not_eq: + assumes "\x. x < Suc n \ f x \" and "\ (is_init_of (f \ n) g)" + shows "f \ g" + using is_init_of_def assms by auto + +lemma all_init_eq_imp_fun_eq: + assumes "f \ \" and "g \ \" and "\n. f \ n = g \ n" + shows "f = g" +proof + fix n + from assms have "prefix f n = prefix g n" + by (metis init_def list_decode_encode) + then have "the (f n) = the (g n)" + unfolding init_def prefix_def by simp + then show "f n = g n" + using assms(1,2) by (meson R1_imp_total1 option.expand total1E) +qed + +corollary neq_fun_neq_init: + assumes "f \ \" and "g \ \" and "f \ g" + shows "\n. f \ n \ g \ n" + using assms all_init_eq_imp_fun_eq by auto + +lemma eq_init_forall_le: + assumes "f \ n = g \ n" and "m \ n" + shows "f \ m = g \ m" +proof - + from assms(1) have "prefix f n = prefix g n" + by (metis init_def list_decode_encode) + then have "the (f k) = the (g k)" if "k \ n" for k + using prefix_def that by auto + then have "the (f k) = the (g k)" if "k \ m" for k + using assms(2) that by simp + then have "prefix f m = prefix g m" + using prefix_def by simp + then show ?thesis by (simp add: init_def) +qed + +corollary neq_init_forall_ge: + assumes "f \ n \ g \ n" and "m \ n" + shows "f \ m \ g \ m" + using eq_init_forall_le assms by blast + +lemma e_take_init: + assumes "f \ \" and "k < Suc n" + shows "e_take (Suc k) (f \ n) = f \ k" + using assms take_prefix by (simp add: init_def less_Suc_eq_le) + +lemma init_butlast_init: + assumes "total1 f" and "f \ n = e" and "n > 0" + shows "f \ (n - 1) = e_butlast e" +proof - + let ?e = "e_butlast e" + have "e_length e = Suc n" + using assms(2) by auto + then have len: "e_length ?e = n" + by simp + have "f \ (e_length ?e - 1) = ?e" + proof (rule initI) + show "0 < e_length ?e" + using assms(3) len by simp + have "\x. x < e_length e \ f x \= e_nth e x" + using assms(1,2) total1_def \e_length e = Suc n\ by auto + then show "\x. x < e_length ?e \ f x \= e_nth ?e x" + by (simp add: butlast_conv_take) + qed + with len show ?thesis by simp +qed + +text \Some definitions make use of recursive predicates, that is, +$01$-valued functions.\ + +definition RPred1 :: "partial1 set" ("\\<^sub>0\<^sub>1") where + "\\<^sub>0\<^sub>1 \ {f. f \ \ \ (\x. f x \= 0 \ f x \= 1)}" + +lemma RPred1_subseteq_R1: "\\<^sub>0\<^sub>1 \ \" + unfolding RPred1_def by auto + +lemma const0_in_RPred1: "(\_. Some 0) \ \\<^sub>0\<^sub>1" + using RPred1_def const_in_Prim1 by fast + +lemma RPred1_altdef: "\\<^sub>0\<^sub>1 = {f. f \ \ \ (\x. the (f x) \ 1)}" + (is "\\<^sub>0\<^sub>1 = ?S") +proof + show "\\<^sub>0\<^sub>1 \ ?S" + proof + fix f + assume f: "f \ \\<^sub>0\<^sub>1" + with RPred1_def have "f \ \" by auto + from f have "\x. f x \= 0 \ f x \= 1" + by (simp add: RPred1_def) + then have "\x. the (f x) \ 1" + by (metis eq_refl less_Suc_eq_le zero_less_Suc option.sel) + with `f \ \` show "f \ ?S" by simp + qed + show "?S \ \\<^sub>0\<^sub>1" + proof + fix f + assume f: "f \ ?S" + then have "f \ \" by simp + then have total: "\x. f x \" by auto + from f have "\x. the (f x) = 0 \ the (f x) = 1" + by (simp add: le_eq_less_or_eq) + with total have "\x. f x \= 0 \ f x \= 1" + by (metis option.collapse) + then show "f \ \\<^sub>0\<^sub>1" + using `f \ \` RPred1_def by auto + qed +qed + +subsection \NUM\ + +text \A class of recursive functions is in NUM if it can be +embedded in a total numbering. Thus, for learning such classes there is +always a total hypothesis space available.\ + +definition NUM :: "partial1 set set" where + "NUM \ {U. \\\\\<^sup>2. \f\U. \i. \ i = f}" + +definition NUM_wrt :: "partial2 \ partial1 set set" where + "\ \ \\<^sup>2 \ NUM_wrt \ \ {U. \f\U. \i. \ i = f}" + +lemma NUM_I [intro]: + assumes "\ \ \\<^sup>2" and "\f. f \ U \ \i. \ i = f" + shows "U \ NUM" + using assms NUM_def by blast + +lemma NUM_E [dest]: + assumes "U \ NUM" + shows "U \ \" + and "\\\\\<^sup>2. \f\U. \i. \ i = f" + using NUM_def assms by (force, auto) + +lemma NUM_closed_subseteq: + assumes "U \ NUM" and "V \ U" + shows "V \ NUM" + using assms subset_eq[of V U] NUM_I by auto + +text \This is the classical diagonalization proof showing that there is +no total numbering containing all total recursive functions.\ + +lemma R1_not_in_NUM: "\ \ NUM" +proof + assume "\ \ NUM" + then obtain \ where num: "\ \ \\<^sup>2" "\f\\. \i. \ i = f" + by auto + then obtain psi where psi: "recfn 2 psi" "total psi" "eval psi [i, x] = \ i x" for i x + by auto + define d where "d = Cn 1 S [Cn 1 psi [Id 1 0, Id 1 0]]" + then have "recfn 1 d" + using psi(1) by simp + moreover have d: "eval d [x] \= Suc (the (\ x x))" for x + unfolding d_def using num psi by simp + ultimately have "(\x. eval d [x]) \ \" + using R1I by blast + then obtain i where "\ i = (\x. eval d [x])" + using num(2) by auto + then have "\ i i = eval d [i]" by simp + with d have "\ i i \= Suc (the (\ i i))" by simp + then show False + using option.sel[of "Suc (the (\ i i))"] by simp +qed + +text \A hypothesis space that contains a function for every prefix will +come in handy. The following is a total numbering with this property.\ + +definition "r_prenum \ + Cn 2 r_ifless [Id 2 1, Cn 2 r_length [Id 2 0], Cn 2 r_nth [Id 2 0, Id 2 1], r_constn 1 0]" + +lemma r_prenum_prim [simp]: "prim_recfn 2 r_prenum" + unfolding r_prenum_def by simp_all + +lemma r_prenum [simp]: + "eval r_prenum [e, x] \= (if x < e_length e then e_nth e x else 0)" + by (simp add: r_prenum_def) + +definition prenum :: partial2 where + "prenum e x \ Some (if x < e_length e then e_nth e x else 0)" + +lemma prenum_in_R2: "prenum \ \\<^sup>2" + using prenum_def Prim2I[OF r_prenum_prim, of prenum] by simp + +lemma prenum [simp]: "prenum e x \= (if x < e_length e then e_nth e x else 0)" + unfolding prenum_def .. + +lemma prenum_encode: + "prenum (list_encode vs) x \= (if x < length vs then vs ! x else 0)" + using prenum_def by (cases "x < length vs") simp_all + +text \Prepending a list of numbers to a function:\ + +definition prepend :: "nat list \ partial1 \ partial1" (infixr "\" 64) where + "vs \ f \ \x. if x < length vs then Some (vs ! x) else f (x - length vs)" + +lemma prepend [simp]: + "(vs \ f) x = (if x < length vs then Some (vs ! x) else f (x - length vs))" + unfolding prepend_def .. + +lemma prepend_total: "total1 f \ total1 (vs \ f)" + unfolding total1_def by simp + +lemma prepend_at_less: + assumes "n < length vs" + shows "(vs \ f) n \= vs ! n" + using assms by simp + +lemma prepend_at_ge: + assumes "n \ length vs" + shows "(vs \ f) n = f (n - length vs)" + using assms by simp + +lemma prefix_prepend_less: + assumes "n < length vs" + shows "prefix (vs \ f) n = take (Suc n) vs" + using assms length_prefix by (intro nth_equalityI) simp_all + +lemma prepend_eqI: + assumes "\x. x < length vs \ g x \= vs ! x" + and "\x. g (length vs + x) = f x" + shows "g = vs \ f" +proof + fix x + show "g x = (vs \ f) x" + proof (cases "x < length vs") + case True + then show ?thesis using assms by simp + next + case False + then show ?thesis + using assms prepend by (metis add_diff_inverse_nat) + qed +qed + +fun r_prepend :: "nat list \ recf \ recf" where + "r_prepend [] r = r" +| "r_prepend (v # vs) r = + Cn 1 (r_lifz (r_const v) (Cn 1 (r_prepend vs r) [r_dec])) [Id 1 0, Id 1 0]" + +lemma r_prepend_recfn: + assumes "recfn 1 r" + shows "recfn 1 (r_prepend vs r)" + using assms by (induction vs) simp_all + +lemma r_prepend: + assumes "recfn 1 r" + shows "eval (r_prepend vs r) [x] = + (if x < length vs then Some (vs ! x) else eval r [x - length vs])" +proof (induction vs arbitrary: x) + case Nil + then show ?case using assms by simp +next + case (Cons v vs) + show ?case + using assms Cons by (cases "x = 0") (auto simp add: r_prepend_recfn) +qed + +lemma r_prepend_total: + assumes "recfn 1 r" and "total r" + shows "eval (r_prepend vs r) [x] \= + (if x < length vs then vs ! x else the (eval r [x - length vs]))" +proof (induction vs arbitrary: x) + case Nil + then show ?case using assms by simp +next + case (Cons v vs) + show ?case + using assms Cons by (cases "x = 0") (auto simp add: r_prepend_recfn) +qed + +lemma prepend_in_P1: + assumes "f \ \

" + shows "vs \ f \ \

" +proof - + obtain r where r: "recfn 1 r" "\x. eval r [x] = f x" + using assms by auto + moreover have "recfn 1 (r_prepend vs r)" + using r r_prepend_recfn by simp + moreover have "eval (r_prepend vs r) [x] = (vs \ f) x" for x + using r r_prepend by simp + ultimately show ?thesis by blast +qed + +lemma prepend_in_R1: + assumes "f \ \" + shows "vs \ f \ \" +proof - + obtain r where r: "recfn 1 r" "total r" "\x. eval r [x] = f x" + using assms by auto + then have "total1 f" + using R1_imp_total1[OF assms] by simp + have "total (r_prepend vs r)" + using r r_prepend_total r_prepend_recfn totalI1[of "r_prepend vs r"] by simp + with r have "total (r_prepend vs r)" by simp + moreover have "recfn 1 (r_prepend vs r)" + using r r_prepend_recfn by simp + moreover have "eval (r_prepend vs r) [x] = (vs \ f) x" for x + using r r_prepend `total1 f` total1E by simp + ultimately show ?thesis by auto +qed + +lemma prepend_associative: "(us @ vs) \ f = us \ vs \ f" (is "?lhs = ?rhs") +proof + fix x + consider + "x < length us" + | "x \ length us \ x < length (us @ vs)" + | "x \ length (us @ vs)" + by linarith + then show "?lhs x = ?rhs x" + proof (cases) + case 1 + then show ?thesis + by (metis le_add1 length_append less_le_trans nth_append prepend_at_less) + next + case 2 + then show ?thesis + by (smt add_diff_inverse_nat add_less_cancel_left length_append nth_append prepend) + next + case 3 + then show ?thesis + using prepend_at_ge by auto + qed +qed + +abbreviation constant_divergent :: partial1 ("\\<^sup>\") where + "\\<^sup>\ \ \_. None" + +abbreviation constant_zero :: partial1 ("0\<^sup>\") where + "0\<^sup>\ \ \_. Some 0" + +lemma almost0_in_R1: "vs \ 0\<^sup>\ \ \" + using RPred1_subseteq_R1 const0_in_RPred1 prepend_in_R1 by auto + +text \The class $U_0$ of all total recursive functions that are almost +everywhere zero will be used several times to construct +(counter-)examples.\ + +definition U0 :: "partial1 set" ("U\<^sub>0") where + "U\<^sub>0 \ {vs \ 0\<^sup>\ |vs. vs \ UNIV}" + +text \The class @{term U0} contains exactly the functions in the +numbering @{term prenum}.\ + +lemma U0_altdef: "U\<^sub>0 = {prenum e| e. e \ UNIV}" (is "U\<^sub>0 = ?W") +proof + show "U\<^sub>0 \ ?W" + proof + fix f + assume "f \ U\<^sub>0" + with U0_def obtain vs where "f = vs \ 0\<^sup>\" + by auto + then have "f = prenum (list_encode vs)" + using prenum_encode by auto + then show "f \ ?W" by auto + qed + show "?W \ U\<^sub>0" + unfolding U0_def by fastforce +qed + +lemma U0_in_NUM: "U\<^sub>0 \ NUM" + using prenum_in_R2 U0_altdef by (intro NUM_I[of prenum]; force) + +text \Every almost-zero function can be represented by $v0^\infty$ for +a list $v$ not ending in zero.\ + +lemma almost0_canonical: + assumes "f = vs \ 0\<^sup>\" and "f \ 0\<^sup>\" + obtains ws where "length ws > 0" and "last ws \ 0" and "f = ws \ 0\<^sup>\" +proof - + let ?P = "\k. k < length vs \ vs ! k \ 0" + from assms have "vs \ []" + by auto + then have ex: "\k 0" + using assms by auto + define m where "m = Greatest ?P" + moreover have le: "\y. ?P y \ y \ length vs" + by simp + ultimately have "?P m" + using ex GreatestI_ex_nat[of ?P "length vs"] by simp + have not_gr: "\ ?P k" if "k > m" for k + using Greatest_le_nat[of ?P _ "length vs"] m_def ex le not_less that by blast + let ?ws = "take (Suc m) vs" + have "vs \ 0\<^sup>\ = ?ws \ 0\<^sup>\" + proof + fix x + show "(vs \ 0\<^sup>\) x = (?ws \ 0\<^sup>\) x" + proof (cases "x < Suc m") + case True + then show ?thesis using `?P m` by simp + next + case False + moreover from this have "(?ws \ 0\<^sup>\) x \= 0" + by simp + ultimately show ?thesis + using not_gr by (cases "x < length vs") simp_all + qed + qed + then have "f = ?ws \ 0\<^sup>\" + using assms(1) by simp + moreover have "length ?ws > 0" + by (simp add: \vs \ []\) + moreover have "last ?ws \ 0" + by (simp add: \?P m\ take_Suc_conv_app_nth) + ultimately show ?thesis using that by blast +qed + + +section \Types of inference\label{s:inference_types}\ + +text \This section introduces all inference types that we are going to +consider together with some of their simple properties. All these inference +types share the following condition, which essentially says that everything +must be computable:\ + +abbreviation environment :: "partial2 \ (partial1 set) \ partial1 \ bool" where + "environment \ U s \ \ \ \

\<^sup>2 \ U \ \ \ s \ \

\ (\f\U. \n. s (f \ n) \)" + + +subsection \LIM: Learning in the limit\ + +text \A strategy $S$ learns a class $U$ in the limit with respect to a +hypothesis space @{term "\ \ \

\<^sup>2"} if for all $f\in U$, the +sequence $(S(f^n))_{n\in\mathbb{N}}$ converges to an $i$ with $\psi_i = f$. +Convergence for a sequence of natural numbers means that almost all elements +are the same. We express this with the following notation.\ + +abbreviation Almost_All :: "(nat \ bool) \ bool" (binder "\\<^sup>\" 10) where + "\\<^sup>\n. P n \ \n\<^sub>0. \n\n\<^sub>0. P n" + +definition learn_lim :: "partial2 \ (partial1 set) \ partial1 \ bool" where + "learn_lim \ U s \ + environment \ U s \ + (\f\U. \i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i))" + +lemma learn_limE: + assumes "learn_lim \ U s" + shows "environment \ U s" + and "\f. f \ U \ \i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i)" + using assms learn_lim_def by auto + +lemma learn_limI: + assumes "environment \ U s" + and "\f. f \ U \ \i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i)" + shows "learn_lim \ U s" + using assms learn_lim_def by auto + +definition LIM_wrt :: "partial2 \ partial1 set set" where + "LIM_wrt \ \ {U. \s. learn_lim \ U s}" + +definition Lim :: "partial1 set set" ("LIM") where + "LIM \ {U. \\ s. learn_lim \ U s}" + +text \LIM is closed under the the subset relation.\ + +lemma learn_lim_closed_subseteq: + assumes "learn_lim \ U s" and "V \ U" + shows "learn_lim \ V s" + using assms learn_lim_def by auto + +corollary LIM_closed_subseteq: + assumes "U \ LIM" and "V \ U" + shows "V \ LIM" + using assms learn_lim_closed_subseteq by (smt Lim_def mem_Collect_eq) + +text \Changing the hypothesis infinitely often precludes learning in +the limit.\ + +lemma infinite_hyp_changes_not_Lim: + assumes "f \ U" and "\n. \m\<^sub>1>n. \m\<^sub>2>n. s (f \ m\<^sub>1) \ s (f \ m\<^sub>2)" + shows "\ learn_lim \ U s" + using assms learn_lim_def by (metis less_imp_le) + +lemma always_hyp_change_not_Lim: + assumes "\x. s (f \ (Suc x)) \ s (f \ x)" + shows "\ learn_lim \ {f} s" + using assms learn_limE by (metis le_SucI order_refl singletonI) + +text \Guessing a wrong hypothesis infinitely often precludes learning +in the limit.\ + +lemma infinite_hyp_wrong_not_Lim: + assumes "f \ U" and "\n. \m>n. \ (the (s (f \ m))) \ f" + shows "\ learn_lim \ U s" + using assms learn_limE by (metis less_imp_le option.sel) + +text \Converging to the same hypothesis on two functions precludes +learning in the limit.\ + +lemma same_hyp_for_two_not_Lim: + assumes "f\<^sub>1 \ U" + and "f\<^sub>2 \ U" + and "f\<^sub>1 \ f\<^sub>2" + and "\n\n\<^sub>1. s (f\<^sub>1 \ n) = h" + and "\n\n\<^sub>2. s (f\<^sub>2 \ n) = h" + shows "\ learn_lim \ U s" + using assms learn_limE by (metis le_cases option.sel) + +text \Every class that can be learned in the limit can be learned in +the limit with respect to any Gödel numbering. We prove a generalization in +which hypotheses may have to satisfy an extra condition, so we can re-use it +for other inference types later.\ + +lemma learn_lim_extra_wrt_goedel: + fixes extra :: "(partial1 set) \ partial1 \ nat \ partial1 \ bool" + assumes "goedel_numbering \" + and "learn_lim \ U s" + and "\f n. f \ U \ extra U f n (\ (the (s (f \ n))))" + shows "\t. learn_lim \ U t \ (\f\U. \n. extra U f n (\ (the (t (f \ n)))))" +proof - + have env: "environment \ U s" + and lim: "learn_lim \ U s" + and extra: "\f\U. \n. extra U f n (\ (the (s (f \ n))))" + using assms learn_limE by auto + obtain c where c: "c \ \" "\i. \ i = \ (the (c i))" + using env goedel_numberingE[OF assms(1), of \] by auto + define t where "t \ + (\x. if s x \ \ c (the (s x)) \ then Some (the (c (the (s x)))) else None)" + have "t \ \

" + unfolding t_def using env c concat_P1_P1[of c s] by auto + have "t x = (if s x \ then Some (the (c (the (s x)))) else None)" for x + using t_def c(1) R1_imp_total1 by auto + then have t: "t (f \ n) \= the (c (the (s (f \ n))))" if "f \ U" for f n + using lim learn_limE that by simp + have "learn_lim \ U t" + proof (rule learn_limI) + show "environment \ U t" + using t by (simp add: \t \ \

\ env goedel_numbering_P2[OF assms(1)]) + show "\i. \ i = f \ (\\<^sup>\n. t (f \ n) \= i)" if "f \ U" for f + proof - + from lim learn_limE(2) obtain i n\<^sub>0 where + i: "\ i = f \ (\n\n\<^sub>0. s (f \ n) \= i)" + using \f \ U\ by blast + let ?j = "the (c i)" + have "\ ?j = f" + using c(2) i by simp + moreover have "t (f \ n) \= ?j" if "n \ n\<^sub>0" for n + by (simp add: \f \ U\ i t that) + ultimately show ?thesis by auto + qed + qed + moreover have "extra U f n (\ (the (t (f \ n))))" if "f \ U" for f n + proof - + from t have "the (t (f \ n)) = the (c (the (s (f \ n))))" + by (simp add: that) + then have "\ (the (t (f \ n))) = \ (the (s (f \ n)))" + using c(2) by simp + with extra show ?thesis using that by simp + qed + ultimately show ?thesis by auto +qed + +lemma learn_lim_wrt_goedel: + assumes "goedel_numbering \" and "learn_lim \ U s" + shows "\t. learn_lim \ U t" + using assms learn_lim_extra_wrt_goedel[where ?extra="\U f n h. True"] + by simp + +lemma LIM_wrt_phi_eq_Lim: "LIM_wrt \ = LIM" + using LIM_wrt_def Lim_def learn_lim_wrt_goedel[OF goedel_numbering_phi] + by blast + + +subsection \BC: Behaviorally correct learning in the limit\ + +text \Behaviorally correct learning in the limit relaxes LIM by +requiring that the strategy almost always output an index for the target +function, but not necessarily the same index. In other words convergence of +$(S(f^n))_{n\in\mathbb{N}}$ is replaced by convergence of +$(\psi_{S(f^n)})_{n\in\mathbb{N}}$.\ + +definition learn_bc :: "partial2 \ (partial1 set) \ partial1 \ bool" where + "learn_bc \ U s \ + environment \ U s \ + (\f\U. \\<^sup>\n. \ (the (s (f \ n))) = f)" + +lemma learn_bcE: + assumes "learn_bc \ U s" + shows "environment \ U s" + and "\f. f \ U \ \\<^sup>\n. \ (the (s (f \ n))) = f" + using assms learn_bc_def by auto + +lemma learn_bcI: + assumes "environment \ U s" + and "\f. f \ U \ \\<^sup>\n. \ (the (s (f \ n))) = f" + shows "learn_bc \ U s" + using assms learn_bc_def by auto + +definition BC_wrt :: "partial2 \ partial1 set set" where + "BC_wrt \ \ {U. \s. learn_bc \ U s}" + +definition BC :: "partial1 set set" where + "BC \ {U. \\ s. learn_bc \ U s}" + +text \BC is a superset of LIM and closed under the subset relation.\ + +lemma learn_lim_imp_BC: "learn_lim \ U s \ learn_bc \ U s" + using learn_limE learn_bcI[of \ U s] by fastforce + +lemma Lim_subseteq_BC: "LIM \ BC" + using learn_lim_imp_BC Lim_def BC_def by blast + +lemma learn_bc_closed_subseteq: + assumes "learn_bc \ U s" and "V \ U" + shows "learn_bc \ V s" + using assms learn_bc_def by auto + +corollary BC_closed_subseteq: + assumes "U \ BC" and "V \ U" + shows "V \ BC" + using assms by (smt BC_def learn_bc_closed_subseteq mem_Collect_eq) + +text \Just like with LIM, guessing a wrong hypothesis infinitely often +precludes BC-style learning.\ + +lemma infinite_hyp_wrong_not_BC: + assumes "f \ U" and "\n. \m>n. \ (the (s (f \ m))) \ f" + shows "\ learn_bc \ U s" +proof + assume "learn_bc \ U s" + then obtain n\<^sub>0 where "\n\n\<^sub>0. \ (the (s (f \ n))) = f" + using learn_bcE assms(1) by metis + with assms(2) show False using less_imp_le by blast +qed + +text \The proof that Gödel numberings suffice as hypothesis spaces for +BC is similar to the one for @{thm[source] learn_lim_extra_wrt_goedel}. We do +not need the @{term extra} part for BC, but we get it for free.\ + +lemma learn_bc_extra_wrt_goedel: + fixes extra :: "(partial1 set) \ partial1 \ nat \ partial1 \ bool" + assumes "goedel_numbering \" + and "learn_bc \ U s" + and "\f n. f \ U \ extra U f n (\ (the (s (f \ n))))" + shows "\t. learn_bc \ U t \ (\f\U. \n. extra U f n (\ (the (t (f \ n)))))" +proof - + have env: "environment \ U s" + and lim: "learn_bc \ U s" + and extra: "\f\U. \n. extra U f n (\ (the (s (f \ n))))" + using assms learn_bc_def by auto + obtain c where c: "c \ \" "\i. \ i = \ (the (c i))" + using env goedel_numberingE[OF assms(1), of \] by auto + define t where + "t = (\x. if s x \ \ c (the (s x)) \ then Some (the (c (the (s x)))) else None)" + have "t \ \

" + unfolding t_def using env c concat_P1_P1[of c s] by auto + have "t x = (if s x \ then Some (the (c (the (s x)))) else None)" for x + using t_def c(1) R1_imp_total1 by auto + then have t: "t (f \ n) \= the (c (the (s (f \ n))))" if "f \ U" for f n + using lim learn_bcE(1) that by simp + have "learn_bc \ U t" + proof (rule learn_bcI) + show "environment \ U t" + using t by (simp add: \t \ \

\ env goedel_numbering_P2[OF assms(1)]) + show "\\<^sup>\n. \ (the (t (f \ n))) = f" if "f \ U" for f + proof - + obtain n\<^sub>0 where "\n\n\<^sub>0. \ (the (s (f \ n))) = f" + using lim learn_bcE(2) \f \ U\ by blast + then show ?thesis using that t c(2) by auto + qed + qed + moreover have "extra U f n (\ (the (t (f \ n))))" if "f \ U" for f n + proof - + from t have "the (t (f \ n)) = the (c (the (s (f \ n))))" + by (simp add: that) + then have "\ (the (t (f \ n))) = \ (the (s (f \ n)))" + using c(2) by simp + with extra show ?thesis using that by simp + qed + ultimately show ?thesis by auto +qed + +corollary learn_bc_wrt_goedel: + assumes "goedel_numbering \" and "learn_bc \ U s" + shows "\t. learn_bc \ U t" + using assms learn_bc_extra_wrt_goedel[where ?extra="\_ _ _ _. True"] by simp + +corollary BC_wrt_phi_eq_BC: "BC_wrt \ = BC" + using learn_bc_wrt_goedel goedel_numbering_phi BC_def BC_wrt_def by blast + + +subsection \CONS: Learning in the limit with consistent hypotheses\ + +text \A hypothesis is \emph{consistent} if it matches all values in the +prefix given to the strategy. Consistent learning in the limit requires the +strategy to output only consistent hypotheses for prefixes from the class.\ + +definition learn_cons :: "partial2 \ (partial1 set) \ partial1 \ bool" where + "learn_cons \ U s \ + learn_lim \ U s \ + (\f\U. \n. \k\n. \ (the (s (f \ n))) k = f k)" + +definition CONS_wrt :: "partial2 \ partial1 set set" where + "CONS_wrt \ \ {U. \s. learn_cons \ U s}" + +definition CONS :: "partial1 set set" where + "CONS \ {U. \\ s. learn_cons \ U s}" + +lemma CONS_subseteq_Lim: "CONS \ LIM" + using CONS_def Lim_def learn_cons_def by blast + +lemma learn_consI: + assumes "environment \ U s" + and "\f. f \ U \ \i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i)" + and "\f n. f \ U \ \k\n. \ (the (s (f \ n))) k = f k" + shows "learn_cons \ U s" + using assms learn_lim_def learn_cons_def by simp + +text \If a consistent strategy converges, it automatically converges to +a correct hypothesis. Thus we can remove @{term "\ i = f"} from the second +assumption in the previous lemma.\ + +lemma learn_consI2: + assumes "environment \ U s" + and "\f. f \ U \ \i. \\<^sup>\n. s (f \ n) \= i" + and "\f n. f \ U \ \k\n. \ (the (s (f \ n))) k = f k" + shows "learn_cons \ U s" +proof (rule learn_consI) + show "environment \ U s" + and cons: "\f n. f \ U \ \k\n. \ (the (s (f \ n))) k = f k" + using assms by simp_all + show "\i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i)" if "f \ U" for f + proof - + from that assms(2) obtain i n\<^sub>0 where i_n0: "\n\n\<^sub>0. s (f \ n) \= i" + by blast + have "\ i x = f x" for x + proof (cases "x \ n\<^sub>0") + case True + then show ?thesis + using i_n0 cons that by fastforce + next + case False + moreover have "\k\x. \ (the (s (f \ x))) k = f k" + using cons that by simp + ultimately show ?thesis using i_n0 by simp + qed + with i_n0 show ?thesis by auto + qed +qed + +lemma learn_consE: + assumes "learn_cons \ U s" + shows "environment \ U s" + and "\f. f \ U \ \i n\<^sub>0. \ i = f \ (\n\n\<^sub>0. s (f \ n) \= i)" + and "\f n. f \ U \ \k\n. \ (the (s (f \ n))) k = f k" + using assms learn_cons_def learn_lim_def by auto + +lemma learn_cons_wrt_goedel: + assumes "goedel_numbering \" and "learn_cons \ U s" + shows "\t. learn_cons \ U t" + using learn_cons_def assms + learn_lim_extra_wrt_goedel[where ?extra="\U f n h. \k\n. h k = f k"] + by auto + +lemma CONS_wrt_phi_eq_CONS: "CONS_wrt \ = CONS" + using CONS_wrt_def CONS_def learn_cons_wrt_goedel goedel_numbering_phi + by blast + +lemma learn_cons_closed_subseteq: + assumes "learn_cons \ U s" and "V \ U" + shows "learn_cons \ V s" + using assms learn_cons_def learn_lim_closed_subseteq by auto + +lemma CONS_closed_subseteq: + assumes "U \ CONS" and "V \ U" + shows "V \ CONS" + using assms learn_cons_closed_subseteq by (smt CONS_def mem_Collect_eq) + +text \A consistent strategy cannot output the same hypothesis for two +different prefixes from the class to be learned.\ + +lemma same_hyp_different_init_not_cons: + assumes "f \ U" + and "g \ U" + and "f \ n \ g \ n" + and "s (f \ n) = s (g \ n)" + shows "\ learn_cons \ U s" + unfolding learn_cons_def by (auto, metis assms init_eqI) + + +subsection \TOTAL: Learning in the limit with total hypotheses\ + +text \Total learning in the limit requires the strategy to hypothesize +only total functions for prefixes from the class.\ + +definition learn_total :: "partial2 \ (partial1 set) \ partial1 \ bool" where + "learn_total \ U s \ + learn_lim \ U s \ + (\f\U. \n. \ (the (s (f \ n))) \ \)" + +definition TOTAL_wrt :: "partial2 \ partial1 set set" where + "TOTAL_wrt \ \ {U. \s. learn_total \ U s}" + +definition TOTAL :: "partial1 set set" where + "TOTAL \ {U. \\ s. learn_total \ U s}" + +lemma TOTAL_subseteq_LIM: "TOTAL \ LIM" + unfolding TOTAL_def Lim_def using learn_total_def by auto + +lemma learn_totalI: + assumes "environment \ U s" + and "\f. f \ U \ \i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i)" + and "\f n. f \ U \ \ (the (s (f \ n))) \ \" + shows "learn_total \ U s" + using assms learn_lim_def learn_total_def by auto + +lemma learn_totalE: + assumes "learn_total \ U s" + shows "environment \ U s" + and "\f. f \ U \ \i n\<^sub>0. \ i = f \ (\n\n\<^sub>0. s (f \ n) \= i)" + and "\f n. f \ U \ \ (the (s (f \ n))) \ \" + using assms learn_lim_def learn_total_def by auto + +lemma learn_total_wrt_goedel: + assumes "goedel_numbering \" and "learn_total \ U s" + shows "\t. learn_total \ U t" + using learn_total_def assms learn_lim_extra_wrt_goedel[where ?extra="\U f n h. h \ \"] + by auto + +lemma TOTAL_wrt_phi_eq_TOTAL: "TOTAL_wrt \ = TOTAL" + using TOTAL_wrt_def TOTAL_def learn_total_wrt_goedel goedel_numbering_phi + by blast + +lemma learn_total_closed_subseteq: + assumes "learn_total \ U s" and "V \ U" + shows "learn_total \ V s" + using assms learn_total_def learn_lim_closed_subseteq by auto + +lemma TOTAL_closed_subseteq: + assumes "U \ TOTAL" and "V \ U" + shows "V \ TOTAL" + using assms learn_total_closed_subseteq by (smt TOTAL_def mem_Collect_eq) + + +subsection \CP: Learning in the limit with class-preserving hypotheses\ + +text \Class-preserving learning in the limit requires all hypotheses +for prefixes from the class to be functions from the class.\ + +definition learn_cp :: "partial2 \ (partial1 set) \ partial1 \ bool" where + "learn_cp \ U s \ + learn_lim \ U s \ + (\f\U. \n. \ (the (s (f \ n))) \ U)" + +definition CP_wrt :: "partial2 \ partial1 set set" where + "CP_wrt \ \ {U. \s. learn_cp \ U s}" + +definition CP :: "partial1 set set" where + "CP \ {U. \\ s. learn_cp \ U s}" + +lemma learn_cp_wrt_goedel: + assumes "goedel_numbering \" and "learn_cp \ U s" + shows "\t. learn_cp \ U t" + using learn_cp_def assms learn_lim_extra_wrt_goedel[where ?extra="\U f n h. h \ U"] + by auto + +corollary CP_wrt_phi: "CP = CP_wrt \" + using learn_cp_wrt_goedel[OF goedel_numbering_phi] + by (smt CP_def CP_wrt_def Collect_cong) + +lemma learn_cpI: + assumes "environment \ U s" + and "\f. f \ U \ \i. \ i = f \ (\\<^sup>\n. s (f \ n) \= i)" + and "\f n. f \ U \ \ (the (s (f \ n))) \ U" + shows "learn_cp \ U s" + using assms learn_cp_def learn_lim_def by auto + +lemma learn_cpE: + assumes "learn_cp \ U s" + shows "environment \ U s" + and "\f. f \ U \ \i n\<^sub>0. \ i = f \ (\n\n\<^sub>0. s (f \ n) \= i)" + and "\f n. f \ U \ \ (the (s (f \ n))) \ U" + using assms learn_lim_def learn_cp_def by auto + +text \Since classes contain only total functions, a CP strategy is also +a TOTAL strategy.\ + +lemma learn_cp_imp_total: "learn_cp \ U s \ learn_total \ U s" + using learn_cp_def learn_total_def learn_lim_def by auto + +lemma CP_subseteq_TOTAL: "CP \ TOTAL" + using learn_cp_imp_total CP_def TOTAL_def by blast + + +subsection \FIN: Finite learning\ + +text \In general it is undecidable whether a LIM strategy has reached +its final hypothesis. By contrast, in finite learning (also called ``one-shot +learning'') the strategy signals when it is ready to output a hypothesis. Up +until then it outputs a ``don't know yet'' value. This value is represented +by zero and the actual hypothesis $i$ by $i + 1$.\ + +definition learn_fin :: "partial2 \ partial1 set \ partial1 \ bool" where + "learn_fin \ U s \ + environment \ U s \ + (\f \ U. \i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i))" + +definition FIN_wrt :: "partial2 \ partial1 set set" where + "FIN_wrt \ \ {U. \s. learn_fin \ U s}" + +definition FIN :: "partial1 set set" where + "FIN \ {U. \\ s. learn_fin \ U s}" + +lemma learn_finI: + assumes "environment \ U s" + and "\f. f \ U \ + \i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + shows "learn_fin \ U s" + using assms learn_fin_def by auto + +lemma learn_finE: + assumes "learn_fin \ U s" + shows "environment \ U s" + and "\f. f \ U \ + \i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + using assms learn_fin_def by auto + +lemma learn_fin_closed_subseteq: + assumes "learn_fin \ U s" and "V \ U" + shows "learn_fin \ V s" + using assms learn_fin_def by auto + +lemma learn_fin_wrt_goedel: + assumes "goedel_numbering \" and "learn_fin \ U s" + shows "\t. learn_fin \ U t" +proof - + have env: "environment \ U s" + and fin: "\f. f \ U \ + \i n\<^sub>0. \ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + using assms(2) learn_finE by auto + obtain c where c: "c \ \" "\i. \ i = \ (the (c i))" + using env goedel_numberingE[OF assms(1), of \] by auto + define t where "t \ + \x. if s x \ then None + else if s x = Some 0 then Some 0 + else Some (Suc (the (c (the (s x) - 1))))" + have "t \ \

" + proof - + from c obtain rc where rc: + "recfn 1 rc" + "total rc" + "\x. c x = eval rc [x]" + by auto + from env obtain rs where rs: "recfn 1 rs" "\x. s x = eval rs [x]" + by auto + then have "eval rs [f \ n] \" if "f \ U" for f n + using env that by simp + define rt where "rt = Cn 1 r_ifz [rs, Z, Cn 1 S [Cn 1 rc [Cn 1 r_dec [rs]]]]" + then have "recfn 1 rt" + using rc(1) rs(1) by simp + have "eval rt [x] \" if "eval rs [x] \" for x + using rc(1) rs(1) rt_def that by auto + moreover have "eval rt [x] \= 0" if "eval rs [x] \= 0" for x + using rt_def that rc(1,2) rs(1) by simp + moreover have "eval rt [x] \= Suc (the (c (the (s x) - 1)))" if "eval rs [x] \\ 0" for x + using rt_def that rc rs by auto + ultimately have "eval rt [x] = t x" for x + by (simp add: rs(2) t_def) + with `recfn 1 rt` show ?thesis by auto + qed + have t: "t (f \ n) \= + (if s (f \ n) = Some 0 then 0 else Suc (the (c (the (s (f \ n)) - 1))))" + if "f \ U" for f n + using that env by (simp add: t_def) + have "learn_fin \ U t" + proof (rule learn_finI) + show "environment \ U t" + using t by (simp add: \t \ \

\ env goedel_numbering_P2[OF assms(1)]) + show "\i n\<^sub>0. \ i = f \ (\n0. t (f \ n) \= 0) \ (\n\n\<^sub>0. t (f \ n) \= Suc i)" + if "f \ U" for f + proof - + from fin obtain i n\<^sub>0 where + i: "\ i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + using \f \ U\ by blast + let ?j = "the (c i)" + have "\ ?j = f" + using c(2) i by simp + moreover have "\n0. t (f \ n) \= 0" + using t[OF that] i by simp + moreover have "t (f \ n) \= Suc ?j" if "n \ n\<^sub>0" for n + using that i t[OF `f \ U`] by simp + ultimately show ?thesis by auto + qed + qed + then show ?thesis by auto +qed + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/LIM_BC.thy b/thys/Inductive_Inference/LIM_BC.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/LIM_BC.thy @@ -0,0 +1,1329 @@ +section \LIM is a proper subset of BC\label{s:lim_bc}\ + +theory LIM_BC + imports Lemma_R +begin + +text \The proper inclusion of LIM in BC has been proved by +Barzdin~\cite{b-ttlsf-74} (see also Case and Smith~\cite{cs-cicmii-83}). The +proof constructs a class $V \in \mathrm{BC} - \mathrm{LIM}$ by +diagonalization against all LIM strategies. Exploiting Lemma~R for LIM, we +can assume that all such strategies are total functions. From the effective +version of this lemma we derive a numbering @{term "\ \ +\\<^sup>2"} such that for all $U \in \mathrm{LIM}$ there is an $i$ with +$U\in \mathrm{LIM}_\varphi(\sigma_i)$. The idea behind $V$ +is for every $i$ to construct a class $V_i$ of cardinality one or two such +that $V_i \notin \mathrm{LIM}_\varphi(\sigma_i)$. It then follows that the +union $V := \bigcup_i V_i$ cannot be learned by any $\sigma_i$ and thus $V +\notin \mathrm{LIM}$. At the same time, the construction ensures that the +functions in $V$ are ``predictable enough'' to be learnable in the BC sense. + +At the core is a process that maintains a state $(b, k)$ of a list $b$ of +numbers and an index $k < |b|$ into this list. We imagine $b$ to be the +prefix of the function being constructed, except for position $k$ where +we imagine $b$ to have a ``gap''; that is, $b_k$ is not defined yet. +Technically, we will always have $b_k = 0$, so $b$ also represents the prefix +after the ``gap is filled'' with 0, whereas $b_{k:=1}$ represents the prefix +where the gap is filled with 1. For every $i \in \mathbb{N}$, the process +starts in state $(i0, 1)$ and computes the next state from a given state +$(b,k)$ as follows: +\begin{enumerate} +\item if $ \sigma_i(b_{ + + +subsection \Enumerating enough total strategies\ + +text \For the construction of $\sigma$ we need the function @{term +r_limr} from the effective version of Lemma~R for LIM.\ + +definition "r_sigma \ Cn 2 r_phi [Cn 2 r_limr [Id 2 0], Id 2 1]" + +lemma r_sigma_recfn: "recfn 2 r_sigma" + unfolding r_sigma_def using r_limr_recfn by simp + +lemma r_sigma: "eval r_sigma [i, x] = \ (the (eval r_limr [i])) x" + unfolding r_sigma_def phi_def using r_sigma_recfn r_limr_total r_limr_recfn + by simp + +lemma r_sigma_total: "total r_sigma" + using r_sigma r_limr r_sigma_recfn totalI2[of r_sigma] by simp + +abbreviation sigma :: partial2 ("\") where + "\ i x \ eval r_sigma [i, x]" + +lemma sigma: "\ i = \ (the (eval r_limr [i]))" + using r_sigma by simp + +text \The numbering @{term \} does indeed enumerate enough total +strategies for every LIM learning problem.\ + +lemma learn_lim_sigma: + assumes "learn_lim \ U (\ i)" + shows "learn_lim \ U (\ i)" + using assms sigma r_limr by simp + + +subsection \The diagonalization process\ + +text \The following function represents the process described above. It +computes the next state from a given state $(b, k)$.\ + +definition "r_next \ + Cn 1 r_ifeq + [Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], r_pdec1], + Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], Cn 1 r_take [r_pdec2, r_pdec1]], + Cn 1 r_ifeq + [Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], Cn 1 r_update [r_pdec1, r_pdec2, r_const 1]], + Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], Cn 1 r_take [r_pdec2, r_pdec1]], + Cn 1 r_prod_encode [Cn 1 r_snoc [r_pdec1, Z], r_pdec2], + Cn 1 r_prod_encode + [Cn 1 r_snoc + [Cn 1 r_update [r_pdec1, r_pdec2, r_const 1], Z], Cn 1 r_length [r_pdec1]]], + Cn 1 r_prod_encode [Cn 1 r_snoc [r_pdec1, Z], Cn 1 r_length [r_pdec1]]]" + +lemma r_next_recfn: "recfn 1 r_next" + unfolding r_next_def using r_sigma_recfn by simp + +text \The three conditions distinguished in @{term r_next} correspond +to Steps 1, 2, and 3 of the process: hypothesis change when the gap is +filled with 0; hypothesis change when the gap is filled with 1; or +no hypothesis change either way.\ + +abbreviation "change_on_0 b k \ \ (e_hd b) b \ \ (e_hd b) (e_take k b)" + +abbreviation "change_on_1 b k \ + \ (e_hd b) b = \ (e_hd b) (e_take k b) \ + \ (e_hd b) (e_update b k 1) \ \ (e_hd b) (e_take k b)" + +abbreviation "change_on_neither b k \ + \ (e_hd b) b = \ (e_hd b) (e_take k b) \ + \ (e_hd b) (e_update b k 1) = \ (e_hd b) (e_take k b)" + +lemma change_conditions: + obtains + (on_0) "change_on_0 b k" + | (on_1) "change_on_1 b k" + | (neither) "change_on_neither b k" + by auto + +lemma r_next: + assumes "arg = prod_encode (b, k)" + shows "change_on_0 b k \ eval r_next [arg] \= prod_encode (e_snoc b 0, e_length b)" + and "change_on_1 b k \ + eval r_next [arg] \= prod_encode (e_snoc (e_update b k 1) 0, e_length b)" + and "change_on_neither b k \ eval r_next [arg] \= prod_encode (e_snoc b 0, k)" +proof - + let ?bhd = "Cn 1 r_hd [r_pdec1]" + let ?bup = "Cn 1 r_update [r_pdec1, r_pdec2, r_const 1]" + let ?bk = "Cn 1 r_take [r_pdec2, r_pdec1]" + let ?bap = "Cn 1 r_snoc [r_pdec1, Z]" + let ?len = "Cn 1 r_length [r_pdec1]" + let ?thenthen = "Cn 1 r_prod_encode [?bap, r_pdec2]" + let ?thenelse = "Cn 1 r_prod_encode [Cn 1 r_snoc [?bup, Z], ?len]" + let ?else = "Cn 1 r_prod_encode [?bap, ?len]" + have bhd: "eval ?bhd [arg] \= e_hd b" + using assms by simp + have bup: "eval ?bup [arg] \= e_update b k 1" + using assms by simp + have bk: "eval ?bk [arg] \= e_take k b" + using assms by simp + have bap: "eval ?bap [arg] \= e_snoc b 0" + using assms by simp + have len: "eval ?len [arg] \= e_length b" + using assms by simp + have else_: "eval ?else [arg] \= prod_encode (e_snoc b 0, e_length b)" + using bap len by simp + have thenthen: "eval ?thenthen [arg] \= prod_encode (e_snoc b 0, k)" + using bap assms by simp + have thenelse: "eval ?thenelse [arg] \= prod_encode (e_snoc (e_update b k 1) 0, e_length b)" + using bup len by simp + have then_: + "eval + (Cn 1 r_ifeq [Cn 1 r_sigma [?bhd, ?bup], Cn 1 r_sigma [?bhd, ?bk], ?thenthen, ?thenelse]) + [arg] \= + (if the (\ (e_hd b) (e_update b k 1)) = the (\ (e_hd b) (e_take k b)) + then prod_encode (e_snoc b 0, k) + else prod_encode (e_snoc (e_update b k 1) 0, e_length b))" + (is "eval ?then [arg] \= ?then_eval") + using bhd bup bk thenthen thenelse r_sigma r_sigma_recfn r_limr R1_imp_total1 by simp + have *: "eval r_next [arg] \= + (if the (\ (e_hd b) b) = the (\ (e_hd b) (e_take k b)) + then ?then_eval + else prod_encode (e_snoc b 0, e_length b))" + unfolding r_next_def + using bhd bk then_ else_ r_sigma r_sigma_recfn r_limr R1_imp_total1 assms + by simp + have r_sigma_neq: "eval r_sigma [x\<^sub>1, y\<^sub>1] \ eval r_sigma [x\<^sub>2, y\<^sub>2] \ + the (eval r_sigma [x\<^sub>1, y\<^sub>1]) \ the (eval r_sigma [x\<^sub>2, y\<^sub>2])" + for x\<^sub>1 y\<^sub>1 x\<^sub>2 y\<^sub>2 + using r_sigma r_limr totalE[OF r_sigma_total r_sigma_recfn] r_sigma_recfn r_sigma_total + by (metis One_nat_def Suc_1 length_Cons list.size(3) option.expand) + { + assume "change_on_0 b k" + then show "eval r_next [arg] \= prod_encode (e_snoc b 0, e_length b)" + using * r_sigma_neq by simp + next + assume "change_on_1 b k" + then show "eval r_next [arg] \= prod_encode (e_snoc (e_update b k 1) 0, e_length b)" + using * r_sigma_neq by simp + next + assume "change_on_neither b k" + then show "eval r_next [arg] \= prod_encode (e_snoc b 0, k)" + using * r_sigma_neq by simp + } +qed + +lemma r_next_total: "total r_next" +proof (rule totalI1) + show "recfn 1 r_next" + using r_next_recfn by simp + show "eval r_next [x] \" for x + proof - + obtain b k where "x = prod_encode (b, k)" + using prod_encode_pdec'[of x] by metis + then show ?thesis using r_next by fast + qed +qed + +text \The next function computes the state of the process after +any number of iterations.\ + +definition "r_state \ + Pr 1 + (Cn 1 r_prod_encode [Cn 1 r_snoc [Cn 1 r_singleton_encode [Id 1 0], Z], r_const 1]) + (Cn 3 r_next [Id 3 1])" + +lemma r_state_recfn: "recfn 2 r_state" + unfolding r_state_def using r_next_recfn by simp + +lemma r_state_at_0: "eval r_state [0, i] \= prod_encode (list_encode [i, 0], 1)" +proof - + let ?f = "Cn 1 r_prod_encode [Cn 1 r_snoc [Cn 1 r_singleton_encode [Id 1 0], Z], r_const 1]" + have "eval r_state [0, i] = eval ?f [i]" + unfolding r_state_def using r_next_recfn by simp + also have "... \= prod_encode (list_encode [i, 0], 1)" + by (simp add: list_decode_singleton) + finally show ?thesis . +qed + +lemma r_state_total: "total r_state" + unfolding r_state_def + using r_next_recfn totalE[OF r_next_total r_next_recfn] totalI3[of "Cn 3 r_next [Id 3 1]"] + by (intro Pr_total) auto + +text \We call the components of a state $(b, k)$ the \emph{block} $b$ +and the \emph{gap} $k$.\ + +definition block :: "nat \ nat \ nat" where + "block i t \ pdec1 (the (eval r_state [t, i]))" + +definition gap :: "nat \ nat \ nat" where + "gap i t \ pdec2 (the (eval r_state [t, i]))" + +lemma state_at_0: + "block i 0 = list_encode [i, 0]" + "gap i 0 = 1" + unfolding block_def gap_def r_state_at_0 by simp_all + +text \Some lemmas describing the behavior of blocks and gaps in +one iteration of the process:\ + +lemma state_Suc: + assumes "b = block i t" and "k = gap i t" + shows "block i (Suc t) = pdec1 (the (eval r_next [prod_encode (b, k)]))" + and "gap i (Suc t) = pdec2 (the (eval r_next [prod_encode (b, k)]))" +proof - + have "eval r_state [Suc t, i] = + eval (Cn 3 r_next [Id 3 1]) [t, the (eval r_state [t, i]), i]" + using r_state_recfn r_next_recfn totalE[OF r_state_total r_state_recfn, of "[t, i]"] + by (simp add: r_state_def) + also have "... = eval r_next [the (eval r_state [t, i])]" + using r_next_recfn by simp + also have "... = eval r_next [prod_encode (b, k)]" + using assms block_def gap_def by simp + finally have "eval r_state [Suc t, i] = eval r_next [prod_encode (b, k)]" . + then show + "block i (Suc t) = pdec1 (the (eval r_next [prod_encode (b, k)]))" + "gap i (Suc t) = pdec2 (the (eval r_next [prod_encode (b, k)]))" + by (simp add: block_def, simp add: gap_def) +qed + +lemma gap_Suc: + assumes "b = block i t" and "k = gap i t" + shows "change_on_0 b k \ gap i (Suc t) = e_length b" + and "change_on_1 b k \ gap i (Suc t) = e_length b" + and "change_on_neither b k\ gap i (Suc t) = k" + using assms r_next state_Suc by simp_all + +lemma block_Suc: + assumes "b = block i t" and "k = gap i t" + shows "change_on_0 b k \ block i (Suc t) = e_snoc b 0" + and "change_on_1 b k \ block i (Suc t) = e_snoc (e_update b k 1) 0" + and "change_on_neither b k\ block i (Suc t) = e_snoc b 0" + using assms r_next state_Suc by simp_all + +text \Non-gap positions in the block remain unchanged after an +iteration.\ + +lemma block_stable: + assumes "j < e_length (block i t)" and "j \ gap i t" + shows "e_nth (block i t) j = e_nth (block i (Suc t)) j" +proof - + from change_conditions[of "block i t" "gap i t"] show ?thesis + using assms block_Suc gap_Suc + by (cases, (simp_all add: nth_append)) +qed + +text \Next are some properties of @{term block} and @{term gap}.\ + +lemma gap_in_block: "gap i t < e_length (block i t)" +proof (induction t) + case 0 + then show ?case by (simp add: state_at_0) +next + case (Suc t) + with change_conditions[of "block i t" "gap i t"] show ?case + proof (cases) + case on_0 + then show ?thesis by (simp add: block_Suc(1) gap_Suc(1)) + next + case on_1 + then show ?thesis by (simp add: block_Suc(2) gap_Suc(2)) + next + case neither + then show ?thesis using Suc.IH block_Suc(3) gap_Suc(3) by force + qed +qed + +lemma length_block: "e_length (block i t) = Suc (Suc t)" +proof (induction t) + case 0 + then show ?case by (simp add: state_at_0) +next + case (Suc t) + with change_conditions[of "block i t" "gap i t"] show ?case + by (cases, simp_all add: block_Suc gap_Suc) +qed + +lemma gap_gr0: "gap i t > 0" +proof (induction t) + case 0 + then show ?case by (simp add: state_at_0) +next + case (Suc t) + with change_conditions[of "block i t" "gap i t"] show ?case + using length_block by (cases, simp_all add: block_Suc gap_Suc) +qed + +lemma hd_block: "e_hd (block i t) = i" +proof (induction t) + case 0 + then show ?case by (simp add: state_at_0) +next + case (Suc t) + from change_conditions[of "block i t" "gap i t"] show ?case + proof (cases) + case on_0 + then show ?thesis + using Suc block_Suc(1) length_block by (metis e_hd_snoc gap_Suc(1) gap_gr0) + next + case on_1 + let ?b = "block i t" and ?k = "gap i t" + have "?k > 0" + using gap_gr0 Suc by simp + then have "e_nth (e_update ?b ?k 1) 0 = e_nth ?b 0" + by simp + then have *: "e_hd (e_update ?b ?k 1) = e_hd ?b" + using e_hd_nth0 gap_Suc(2)[of _ i t] gap_gr0 on_1 by (metis e_length_update) + from on_1 have "block i (Suc t) = e_snoc (e_update ?b ?k 1) 0" + by (simp add: block_Suc(2)) + then show ?thesis + using e_hd_0 e_hd_snoc Suc length_block `?k > 0` * + by (metis e_length_update gap_Suc(2) gap_gr0 on_1) + next + case neither + then show ?thesis using Suc block_Suc(3) length_block by simp + qed +qed + +text \Formally, a block always ends in zero, even if it ends in a gap.\ + +lemma last_block: "e_nth (block i t) (gap i t) = 0" +proof (induction t) + case 0 + then show ?case by (simp add: state_at_0) +next + case (Suc t) + from change_conditions[of "block i t" "gap i t"] show ?case + proof cases + case on_0 + then show ?thesis using Suc by (simp add: block_Suc(1) gap_Suc(1)) + next + case on_1 + then show ?thesis using Suc by (simp add: block_Suc(2) gap_Suc(2) nth_append) + next + case neither + then have + "block i (Suc t) = e_snoc (block i t) 0" + "gap i (Suc t) = gap i t" + by (simp_all add: gap_Suc(3) block_Suc(3)) + then show ?thesis + using Suc gap_in_block by (simp add: nth_append) + qed +qed + +lemma gap_le_Suc: "gap i t \ gap i (Suc t)" + using change_conditions[of "block i t" "gap i t"] + gap_Suc gap_in_block less_imp_le[of "gap i t" "e_length (block i t)"] + by (cases) simp_all + +lemma gap_monotone: + assumes "t\<^sub>1 \ t\<^sub>2" + shows "gap i t\<^sub>1 \ gap i t\<^sub>2" +proof - + have "gap i t\<^sub>1 \ gap i (t\<^sub>1 + j)" for j + proof (induction j) + case 0 + then show ?case by simp + next + case (Suc j) + then show ?case using gap_le_Suc dual_order.trans by fastforce + qed + then show ?thesis using assms le_Suc_ex by blast +qed + +text \We need some lemmas relating the shape of the next state +to the hypothesis change conditions in Steps 1, 2, and 3.\ + +lemma state_change_on_neither: + assumes "gap i (Suc t) = gap i t" + shows "change_on_neither (block i t) (gap i t)" + and "block i (Suc t) = e_snoc (block i t) 0" +proof - + let ?b = "block i t" and ?k = "gap i t" + have "?k < e_length ?b" + using gap_in_block by simp + from change_conditions[of ?b ?k] show "change_on_neither (block i t) (gap i t)" + proof (cases) + case on_0 + then show ?thesis + using \?k < e_length ?b\ assms gap_Suc(1) by auto + next + case on_1 + then show ?thesis using assms gap_Suc(2) by auto + next + case neither + then show ?thesis by simp + qed + then show "block i (Suc t) = e_snoc (block i t) 0" + using block_Suc(3) by simp +qed + +lemma state_change_on_either: + assumes "gap i (Suc t) \ gap i t" + shows "\ change_on_neither (block i t) (gap i t)" + and "gap i (Suc t) = e_length (block i t)" +proof - + let ?b = "block i t" and ?k = "gap i t" + show "\ change_on_neither (block i t) (gap i t)" + proof + assume "change_on_neither (block i t) (gap i t)" + then have "gap i (Suc t) = ?k" + by (simp add: gap_Suc(3)) + with assms show False by simp + qed + then show "gap i (Suc t) = e_length (block i t)" + using gap_Suc(1) gap_Suc(2) by blast +qed + +text \Next up is the definition of $\tau$. In every iteration the +process determines $\tau_i(x)$ for some $x$ either by appending 0 to the +current block $b$, or by filling the current gap $k$. In the former case, +the value is determined for $x = |b|$, in the latter for $x = k$.\ + +text \For $i$ and $x$ the function @{term r_dettime} computes in which +iteration the process for $i$ determines the value $\tau_i(x)$. This is the +first iteration in which the block is long enough to contain position $x$ and +in which $x$ is not the gap. If $\tau_i(x)$ is never determined, because Case~2 is +reached with $k = x$, then @{term r_dettime} diverges.\ + +abbreviation determined :: "nat \ nat \ bool" where + "determined i x \ \t. x < e_length (block i t) \ x \ gap i t" + +lemma determined_0: "determined i 0" + using gap_gr0[of i 0] gap_in_block[of i 0] by force + +definition "r_dettime \ + Mn 2 + (Cn 3 r_and + [Cn 3 r_less + [Id 3 2, Cn 3 r_length [Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]]], + Cn 3 r_neq + [Id 3 2, Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]]])" + +lemma r_dettime_recfn: "recfn 2 r_dettime" + unfolding r_dettime_def using r_state_recfn by simp + +abbreviation dettime :: partial2 where + "dettime i x \ eval r_dettime [i, x]" + +lemma r_dettime: + shows "determined i x \ dettime i x \= (LEAST t. x < e_length (block i t) \ x \ gap i t)" + and "\ determined i x \ dettime i x \" +proof - + define f where "f = + (Cn 3 r_and + [Cn 3 r_less + [Id 3 2, Cn 3 r_length [Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]]], + Cn 3 r_neq + [Id 3 2, Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]]])" + then have "r_dettime = Mn 2 f" + unfolding f_def r_dettime_def by simp + have "recfn 3 f" + unfolding f_def using r_state_recfn by simp + then have "total f" + unfolding f_def using Cn_total r_state_total Mn_free_imp_total by simp + have f: "eval f [t, i, x] \= (if x < e_length (block i t) \ x \ gap i t then 0 else 1)" for t + proof - + let ?b = "Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]" + let ?k = "Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]" + have "eval ?b [t, i, x] \= pdec1 (the (eval r_state [t, i]))" + using r_state_recfn r_state_total by simp + then have b: "eval ?b [t, i, x] \= block i t" + using block_def by simp + have "eval ?k [t, i, x] \= pdec2 (the (eval r_state [t, i]))" + using r_state_recfn r_state_total by simp + then have k: "eval ?k [t, i, x] \= gap i t" + using gap_def by simp + have "eval + (Cn 3 r_neq [Id 3 2, Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]]) + [t, i, x] \= + (if x \ gap i t then 0 else 1)" + using b k r_state_recfn r_state_total by simp + moreover have "eval + (Cn 3 r_less + [Id 3 2, Cn 3 r_length [Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]]]) + [t, i, x] \= + (if x < e_length (block i t) then 0 else 1)" + using b k r_state_recfn r_state_total by simp + ultimately show ?thesis + unfolding f_def using b k r_state_recfn r_state_total by simp + qed + { + assume "determined i x" + with f have "\t. eval f [t, i, x] \= 0" by simp + then have "dettime i x \= (LEAST t. eval f [t, i, x] \= 0)" + using `total f` `r_dettime = Mn 2 f` r_dettime_recfn `recfn 3 f` + eval_Mn_total[of 2 f "[i, x]"] + by simp + then show "dettime i x \= (LEAST t. x < e_length (block i t) \ x \ gap i t)" + using f by simp + next + assume "\ determined i x" + with f have "\ (\t. eval f [t, i, x] \= 0)" by simp + then have "dettime i x \" + using `total f` `r_dettime = Mn 2 f` r_dettime_recfn `recfn 3 f` + eval_Mn_total[of 2 f "[i, x]"] + by simp + with f show "dettime i x \" by simp + } +qed + +lemma r_dettimeI: + assumes "x < e_length (block i t) \ x \ gap i t" + and "\T. x < e_length (block i T) \ x \ gap i T \ t \ T" + shows "dettime i x \= t" +proof - + let ?P = "\T. x < e_length (block i T) \ x \ gap i T" + have "determined i x" + using assms(1) by auto + moreover have "Least ?P = t" + using assms Least_equality[of ?P t] by simp + ultimately show ?thesis using r_dettime by simp +qed + +lemma r_dettime_0: "dettime i 0 \= 0" + using r_dettimeI[of _ i 0] determined_0 gap_gr0[of i 0] gap_in_block[of i 0] + by fastforce + +text \Computing the value of $\tau_i(x)$ works by running the process +@{term r_state} for @{term "dettime i x"} iterations and taking the value at +index $x$ of the resulting block.\ + +definition "r_tau \ Cn 2 r_nth [Cn 2 r_pdec1 [Cn 2 r_state [r_dettime, Id 2 0]], Id 2 1]" + +lemma r_tau_recfn: "recfn 2 r_tau" + unfolding r_tau_def using r_dettime_recfn r_state_recfn by simp + +abbreviation tau :: partial2 ("\") where + "\ i x \ eval r_tau [i, x]" + +lemma tau_in_P2: "\ \ \

\<^sup>2" + using r_tau_recfn by auto + +lemma tau_diverg: + assumes "\ determined i x" + shows "\ i x \" + unfolding r_tau_def using assms r_dettime r_dettime_recfn r_state_recfn by simp + +lemma tau_converg: + assumes "determined i x" + shows "\ i x \= e_nth (block i (the (dettime i x))) x" +proof - + from assms obtain t where t: "dettime i x \= t" + using r_dettime(1) by blast + then have "eval (Cn 2 r_state [r_dettime, Id 2 0]) [i, x] = eval r_state [t, i]" + using r_state_recfn r_dettime_recfn by simp + moreover have "eval r_state [t, i] \" + using r_state_total r_state_recfn by simp + ultimately have "eval (Cn 2 r_pdec1 [Cn 2 r_state [r_dettime, Id 2 0]]) [i, x] = + eval r_pdec1 [the (eval r_state [t, i])]" + using r_state_recfn r_dettime_recfn by simp + then show ?thesis + unfolding r_tau_def using r_state_recfn r_dettime_recfn t block_def by simp +qed + +lemma tau_converg': + assumes "dettime i x \= t" + shows "\ i x \= e_nth (block i t) x" + using assms tau_converg[of x i] r_dettime(2)[of x i] by fastforce + +lemma tau_at_0: "\ i 0 \= i" +proof - + have "\ i 0 \= e_nth (block i 0) 0" + using tau_converg'[OF r_dettime_0] by simp + then show ?thesis using block_def by (simp add: r_state_at_0) +qed + +lemma state_unchanged: + assumes "gap i t - 1 \ y" and "y \ t" + shows "gap i t = gap i y" +proof - + have "gap i t = gap i (gap i t - 1)" + proof (induction t) + case 0 + then show ?case by (simp add: gap_def r_state_at_0) + next + case (Suc t) + show ?case + proof (cases "gap i (Suc t) = t + 2") + case True + then show ?thesis by simp + next + case False + then show ?thesis + using Suc state_change_on_either(2) length_block by force + qed + qed + moreover have "gap i (gap i t - 1) \ gap i y" + using assms(1) gap_monotone by simp + moreover have "gap i y \ gap i t" + using assms(2) gap_monotone by simp + ultimately show ?thesis by simp +qed + +text \The values of the non-gap indices $x$ of every block created in +the diagonalization process equal $\tau_i(x)$.\ + +lemma tau_eq_state: + assumes "j < e_length (block i t)" and "j \ gap i t" + shows "\ i j \= e_nth (block i t) j" + using assms +proof (induction t) + case 0 + then have "j = 0" + using gap_gr0[of i 0] gap_in_block[of i 0] length_block[of i 0] by simp + then have "\ (e_hd (block i t)) j \= e_nth (block i (the (dettime i 0))) 0" + using determined_0 tau_converg hd_block by simp + then have "\ (e_hd (block i t)) j \= e_nth (block i 0) 0" + using r_dettime_0 by simp + then show ?case using \j = 0\ r_dettime_0 tau_converg' by simp +next + case (Suc t) + let ?b = "block i t" + let ?bb = "block i (Suc t)" + let ?k = "gap i t" + let ?kk = "gap i (Suc t)" + show ?case + proof (cases "?kk = ?k") + case kk_eq_k: True + then have bb_b0: "?bb = e_snoc ?b 0" + using state_change_on_neither by simp + show "\ i j \= e_nth ?bb j" + proof (cases "j < e_length ?b") + case True + then have "e_nth ?bb j = e_nth ?b j" + using bb_b0 by (simp add: nth_append) + moreover have "j \ ?k" + using Suc kk_eq_k by simp + ultimately show ?thesis using Suc True by simp + next + case False + then have j: "j = e_length ?b" + using Suc.prems(1) length_block by auto + then have "e_nth ?bb j = 0" + using bb_b0 by simp + have "dettime i j \= Suc t" + proof (rule r_dettimeI) + show "j < e_length ?bb \ j \ ?kk" + using Suc.prems(1,2) by linarith + show "\T. j < e_length (block i T) \ j \ gap i T \ Suc t \ T" + using length_block j by simp + qed + with tau_converg' show ?thesis by simp + qed + next + case False + then have kk_lenb: "?kk = e_length ?b" + using state_change_on_either by simp + then show ?thesis + proof (cases "j = ?k") + case j_eq_k: True + have "dettime i j \= Suc t" + proof (rule r_dettimeI) + show "j < e_length ?bb \ j \ ?kk" + using Suc.prems(1,2) by simp + show "Suc t \ T" if "j < e_length (block i T) \ j \ gap i T" for T + proof (rule ccontr) + assume "\ (Suc t \ T)" + then have "T < Suc t" by simp + then show False + proof (cases "T < ?k - 1") + case True + then have "e_length (block i T) = T + 2" + using length_block by simp + then have "e_length (block i T) < ?k + 1" + using True by simp + then have "e_length (block i T) \ ?k" by simp + then have "e_length (block i T) \ j" + using j_eq_k by simp + then show False + using that by simp + next + case False + then have "?k - 1 \ T" and "T \ t" + using `T < Suc t` by simp_all + with state_unchanged have "gap i t = gap i T" by blast + then show False + using j_eq_k that by simp + qed + qed + qed + then show ?thesis using tau_converg' by simp + next + case False + then have "j < e_length ?b" + using kk_lenb Suc.prems(1,2) length_block by auto + then show ?thesis using Suc False block_stable by fastforce + qed + qed +qed + +lemma tau_eq_state': + assumes "j < t + 2" and " j \ gap i t" + shows "\ i j \= e_nth (block i t) j" + using assms tau_eq_state length_block by simp + +text \We now consider the two cases described in the proof sketch. +In Case~2 there is a gap that never gets filled, or equivalently there is +a rightmost gap.\ + +abbreviation "case_two i \ (\t. \T. gap i T \ gap i t)" + +abbreviation "case_one i \ \ case_two i" + +text \Another characterization of Case~2 is that from some iteration on +only @{term change_on_neither} holds.\ + +lemma case_two_iff_forever_neither: + "case_two i \ (\t. \T\t. change_on_neither (block i T) (gap i T))" +proof + assume "\t. \T\t. change_on_neither (block i T) (gap i T)" + then obtain t where t: "\T\t. change_on_neither (block i T) (gap i T)" + by auto + have "(gap i T) \ (gap i t)" for T + proof (cases "T \ t") + case True + then show ?thesis using gap_monotone by simp + next + case False + then show ?thesis + proof (induction T) + case 0 + then show ?case by simp + next + case (Suc T) + with t have "change_on_neither ((block i T)) ((gap i T))" + by simp + then show ?case + using Suc.IH state_change_on_either(1)[of i T] gap_monotone[of T t i] + by metis + qed + qed + then show "\t. \T. gap i T \ gap i t" + by auto +next + assume "\t. \T. gap i T \ gap i t" + then obtain t where t: "\T. gap i T \ gap i t" + by auto + have "change_on_neither (block i T) (gap i T)" if "T\t" for T + proof - + have T: "(gap i T) \ (gap i t)" + using gap_monotone that by simp + show ?thesis + proof (rule ccontr) + assume "\ change_on_neither (block i T) (gap i T)" + then have "change_on_0 (block i T) (gap i T) \ change_on_1 (block i T) (gap i T)" + by simp + then have "gap i (Suc T) > gap i T" + using gap_le_Suc[of i] state_change_on_either(2)[of i] state_change_on_neither(1)[of i] + dual_order.strict_iff_order + by blast + with T have "gap i (Suc T) > gap i t" by simp + with t show False + using not_le by auto + qed + qed + then show "\t. \T\t. change_on_neither (block i T) (gap i T)" + by auto +qed + +text \In Case~1, $\tau_i$ is total.\ + +lemma case_one_tau_total: + assumes "case_one i" + shows "\ i x \" +proof (cases "x = gap i x") + case True + from assms have "\t. \T. gap i T > gap i t" + using le_less_linear gap_def[of i x] by blast + then obtain T where T: "gap i T > gap i x" + by auto + then have "T > x" + using gap_monotone leD le_less_linear by blast + then have "x < T + 2" by simp + moreover from T True have "x \ gap i T" by simp + ultimately show ?thesis using tau_eq_state' by simp +next + case False + moreover have "x < x + 2" by simp + ultimately show ?thesis using tau_eq_state' by blast +qed + +text \In Case~2, $\tau_i$ is undefined only at the gap that never gets filled.\ + +lemma case_two_tau_not_quite_total: + assumes "\T. gap i T \ gap i t" + shows "\ i (gap i t) \" + and "x \ gap i t \ \ i x \" +proof - + let ?k = "gap i t" + have "\ determined i ?k" + proof + assume "determined i ?k" + then obtain T where T: "?k < e_length (block i T) \ ?k \ gap i T" + by auto + with assms have snd_le: "gap i T < ?k" + by (simp add: dual_order.strict_iff_order) + then have "T < t" + using gap_monotone by (metis leD le_less_linear) + from T length_block have "?k < T + 2" by simp + moreover have "?k \ T + 1" + using T state_change_on_either(2) \T < t\ state_unchanged + by (metis Suc_eq_plus1 Suc_leI add_diff_cancel_right' le_add1 nat_neq_iff) + ultimately have "?k \ T" by simp + then have "gap i T = gap i ?k" + using state_unchanged[of i T "?k"] \?k < T + 2\ snd_le by simp + then show False + by (metis diff_le_self state_unchanged leD nat_le_linear gap_monotone snd_le) + qed + with tau_diverg show "\ i ?k \" by simp + + assume "x \ ?k" + show "\ i x \" + proof (cases "x < t + 2") + case True + with `x \ ?k` tau_eq_state' show ?thesis by simp + next + case False + then have "gap i x = ?k" + using assms by (simp add: dual_order.antisym gap_monotone) + with `x \ ?k` have "x \ gap i x" by simp + then show ?thesis using tau_eq_state'[of x x] by simp + qed +qed + +lemma case_two_tau_almost_total: + assumes "\t. \T. gap i T \ gap i t" (is "\t. ?P t") + shows "\ i (gap i (Least ?P)) \" + and "x \ gap i (Least ?P) \ \ i x \" +proof - + from assms have "?P (Least ?P)" + using LeastI_ex[of ?P] by simp + then show "\ i (gap i (Least ?P)) \" and "x \ gap i (Least ?P) \ \ i x \" + using case_two_tau_not_quite_total by simp_all +qed + +text \Some more properties of $\tau$.\ + +lemma init_tau_gap: "(\ i) \ (gap i t - 1) = e_take (gap i t) (block i t)" +proof (intro initI') + show 1: "e_length (e_take (gap i t) (block i t)) = Suc (gap i t - 1)" + proof - + have "gap i t > 0" + using gap_gr0 by simp + moreover have "gap i t < e_length (block i t)" + using gap_in_block by simp + ultimately have "e_length (e_take (gap i t) (block i t)) = gap i t" + by simp + then show ?thesis using gap_gr0 by simp + qed + show "\ i x \= e_nth (e_take (gap i t) (block i t)) x" if "x < Suc (gap i t - 1)" for x + proof - + have x_le: "x < gap i t" + using that gap_gr0 by simp + then have "x < e_length (block i t)" + using gap_in_block less_trans by blast + then have *: "\ i x \= e_nth (block i t) x" + using x_le tau_eq_state by auto + have "x < e_length (e_take (gap i t) (block i t))" + using x_le 1 by simp + then have "e_nth (block i t) x = e_nth (e_take (gap i t) (block i t)) x" + using x_le by simp + then show ?thesis using * by simp + qed +qed + +lemma change_on_0_init_tau: + assumes "change_on_0 (block i t) (gap i t)" + shows "(\ i) \ (t + 1) = block i t" +proof (intro initI') + let ?b = "block i t" and ?k = "gap i t" + show "e_length (block i t) = Suc (t + 1)" + using length_block by simp + show "(\ i) x \= e_nth (block i t) x" if "x < Suc (t + 1)" for x + proof (cases "x = ?k") + case True + have "gap i (Suc t) = e_length ?b" and b: "block i (Suc t) = e_snoc ?b 0" + using gap_Suc(1) block_Suc(1) assms by simp_all + then have "x < e_length (block i (Suc t))" "x \ gap i (Suc t)" + using that length_block by simp_all + then have "\ i x \= e_nth (block i (Suc t)) x" + using tau_eq_state by simp + then show ?thesis using that assms b by (simp add: nth_append) + next + case False + then show ?thesis using that assms tau_eq_state' by simp + qed +qed + +lemma change_on_0_hyp_change: + assumes "change_on_0 (block i t) (gap i t)" + shows "\ i ((\ i) \ (t + 1)) \ \ i ((\ i) \ (gap i t - 1))" + using assms hd_block init_tau_gap change_on_0_init_tau by simp + +lemma change_on_1_init_tau: + assumes "change_on_1 (block i t) (gap i t)" + shows "(\ i) \ (t + 1) = e_update (block i t) (gap i t) 1" +proof (intro initI') + let ?b = "block i t" and ?k = "gap i t" + show "e_length (e_update ?b ?k 1) = Suc (t + 1)" + using length_block by simp + show "(\ i) x \= e_nth (e_update ?b ?k 1) x" if "x < Suc (t + 1)" for x + proof (cases "x = ?k") + case True + have "gap i (Suc t) = e_length ?b" and b: "block i (Suc t) = e_snoc (e_update ?b ?k 1) 0" + using gap_Suc(2) block_Suc(2) assms by simp_all + then have "x < e_length (block i (Suc t))" "x \ gap i (Suc t)" + using that length_block by simp_all + then have "\ i x \= e_nth (block i (Suc t)) x" + using tau_eq_state by simp + then show ?thesis using that assms b nth_append by (simp add: nth_append) + next + case False + then show ?thesis using that assms tau_eq_state' by simp + qed +qed + +lemma change_on_1_hyp_change: + assumes "change_on_1 (block i t) (gap i t)" + shows "\ i ((\ i) \ (t + 1)) \ \ i ((\ i) \ (gap i t - 1))" + using assms hd_block init_tau_gap change_on_1_init_tau by simp + +lemma change_on_either_hyp_change: + assumes "\ change_on_neither (block i t) (gap i t)" + shows "\ i ((\ i) \ (t + 1)) \ \ i ((\ i) \ (gap i t - 1))" + using assms change_on_0_hyp_change change_on_1_hyp_change by auto + +lemma filled_gap_0_init_tau: + assumes "f\<^sub>0 = (\ i)((gap i t):=Some 0)" + shows "f\<^sub>0 \ (t + 1) = block i t" +proof (intro initI') + show len: "e_length (block i t) = Suc (t + 1)" + using assms length_block by auto + show "f\<^sub>0 x \= e_nth (block i t) x" if "x < Suc (t + 1)" for x + proof (cases "x = gap i t") + case True + then show ?thesis using assms last_block by auto + next + case False + then show ?thesis using assms len tau_eq_state that by auto + qed +qed + +lemma filled_gap_1_init_tau: + assumes "f\<^sub>1 = (\ i)((gap i t):=Some 1)" + shows "f\<^sub>1 \ (t + 1) = e_update (block i t) (gap i t) 1" +proof (intro initI') + show len: "e_length (e_update (block i t) (gap i t) 1) = Suc (t + 1)" + using e_length_update length_block by simp + show "f\<^sub>1 x \= e_nth (e_update (block i t) (gap i t) 1) x" if "x < Suc (t + 1)" for x + proof (cases "x = gap i t") + case True + moreover have "gap i t < e_length (block i t)" + using gap_in_block by simp + ultimately show ?thesis using assms by simp + next + case False + then show ?thesis using assms len tau_eq_state that by auto + qed +qed + + +subsection \The separating class\ + +text \Next we define the sets $V_i$ from the introductory proof sketch +(page~\pageref{s:lim_bc}).\ + +definition V_bclim :: "nat \ partial1 set" where + "V_bclim i \ + if case_two i + then let k = gap i (LEAST t. \T. gap i T \ gap i t) + in {(\ i)(k:=Some 0), (\ i)(k:=Some 1)} + else {\ i}" + +lemma V_subseteq_R1: "V_bclim i \ \" +proof (cases "case_two i") + case True + define k where "k = gap i (LEAST t. \T. gap i T \ gap i t)" + have "\ i \ \

" + using tau_in_P2 P2_proj_P1 by auto + then have "(\ i)(k:=Some 0) \ \

" and "(\ i)(k:=Some 1) \ \

" + using P1_update_P1 by simp_all + moreover have "total1 ((\ i)(k:=Some v))" for v + using case_two_tau_almost_total(2)[OF True] k_def total1_def by simp + ultimately have "(\ i)(k:=Some 0) \ \" and "(\ i)(k:=Some 1) \ \" + using P1_total_imp_R1 by simp_all + moreover have "V_bclim i = {(\ i)(k:=Some 0), (\ i)(k:=Some 1)}" + using True V_bclim_def k_def by (simp add: Let_def) + ultimately show ?thesis by simp +next + case False + have "V_bclim i = {\ i}" + unfolding V_bclim_def by (simp add: False) + moreover have "\ i \ \" + using total1I case_one_tau_total[OF False] tau_in_P2 P2_proj_P1[of \] P1_total_imp_R1 + by simp + ultimately show ?thesis by simp +qed + +lemma case_one_imp_gap_unbounded: + assumes "case_one i" + shows "\t. gap i t - 1 > n" +proof (induction n) + case 0 + then show ?case + using assms gap_gr0[of i] state_at_0(2)[of i] by (metis diff_is_0_eq gr_zeroI) +next + case (Suc n) + then obtain t where t: "gap i t - 1 > n" + by auto + moreover from assms have "\t. \T. gap i T > gap i t" + using leI by blast + ultimately obtain T where "gap i T > gap i t" + by auto + then have "gap i T - 1 > gap i t - 1" + using gap_gr0[of i] by (simp add: Suc_le_eq diff_less_mono) + with t have "gap i T - 1 > Suc n" by simp + then show ?case by auto +qed + +lemma case_one_imp_not_learn_lim_V: + assumes "case_one i" + shows "\ learn_lim \ (V_bclim i) (\ i)" +proof - + have V_bclim: "V_bclim i = {\ i}" + using assms V_bclim_def by (auto simp add: Let_def) + have "\m\<^sub>1>n. \m\<^sub>2>n. (\ i) ((\ i) \ m\<^sub>1) \ (\ i) ((\ i) \ m\<^sub>2)" for n + proof - + obtain t where t: "gap i t - 1 > n" + using case_one_imp_gap_unbounded[OF assms] by auto + moreover have "\t. \T\t. \ change_on_neither (block i T) (gap i T)" + using assms case_two_iff_forever_neither by blast + ultimately obtain T where T: "T \ t" "\ change_on_neither (block i T) (gap i T)" + by auto + then have "(\ i) ((\ i) \ (T + 1)) \ (\ i) ((\ i) \ (gap i T - 1))" + using change_on_either_hyp_change by simp + moreover have "gap i T - 1 > n" + using t T(1) gap_monotone by (simp add: diff_le_mono less_le_trans) + moreover have "T + 1 > n" + proof - + have "gap i T - 1 \ T" + using gap_in_block length_block by (simp add: le_diff_conv less_Suc_eq_le) + then show ?thesis using `gap i T - 1 > n` by simp + qed + ultimately show ?thesis by auto + qed + with infinite_hyp_changes_not_Lim V_bclim show ?thesis by simp +qed + +lemma case_two_imp_not_learn_lim_V: + assumes "case_two i" + shows "\ learn_lim \ (V_bclim i) (\ i)" +proof - + let ?P = "\t. \T. (gap i T) \ (gap i t)" + let ?t = "LEAST t. ?P t" + let ?k = "gap i ?t" + let ?b = "e_take ?k (block i ?t)" + have t: "\T. gap i T \ gap i ?t" + using assms LeastI_ex[of ?P] by simp + then have neither: "\T\?t. change_on_neither (block i T) (gap i T)" + using gap_le_Suc gap_monotone state_change_on_neither(1) + by (metis (no_types, lifting) antisym) + have gap_T: "\T\?t. gap i T = ?k" + using t gap_monotone antisym_conv by blast + define f\<^sub>0 where "f\<^sub>0 = (\ i)(?k:=Some 0)" + define f\<^sub>1 where "f\<^sub>1 = (\ i)(?k:=Some 1)" + show ?thesis + proof (rule same_hyp_for_two_not_Lim) + show "f\<^sub>0 \ V_bclim i" and "f\<^sub>1 \ V_bclim i" + using assms V_bclim_def f\<^sub>0_def f\<^sub>1_def by (simp_all add: Let_def) + show "f\<^sub>0 \ f\<^sub>1" using f\<^sub>0_def f\<^sub>1_def by (meson map_upd_eqD1 zero_neq_one) + show "\n\Suc ?t. \ i (f\<^sub>0 \ n) = \ i ?b" + proof - + have "\ i (block i T) = \ i (e_take ?k (block i T))" if "T \ ?t" for T + using that gap_T neither hd_block by metis + then have "\ i (block i T) = \ i ?b" if "T \ ?t" for T + by (metis (no_types, lifting) init_tau_gap gap_T that) + then have "\ i (f\<^sub>0 \ (T + 1)) = \ i ?b" if "T \ ?t" for T + using filled_gap_0_init_tau[of f\<^sub>0 i T] f\<^sub>0_def gap_T that + by (metis (no_types, lifting)) + then have "\ i (f\<^sub>0 \ T) = \ i ?b" if "T \ Suc ?t" for T + using that by (metis (no_types, lifting) Suc_eq_plus1 Suc_le_D Suc_le_mono) + then show ?thesis by simp + qed + show "\n\Suc ?t. \ i (f\<^sub>1 \ n) = \ i ?b" + proof - + have "\ i (e_update (block i T) ?k 1) = \ i (e_take ?k (block i T))" if "T \ ?t" for T + using neither by (metis (no_types, lifting) hd_block gap_T that) + then have "\ i (e_update (block i T) ?k 1) = \ i ?b" if "T \ ?t" for T + using that init_tau_gap[of i] gap_T by (metis (no_types, lifting)) + then have "\ i (f\<^sub>1 \ (T + 1)) = \ i ?b" if "T \ ?t" for T + using filled_gap_1_init_tau[of f\<^sub>1 i T] f\<^sub>1_def gap_T that + by (metis (no_types, lifting)) + then have "\ i (f\<^sub>1 \ T) = \ i ?b" if "T \ Suc ?t" for T + using that by (metis (no_types, lifting) Suc_eq_plus1 Suc_le_D Suc_le_mono) + then show ?thesis by simp + qed + qed +qed + +corollary not_learn_lim_V: "\ learn_lim \ (V_bclim i) (\ i)" + using case_one_imp_not_learn_lim_V case_two_imp_not_learn_lim_V + by (cases "case_two i") simp_all + +text \Next we define the separating class.\ + +definition V_BCLIM :: "partial1 set" ("V\<^bsub>BC-LIM\<^esub>") where + "V\<^bsub>BC-LIM\<^esub> \ \i. V_bclim i" + +lemma V_BCLIM_R1: "V\<^bsub>BC-LIM\<^esub> \ \" + using V_BCLIM_def V_subseteq_R1 by auto + +lemma V_BCLIM_not_in_Lim: "V\<^bsub>BC-LIM\<^esub> \ LIM" +proof + assume "V\<^bsub>BC-LIM\<^esub> \ LIM" + then obtain s where s: "learn_lim \ V\<^bsub>BC-LIM\<^esub> s" + using learn_lim_wrt_goedel[OF goedel_numbering_phi] Lim_def by blast + moreover obtain i where "\ i = s" + using s learn_limE(1) phi_universal by blast + ultimately have "learn_lim \ V\<^bsub>BC-LIM\<^esub> (\x. eval r_sigma [i, x])" + using learn_lim_sigma by simp + moreover have "V_bclim i \ V\<^bsub>BC-LIM\<^esub>" + using V_BCLIM_def by auto + ultimately have "learn_lim \ (V_bclim i) (\x. eval r_sigma [i, x])" + using learn_lim_closed_subseteq by simp + then show False + using not_learn_lim_V by simp +qed + + +subsection \The separating class is in BC\ + +text \In order to show @{term "V\<^bsub>BC-LIM\<^esub> \ BC"} we +define a hypothesis space that for every function $\tau_i$ and every list $b$ +of numbers contains a copy of $\tau_i$ with the first $|b|$ values replaced +by $b$.\ + +definition psitau :: partial2 ("\\<^sup>\") where + "\\<^sup>\ b x \ (if x < e_length b then Some (e_nth b x) else \ (e_hd b) x)" + +lemma psitau_in_P2: "\\<^sup>\ \ \

\<^sup>2" +proof - + define r where "r \ + Cn 2 + (r_lifz r_nth (Cn 2 r_tau [Cn 2 r_hd [Id 2 0], Id 2 1])) + [Cn 2 r_less [Id 2 1, Cn 2 r_length [Id 2 0]], Id 2 0, Id 2 1]" + then have "recfn 2 r" + using r_tau_recfn by simp + moreover have "eval r [b, x] = \\<^sup>\ b x" for b x + proof - + let ?f = "Cn 2 r_tau [Cn 2 r_hd [Id 2 0], Id 2 1]" + have "recfn 2 r_nth" "recfn 2 ?f" + using r_tau_recfn by simp_all + then have "eval (r_lifz r_nth ?f) [c, b, x] = + (if c = 0 then eval r_nth [b, x] else eval ?f [b, x])" for c + by simp + moreover have "eval r_nth [b, x] \= e_nth b x" + by simp + moreover have "eval ?f [b, x] = \ (e_hd b) x" + using r_tau_recfn by simp + ultimately have "eval (r_lifz r_nth ?f) [c, b, x] = + (if c = 0 then Some (e_nth b x) else \ (e_hd b) x)" for c + by simp + moreover have "eval (Cn 2 r_less [Id 2 1, Cn 2 r_length [Id 2 0]]) [b, x] \= + (if x < e_length b then 0 else 1)" + by simp + ultimately show ?thesis + unfolding r_def psitau_def using r_tau_recfn by simp + qed + ultimately show ?thesis by auto +qed + +lemma psitau_init: + "\\<^sup>\ (f \ n) x = (if x < Suc n then Some (the (f x)) else \ (the (f 0)) x)" +proof - + let ?e = "f \ n" + have "e_length ?e = Suc n" by simp + moreover have "x < Suc n \ e_nth ?e x = the (f x)" by simp + moreover have "e_hd ?e = the (f 0)" + using hd_init by simp + ultimately show ?thesis using psitau_def by simp +qed + +text \The class @{term V_BCLIM} can be learned BC-style in the +hypothesis space @{term psitau} by the identity function.\ + +lemma learn_bc_V_BCLIM: "learn_bc \\<^sup>\ V\<^bsub>BC-LIM\<^esub> Some" +proof (rule learn_bcI) + show "environment \\<^sup>\ V\<^bsub>BC-LIM\<^esub> Some" + using identity_in_R1 V_BCLIM_R1 psitau_in_P2 by auto + show "\n\<^sub>0. \n\n\<^sub>0. \\<^sup>\ (the (Some (f \ n))) = f" if "f \ V\<^bsub>BC-LIM\<^esub>" for f + proof - + from that V_BCLIM_def obtain i where i: "f \ V_bclim i" + by auto + show ?thesis + proof (cases "case_two i") + case True + let ?P = "\t. \T. (gap i T) \ (gap i t)" + let ?lmin = "LEAST t. ?P t" + define k where "k \ gap i ?lmin" + have V_bclim: "V_bclim i = {(\ i)(k:=Some 0), (\ i)(k:=Some 1)}" + using True V_bclim_def k_def by (simp add: Let_def) + moreover have "0 < k" + using gap_gr0[of i] k_def by simp + ultimately have "f 0 \= i" + using tau_at_0[of i] i by auto + have "\\<^sup>\ (f \ n) = f" if "n \ k" for n + proof + fix x + show "\\<^sup>\ (f \ n) x = f x" + proof (cases "x \ n") + case True + then show ?thesis + using R1_imp_total1 V_subseteq_R1 i psitau_init by fastforce + next + case False + then have "\\<^sup>\ (f \ n) x = \ (the (f 0)) x" + using psitau_init by simp + then have "\\<^sup>\ (f \ n) x = \ i x" + using `f 0 \= i` by simp + moreover have "f x = \ i x" + using False V_bclim i that by auto + ultimately show ?thesis by simp + qed + qed + then show ?thesis by auto + next + case False + then have "V_bclim i = {\ i}" + using V_bclim_def by (auto simp add: Let_def) + then have f: "f = \ i" + using i by simp + have "\\<^sup>\ (f \ n) = f" for n + proof + fix x + show "\\<^sup>\ (f \ n) x = f x" + proof (cases "x \ n") + case True + then show ?thesis + using R1_imp_total1 V_BCLIM_R1 psitau_init that by auto + next + case False + then show ?thesis by (simp add: f psitau_init tau_at_0) + qed + qed + then show ?thesis by simp + qed + qed +qed + +text \Finally, the main result of this section:\ + +theorem Lim_subset_BC: "LIM \ BC" + using learn_bc_V_BCLIM BC_def Lim_subseteq_BC V_BCLIM_not_in_Lim by auto + +end diff --git a/thys/Inductive_Inference/Lemma_R.thy b/thys/Inductive_Inference/Lemma_R.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/Lemma_R.thy @@ -0,0 +1,2114 @@ +section \Lemma R\label{s:lemma_r}\ + +theory Lemma_R + imports Inductive_Inference_Basics +begin + +text \A common technique for constructing a class that cannot be +learned is diagonalization against all strategies (see, for instance, +Section~\ref{s:lim_bc}). Similarly, the typical way of proving that a class +cannot be learned is by assuming there is a strategy and deriving a +contradiction. Both techniques are easier to carry out if one has to consider +only \emph{total} recursive strategies. This is not possible in general, +since after all the definitions of the inference types admit strictly partial +strategies. However, for many inference types one can show that for every +strategy there is a total strategy with at least the same ``learning power''. +Results to that effect are called Lemma~R. + +Lemma~R comes in different strengths depending on how general the +construction of the total recursive strategy is. CONS is the only inference +type considered here for which not even a weak form of Lemma~R holds.\ + + +subsection \Strong Lemma R for LIM, FIN, and BC\ + +text \In its strong form Lemma~R says that for any strategy $S$, there +is a total strategy $T$ that learns all classes $S$ learns regardless of +hypothesis space. The strategy $T$ can be derived from $S$ by a delayed +simulation of $S$. More precisely, for input $f^n$, $T$ simulates $S$ for +prefixes $f^0, f^1, \ldots, f^n$ for at most $n$ steps. If $S$ halts on none +of the prefixes, $T$ outputs an arbitrary hypothesis. Otherwise let $k \leq +n$ be maximal such that $S$ halts on $f^k$ in at most $n$ steps. Then $T$ +outputs $S(f^k)$. \ + +text \We reformulate some lemmas for @{term r_result1} to make it easier +to use them with @{term "\"}.\ + +lemma r_result1_converg_phi: + assumes "\ i x \= v" + shows "\t. + (\t'\t. eval r_result1 [t', i, x] \= Suc v) \ + (\t'= 0)" + using assms r_result1_converg' phi_def by simp_all + +lemma r_result1_bivalent': + assumes "eval r_phi [i, x] \= v" + shows "eval r_result1 [t, i, x] \= Suc v \ eval r_result1 [t, i, x] \= 0" + using assms r_result1 r_result_bivalent' r_phi'' by simp + +lemma r_result1_bivalent_phi: + assumes "\ i x \= v" + shows "eval r_result1 [t, i, x] \= Suc v \ eval r_result1 [t, i, x] \= 0" + using assms r_result1_bivalent' phi_def by simp_all + +lemma r_result1_diverg_phi: + assumes "\ i x \" + shows "eval r_result1 [t, i, x] \= 0" + using assms phi_def r_result1_diverg' by simp + +lemma r_result1_some_phi: + assumes "eval r_result1 [t, i, x] \= Suc v" + shows "\ i x \= v" + using assms phi_def r_result1_Some' by simp + +lemma r_result1_saturating': + assumes "eval r_result1 [t, i, x] \= Suc v" + shows "eval r_result1 [t + d, i, x] \= Suc v" + using assms r_result1 r_result_saturating r_phi'' by simp + +lemma r_result1_saturating_the: + assumes "the (eval r_result1 [t, i, x]) > 0" and "t' \ t" + shows "the (eval r_result1 [t', i, x]) > 0" +proof - + from assms(1) obtain v where "eval r_result1 [t, i, x] \= Suc v" + using r_result1_bivalent_phi r_result1_diverg_phi + by (metis inc_induct le_0_eq not_less_zero option.discI option.expand option.sel) + with assms have "eval r_result1 [t', i, x] \= Suc v" + using r_result1_saturating' le_Suc_ex by blast + then show ?thesis by simp +qed + +lemma Greatest_bounded_Suc: + fixes P :: "nat \ nat" + shows "(if P n > 0 then Suc n + else if \j 0 then Suc (GREATEST j. j < n \ P j > 0) else 0) = + (if \j 0 then Suc (GREATEST j. j < Suc n \ P j > 0) else 0)" + (is "?lhs = ?rhs") +proof (cases "\j 0") + case 1: True + show ?thesis + proof (cases "P n > 0") + case True + then have "(GREATEST j. j < Suc n \ P j > 0) = n" + using Greatest_equality[of "\j. j < Suc n \ P j > 0"] by simp + moreover have "?rhs = Suc (GREATEST j. j < Suc n \ P j > 0)" + using 1 by simp + ultimately have "?rhs = Suc n" by simp + then show ?thesis using True by simp + next + case False + then have "?lhs = Suc (GREATEST j. j < n \ P j > 0)" + using 1 by (metis less_SucE) + moreover have "?rhs = Suc (GREATEST j. j < Suc n \ P j > 0)" + using 1 by simp + moreover have "(GREATEST j. j < n \ P j > 0) = + (GREATEST j. j < Suc n \ P j > 0)" + using 1 False by (metis less_SucI less_Suc_eq) + ultimately show ?thesis by simp + qed +next + case False + then show ?thesis by auto +qed + +text \For $n$, $i$, $x$, the next function simulates $\varphi_i$ on all +non-empty prefixes of at most length $n$ of the list $x$ for at most $n$ +steps. It returns the length of the longest such prefix for which $\varphi_i$ +halts, or zero if $\varphi_i$ does not halt for any prefix.\ + +definition "r_delay_aux \ + Pr 2 (r_constn 1 0) + (Cn 4 r_ifz + [Cn 4 r_result1 + [Cn 4 r_length [Id 4 3], Id 4 2, + Cn 4 r_take [Cn 4 S [Id 4 0], Id 4 3]], + Id 4 1, Cn 4 S [Id 4 0]])" + +lemma r_delay_aux_prim: "prim_recfn 3 r_delay_aux" + unfolding r_delay_aux_def by simp_all + +lemma r_delay_aux_total: "total r_delay_aux" + using prim_recfn_total[OF r_delay_aux_prim] . + +lemma r_delay_aux: + assumes "n \ e_length x" + shows "eval r_delay_aux [n, i, x] \= + (if \j 0 + then Suc (GREATEST j. + j < n \ + the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0) + else 0)" +proof - + define z where "z \ + Cn 4 r_result1 + [Cn 4 r_length [Id 4 3], Id 4 2, Cn 4 r_take [Cn 4 S [Id 4 0], Id 4 3]]" + then have z_recfn: "recfn 4 z" by simp + have z: "eval z [j, r, i, x] = eval r_result1 [e_length x, i, e_take (Suc j) x]" + if "j < e_length x" for j r i x + unfolding z_def using that by simp + + define g where "g \ Cn 4 r_ifz [z, Id 4 1, Cn 4 S [Id 4 0]]" + then have g: "eval g [j, r, i, x] \= + (if the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0 then Suc j else r)" + if "j < e_length x" for j r i x + using that z prim_recfn_total z_recfn by simp + + show ?thesis + using assms + proof (induction n) + case 0 + moreover have "eval r_delay_aux [0, i, x] \= 0" + using eval_Pr_0 r_delay_aux_def r_delay_aux_prim r_constn + by (simp add: r_delay_aux_def) + ultimately show ?case by simp + next + case (Suc n) + let ?P = "\j. the (eval r_result1 [e_length x, i, e_take (Suc j) x])" + have "eval r_delay_aux [n, i, x] \" + using Suc by simp + moreover have "eval r_delay_aux [Suc n, i, x] = + eval (Pr 2 (r_constn 1 0) g) [Suc n, i, x]" + unfolding r_delay_aux_def g_def z_def by simp + ultimately have "eval r_delay_aux [Suc n, i, x] = + eval g [n, the (eval r_delay_aux [n, i, x]), i, x]" + using r_delay_aux_prim Suc eval_Pr_converg_Suc + by (simp add: r_delay_aux_def g_def z_def numeral_3_eq_3) + then have "eval r_delay_aux [Suc n, i, x] \= + (if ?P n > 0 then Suc n + else if \j 0 then Suc (GREATEST j. j < n \ ?P j > 0) else 0)" + using g Suc by simp + then have "eval r_delay_aux [Suc n, i, x] \= + (if \j 0 then Suc (GREATEST j. j < Suc n \ ?P j > 0) else 0)" + using Greatest_bounded_Suc[where ?P="?P"] by simp + then show ?case by simp + qed +qed + +text \The next function simulates $\varphi_i$ on all non-empty prefixes +of a list $x$ of length $n$ for at most $n$ steps and outputs the length of +the longest prefix for which $\varphi_i$ halts, or zero if $\varphi_i$ does +not halt for any such prefix.\ + +definition "r_delay \ Cn 2 r_delay_aux [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]" + +lemma r_delay_recfn [simp]: "recfn 2 r_delay" + unfolding r_delay_def by (simp add: r_delay_aux_prim) + +lemma r_delay: + "eval r_delay [i, x] \= + (if \j 0 + then Suc (GREATEST j. + j < e_length x \ the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0) + else 0)" + unfolding r_delay_def using r_delay_aux r_delay_aux_prim by simp + +definition "delay i x \ Some + (if \j 0 + then Suc (GREATEST j. + j < e_length x \ the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0) + else 0)" + +lemma delay_in_R2: "delay \ \\<^sup>2" + using r_delay totalI2 R2I delay_def r_delay_recfn + by (metis (no_types, lifting) numeral_2_eq_2 option.simps(3)) + +lemma delay_le_length: "the (delay i x) \ e_length x" +proof (cases "\j 0") + case True + let ?P = "\j. j < e_length x \ the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0" + from True have "\j. ?P j" by simp + moreover have "\y. ?P y \ y \ e_length x" by simp + ultimately have "?P (Greatest ?P)" + using GreatestI_ex_nat[where ?P="?P"] by blast + then have "Greatest ?P < e_length x" by simp + moreover have "delay i x \= Suc (Greatest ?P)" + using delay_def True by simp + ultimately show ?thesis by auto +next + case False + then show ?thesis using delay_def by auto +qed + +lemma e_take_delay_init: + assumes "f \ \" and "the (delay i (f \ n)) > 0" + shows "e_take (the (delay i (f \ n))) (f \ n) = f \ (the (delay i (f \ n)) - 1)" + using assms e_take_init[of f _ n] length_init[of f n] delay_le_length[of i "f \ n"] + by (metis One_nat_def Suc_le_lessD Suc_pred) + +lemma delay_gr0_converg: + assumes "the (delay i x) > 0" + shows "\ i (e_take (the (delay i x)) x) \" +proof - + let ?P = "\j. j < e_length x \ the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0" + have "\j. ?P j" + proof (rule ccontr) + assume "\ (\j. ?P j)" + then have "delay i x \= 0" + using delay_def by simp + with assms show False by simp + qed + then have d: "the (delay i x) = Suc (Greatest ?P)" + using delay_def by simp + moreover have "\y. ?P y \ y \ e_length x" by simp + ultimately have "?P (Greatest ?P)" + using `\j. ?P j` GreatestI_ex_nat[where ?P="?P"] by blast + then have "the (eval r_result1 [e_length x, i, e_take (Suc (Greatest ?P)) x]) > 0" + by simp + then have "the (eval r_result1 [e_length x, i, e_take (the (delay i x)) x]) > 0" + using d by simp + then show ?thesis using r_result1_diverg_phi by fastforce +qed + +lemma delay_unbounded: + fixes n :: nat + assumes "f \ \" and "\n. \ i (f \ n) \" + shows "\m. the (delay i (f \ m)) > n" +proof - + from assms have "\t. the (eval r_result1 [t, i, f \ n]) > 0" + using r_result1_converg_phi + by (metis le_refl option.exhaust_sel option.sel zero_less_Suc) + then obtain t where t: "the (eval r_result1 [t, i, f \ n]) > 0" + by auto + let ?m = "max n t" + have "Suc ?m \ t" by simp + have m: "the (eval r_result1 [Suc ?m, i, f \ n]) > 0" + proof - + let ?w = "eval r_result1 [t, i, f \ n]" + obtain v where v: "?w \= Suc v" + using t assms(2) r_result1_bivalent_phi by fastforce + have "eval r_result1 [Suc ?m, i, f \ n] = ?w" + using v t r_result1_saturating' `Suc ?m \ t` le_Suc_ex by fastforce + then show ?thesis using t by simp + qed + let ?x = "f \ ?m" + have "the (delay i ?x) > n" + proof - + let ?P = "\j. j < e_length ?x \ the (eval r_result1 [e_length ?x, i, e_take (Suc j) ?x]) > 0" + have "e_length ?x = Suc ?m" by simp + moreover have "e_take (Suc n) ?x = f \ n" + using assms(1) e_take_init by auto + ultimately have "?P n" + using m by simp + have "\y. ?P y \ y \ e_length ?x" by simp + with `?P n` have "n \ (Greatest ?P)" + using Greatest_le_nat[of ?P n "e_length ?x"] by simp + moreover have "the (delay i ?x) = Suc (Greatest ?P)" + using delay_def `?P n` by auto + ultimately show ?thesis by simp + qed + then show ?thesis by auto +qed + +lemma delay_monotone: + assumes "f \ \" and "n\<^sub>1 \ n\<^sub>2" + shows "the (delay i (f \ n\<^sub>1)) \ the (delay i (f \ n\<^sub>2))" + (is "the (delay i ?x1) \ the (delay i ?x2)") +proof (cases "the (delay i (f \ n\<^sub>1)) = 0") + case True + then show ?thesis by simp +next + case False + let ?P1 = "\j. j < e_length ?x1 \ the (eval r_result1 [e_length ?x1, i, e_take (Suc j) ?x1]) > 0" + let ?P2 = "\j. j < e_length ?x2 \ the (eval r_result1 [e_length ?x2, i, e_take (Suc j) ?x2]) > 0" + from False have d1: "the (delay i ?x1) = Suc (Greatest ?P1)" "\j. ?P1 j" + using delay_def option.collapse by fastforce+ + moreover have "\y. ?P1 y \ y \ e_length ?x1" by simp + ultimately have *: "?P1 (Greatest ?P1)" using GreatestI_ex_nat[of ?P1] by blast + let ?j = "Greatest ?P1" + from * have "?j < e_length ?x1" by auto + then have 1: "e_take (Suc ?j) ?x1 = e_take (Suc ?j) ?x2" + using assms e_take_init by auto + from * have 2: "?j < e_length ?x2" using assms(2) by auto + with 1 * have "the (eval r_result1 [e_length ?x1, i, e_take (Suc ?j) ?x2]) > 0" + by simp + moreover have "e_length ?x1 \ e_length ?x2" + using assms(2) by auto + ultimately have "the (eval r_result1 [e_length ?x2, i, e_take (Suc ?j) ?x2]) > 0" + using r_result1_saturating_the by simp + with 2 have "?P2 ?j" by simp + then have d2: "the (delay i ?x2) = Suc (Greatest ?P2)" + using delay_def by auto + have "\y. ?P2 y \ y \ e_length ?x2" by simp + with `?P2 ?j` have "?j \ (Greatest ?P2)" using Greatest_le_nat[of ?P2] by blast + with d1 d2 show ?thesis by simp +qed + +lemma delay_unbounded_monotone: + fixes n :: nat + assumes "f \ \" and "\n. \ i (f \ n) \" + shows "\m\<^sub>0. \m\m\<^sub>0. the (delay i (f \ m)) > n" +proof - + from assms delay_unbounded obtain m\<^sub>0 where "the (delay i (f \ m\<^sub>0)) > n" + by blast + then have "\m\m\<^sub>0. the (delay i (f \ m)) > n" + using assms(1) delay_monotone order.strict_trans2 by blast + then show ?thesis by auto +qed + +text \Now we can define a function that simulates an arbitrary strategy +$\varphi_i$ in a delayed way. The parameter $d$ is the default hypothesis for +when $\varphi_i$ does not halt within the time bound for any prefix.\ + +definition r_totalizer :: "nat \ recf" where + "r_totalizer d \ + Cn 2 + (r_lifz + (r_constn 1 d) + (Cn 2 r_phi + [Id 2 0, Cn 2 r_take [Cn 2 r_delay [Id 2 0, Id 2 1], Id 2 1]])) + [Cn 2 r_delay [Id 2 0, Id 2 1], Id 2 0, Id 2 1]" + +lemma r_totalizer_recfn: "recfn 2 (r_totalizer d)" + unfolding r_totalizer_def by simp + +lemma r_totalizer: + "eval (r_totalizer d) [i, x] = + (if the (delay i x) = 0 then Some d else \ i (e_take (the (delay i x)) x))" +proof - + let ?i = "Cn 2 r_delay [Id 2 0, Id 2 1]" + have "eval ?i [i, x] = eval r_delay [i, x]" for i x + using r_delay_recfn by simp + then have i: "eval ?i [i, x] = delay i x" for i x + using r_delay by (simp add: delay_def) + let ?t = "r_constn 1 d" + have t: "eval ?t [i, x] \= d" for i x by simp + let ?e1 = "Cn 2 r_take [?i, Id 2 1]" + let ?e = "Cn 2 r_phi [Id 2 0, ?e1]" + have "eval ?e1 [i, x] = eval r_take [the (delay i x), x]" for i x + using r_delay i delay_def by simp + then have "eval ?e1 [i, x] \= e_take (the (delay i x)) x" for i x + using delay_le_length by simp + then have e: "eval ?e [i, x] = \ i (e_take (the (delay i x)) x)" + using phi_def by simp + let ?z = "r_lifz ?t ?e" + have recfn_te: "recfn 2 ?t" "recfn 2 ?e" + by simp_all + then have "eval (r_totalizer d) [i, x] = eval (r_lifz ?t ?e) [the (delay i x), i, x]" + for i x + unfolding r_totalizer_def using i r_totalizer_recfn delay_def by simp + then have "eval (r_totalizer d) [i, x] = + (if the (delay i x) = 0 then eval ?t [i, x] else eval ?e [i, x])" + for i x + using recfn_te by simp + then show ?thesis using t e by simp +qed + +lemma r_totalizer_total: "total (r_totalizer d)" +proof (rule totalI2) + show "recfn 2 (r_totalizer d)" using r_totalizer_recfn by simp + show "\x y. eval (r_totalizer d) [x, y] \" + using r_totalizer delay_gr0_converg by simp +qed + +definition totalizer :: "nat \ partial2" where + "totalizer d i x \ + if the (delay i x) = 0 then Some d else \ i (e_take (the (delay i x)) x)" + +lemma totalizer_init: + assumes "f \ \" + shows "totalizer d i (f \ n) = + (if the (delay i (f \ n)) = 0 then Some d + else \ i (f \ (the (delay i (f \ n)) - 1)))" + using assms e_take_delay_init by (simp add: totalizer_def) + +lemma totalizer_in_R2: "totalizer d \ \\<^sup>2" + using totalizer_def r_totalizer r_totalizer_total R2I r_totalizer_recfn + by metis + +text \For LIM, @{term totalizer} works with every default hypothesis +$d$.\ + +lemma lemma_R_for_Lim: + assumes "learn_lim \ U (\ i)" + shows "learn_lim \ U (totalizer d i)" +proof (rule learn_limI) + show env: "environment \ U (totalizer d i)" + using assms learn_limE(1) totalizer_in_R2 by auto + show "\j. \ j = f \ (\\<^sup>\n. totalizer d i (f \ n) \= j)" if "f \ U" for f + proof - + have "f \ \" + using assms env that by auto + from assms learn_limE obtain j n\<^sub>0 where + j: "\ j = f" and + n0: "\n\n\<^sub>0. (\ i) (f \ n) \= j" + using `f \ U` by metis + obtain m\<^sub>0 where m0: "\m\m\<^sub>0. the (delay i (f \ m)) > n\<^sub>0" + using delay_unbounded_monotone `f \ \` \f \ U\ assms learn_limE(1) + by blast + then have "\m\m\<^sub>0. totalizer d i (f \ m) = \ i (e_take (the (delay i (f \ m))) (f \ m))" + using totalizer_def by auto + then have "\m\m\<^sub>0. totalizer d i (f \ m) = \ i (f \ (the (delay i (f \ m)) - 1))" + using e_take_delay_init m0 `f \ \` by auto + with m0 n0 have "\m\m\<^sub>0. totalizer d i (f \ m) \= j" + by auto + with j show ?thesis by auto + qed +qed + +text \The effective version of Lemma~R for LIM states that there is a +total recursive function computing Gödel numbers of total strategies +from those of arbitrary strategies.\ + +lemma lemma_R_for_Lim_effective: + "\g\\. \i. + \ (the (g i)) \ \ \ + (\U \. learn_lim \ U (\ i) \ learn_lim \ U (\ (the (g i))))" +proof - + have "totalizer 0 \ \

\<^sup>2" using totalizer_in_R2 by auto + then obtain g where g: "g \ \" "\i. (totalizer 0) i = \ (the (g i))" + using numbering_translation_for_phi by blast + with totalizer_in_R2 have "\i. \ (the (g i)) \ \" + by (metis R2_proj_R1) + moreover from g(2) lemma_R_for_Lim[where ?d=0] have + "\i U \. learn_lim \ U (\ i) \ learn_lim \ U (\ (the (g i)))" + by simp + ultimately show ?thesis using g(1) by blast +qed + +text \In order for us to use the previous lemma, we need a function +that performs the actual computation:\ + +definition "r_limr \ + SOME g. + recfn 1 g \ + total g \ + (\i. \ (the (eval g [i])) \ \ \ + (\U \. learn_lim \ U (\ i) \ learn_lim \ U (\ (the (eval g [i])))))" + +lemma r_limr_recfn: "recfn 1 r_limr" + and r_limr_total: "total r_limr" + and r_limr: + "\ (the (eval r_limr [i])) \ \" + "learn_lim \ U (\ i) \ learn_lim \ U (\ (the (eval r_limr [i])))" +proof - + let ?P = "\g. + g \ \ \ + (\i. \ (the (g i)) \ \ \ (\U \. learn_lim \ U (\ i) \ learn_lim \ U (\ (the (g i)))))" + let ?Q = "\g. + recfn 1 g \ + total g \ + (\i. \ (the (eval g [i])) \ \ \ + (\U \. learn_lim \ U (\ i) \ learn_lim \ U (\ (the (eval g [i])))))" + have "\g. ?P g" using lemma_R_for_Lim_effective by auto + then obtain g where "?P g" by auto + then obtain g' where g': "recfn 1 g'" "total g'" "\i. eval g' [i] = g i" + by blast + with `?P g` have "?Q g'" by simp + with r_limr_def someI_ex[of ?Q] show + "recfn 1 r_limr" + "total r_limr" + "\ (the (eval r_limr [i])) \ \" + "learn_lim \ U (\ i) \ learn_lim \ U (\ (the (eval r_limr [i])))" + by auto +qed + +text \For BC, too, @{term totalizer} works with every default +hypothesis $d$.\ + +lemma lemma_R_for_BC: + assumes "learn_bc \ U (\ i)" + shows "learn_bc \ U (totalizer d i)" +proof (rule learn_bcI) + show env: "environment \ U (totalizer d i)" + using assms learn_bcE(1) totalizer_in_R2 by auto + show "\n\<^sub>0. \n\n\<^sub>0. \ (the (totalizer d i (f \ n))) = f" if "f \ U" for f + proof - + have "f \ \" + using assms env that by auto + obtain n\<^sub>0 where n0: "\n\n\<^sub>0. \ (the ((\ i) (f \ n))) = f" + using assms learn_bcE `f \ U` by metis + obtain m\<^sub>0 where m0: "\m\m\<^sub>0. the (delay i (f \ m)) > n\<^sub>0" + using delay_unbounded_monotone `f \ \` \f \ U\ assms learn_bcE(1) + by blast + then have "\m\m\<^sub>0. totalizer d i (f \ m) = \ i (e_take (the (delay i (f \ m))) (f \ m))" + using totalizer_def by auto + then have "\m\m\<^sub>0. totalizer d i (f \ m) = \ i (f \ (the (delay i (f \ m)) - 1))" + using e_take_delay_init m0 `f \ \` by auto + with m0 n0 have "\m\m\<^sub>0. \ (the (totalizer d i (f \ m))) = f" + by auto + then show ?thesis by auto + qed +qed + +corollary lemma_R_for_BC_simple: + assumes "learn_bc \ U s" + shows "\s'\\. learn_bc \ U s'" + using assms lemma_R_for_BC totalizer_in_R2 learn_bcE + by (metis R2_proj_R1 learn_bcE(1) phi_universal) + + +text \For FIN the default hypothesis of @{term totalizer} must be +zero, signalling ``don't know yet''.\ + +lemma lemma_R_for_FIN: + assumes "learn_fin \ U (\ i)" + shows "learn_fin \ U (totalizer 0 i)" +proof (rule learn_finI) + show env: "environment \ U (totalizer 0 i)" + using assms learn_finE(1) totalizer_in_R2 by auto + show "\j n\<^sub>0. \ j = f \ + (\n0. totalizer 0 i (f \ n) \= 0) \ + (\n\n\<^sub>0. totalizer 0 i (f \ n) \= Suc j)" + if "f \ U" for f + proof - + have "f \ \" + using assms env that by auto + from assms learn_finE[of \ U "\ i"] obtain j where + j: "\ j = f" and + ex_n0: "\n\<^sub>0. (\n0. (\ i) (f \ n) \= 0) \ (\n\n\<^sub>0. (\ i) (f \ n) \= Suc j)" + using `f \ U` by blast + let ?Q = "\n\<^sub>0. (\n0. (\ i) (f \ n) \= 0) \ (\n\n\<^sub>0. (\ i) (f \ n) \= Suc j)" + define n\<^sub>0 where "n\<^sub>0 = Least ?Q" + with ex_n0 have n0: "?Q n\<^sub>0" "\n0. \ ?Q n" + using LeastI_ex[of ?Q] not_less_Least[of _ ?Q] by blast+ + define m\<^sub>0 where "m\<^sub>0 = (LEAST m\<^sub>0. \m\m\<^sub>0. the (delay i (f \ m)) > n\<^sub>0)" + (is "m\<^sub>0 = Least ?P") + moreover have "\m\<^sub>0. \m\m\<^sub>0. the (delay i (f \ m)) > n\<^sub>0" + using delay_unbounded_monotone `f\\` \f \ U\ assms learn_finE(1) + by simp + ultimately have m0: "?P m\<^sub>0" "\m0. \ ?P m" + using LeastI_ex[of ?P] not_less_Least[of _ ?P] by blast+ + then have "\m\m\<^sub>0. totalizer 0 i (f \ m) = \ i (e_take (the (delay i (f \ m))) (f \ m))" + using totalizer_def by auto + then have "\m\m\<^sub>0. totalizer 0 i (f \ m) = \ i (f \ (the (delay i (f \ m)) - 1))" + using e_take_delay_init m0 `f\\` by auto + with m0 n0 have "\m\m\<^sub>0. totalizer 0 i (f \ m) \= Suc j" + by auto + moreover have "totalizer 0 i (f \ m) \= 0" if "m < m\<^sub>0" for m + proof (cases "the (delay i (f \ m)) = 0") + case True + then show ?thesis by (simp add: totalizer_def) + next + case False + then have "the (delay i (f \ m)) \ n\<^sub>0" + using m0 that `f \ \` delay_monotone by (meson leI order.strict_trans2) + then show ?thesis + using \f \ \\ n0(1) totalizer_init by (simp add: Suc_le_lessD) + qed + ultimately show ?thesis using j by auto + qed +qed + + +subsection \Weaker Lemma R for CP and TOTAL\ + +text \For TOTAL the default hypothesis used by @{term totalizer} +depends on the hypothesis space, because it must refer to a total function in +that space. Consequently the total strategy depends on the hypothesis space, +which makes this form of Lemma~R weaker than the ones in the previous +section.\ + +lemma lemma_R_for_TOTAL: + fixes \ :: partial2 + shows "\d. \U. \i. learn_total \ U (\ i) \ learn_total \ U (totalizer d i)" +proof (cases "\d. \ d \ \") + case True + then obtain d where "\ d \ \" by auto + have "learn_total \ U (totalizer d i)" if "learn_total \ U (\ i)" for U i + proof (rule learn_totalI) + show env: "environment \ U (totalizer d i)" + using that learn_totalE(1) totalizer_in_R2 by auto + show "\f. f \ U \ \j. \ j = f \ (\\<^sup>\n. totalizer d i (f \ n) \= j)" + using that learn_total_def lemma_R_for_Lim[where ?d=d] learn_limE(2) by metis + show "\ (the (totalizer d i (f \ n))) \ \" if "f \ U" for f n + proof (cases "the (delay i (f \ n)) = 0") + case True + then show ?thesis using totalizer_def `\ d \ \` by simp + next + case False + have "f \ \" + using that env by auto + then show ?thesis + using False that `learn_total \ U (\ i)` totalizer_init learn_totalE(3) + by simp + qed + qed + then show ?thesis by auto +next + case False + then show ?thesis using learn_total_def lemma_R_for_Lim by auto +qed + +corollary lemma_R_for_TOTAL_simple: + assumes "learn_total \ U s" + shows "\s'\\. learn_total \ U s'" + using assms lemma_R_for_TOTAL totalizer_in_R2 + by (metis R2_proj_R1 learn_totalE(1) phi_universal) + +text \For CP the default hypothesis used by @{term totalizer} depends +on both the hypothesis space and the class. Therefore the total strategy +depends on both the the hypothesis space and the class, which makes Lemma~R +for CP even weaker than the one for TOTAL.\ + +lemma lemma_R_for_CP: + fixes \ :: partial2 and U :: "partial1 set" + assumes "learn_cp \ U (\ i)" + shows "\d. learn_cp \ U (totalizer d i)" +proof (cases "U = {}") + case True + then show ?thesis using assms learn_cp_def lemma_R_for_Lim by auto +next + case False + then obtain f where "f \ U" by auto + from `f \ U` obtain d where "\ d = f" + using learn_cpE(2)[OF assms] by auto + with `f \ U` have "\ d \ U" by simp + have "learn_cp \ U (totalizer d i)" + proof (rule learn_cpI) + show env: "environment \ U (totalizer d i)" + using assms learn_cpE(1) totalizer_in_R2 by auto + show "\f. f \ U \ \j. \ j = f \ (\\<^sup>\n. totalizer d i (f \ n) \= j)" + using assms learn_cp_def lemma_R_for_Lim[where ?d=d] learn_limE(2) by metis + show "\ (the (totalizer d i (f \ n))) \ U" if "f \ U" for f n + proof (cases "the (delay i (f \ n)) = 0") + case True + then show ?thesis using totalizer_def `\ d \ U` by simp + next + case False + then show ?thesis + using that env assms totalizer_init learn_cpE(3) by auto + qed + qed + then show ?thesis by auto +qed + + +subsection \No Lemma R for CONS\ + +text \This section demonstrates that the class $V_{01}$ of all total +recursive functions $f$ where $f(0)$ or $f(1)$ is a Gödel number of $f$ can +be consistently learned in the limit, but not by a total strategy. This implies +that Lemma~R does not hold for CONS.\ + +definition V01 :: "partial1 set" ("V\<^sub>0\<^sub>1") where + "V\<^sub>0\<^sub>1 = {f. f \ \ \ (\ (the (f 0)) = f \ \ (the (f 1)) = f)}" + + +subsubsection \No total CONS strategy for @{term "V\<^sub>0\<^sub>1"}\label{s:v01_not_total}\ + +text \In order to show that no total strategy can learn @{term +"V\<^sub>0\<^sub>1"} we construct, for each total strategy $S$, one or two +functions in @{term "V\<^sub>0\<^sub>1"} such that $S$ fails for at least one +of them. At the core of this construction is a process that given a total +recursive strategy $S$ and numbers $z, i, j \in \mathbb{N}$ builds a function +$f$ as follows: Set $f(0) = i$ and $f(1) = j$. For $x\geq1$: +\begin{enumerate} +\item[(a)] Check whether $S$ changes its hypothesis when $f^x$ is + extended by 0, that is, if $S(f^x) \neq S(f^x0)$. If so, set $f(x+1) = 0$. +\item[(b)] Otherwise check if $S$ changes its hypothesis when $f^x$ is extended + by $1$, that is, if $S(f^x) \neq S(f^x1)$. If so, set $f(x+1) = 1$. +\item[(c)] If neither happens, set $f(x+1) = z$. +\end{enumerate} +In other words, as long as we can force $S$ to change its hypothesis by +extending the function by 0 or 1, we do just that. Now there are two +cases: +\begin{enumerate} +\item[Case 1.] For all $x\geq1$ either (a) or (b) occurs; then $S$ + changes its hypothesis on $f$ all the time and thus does not learn $f$ in + the limit (not to mention consistently). The value of $z$ makes no + difference in this case. +\item[Case 2.] For some minimal $x$, (c) occurs, that is, + there is an $f^x$ such that $h := S(f^x) = S(f^x0) = S(f^x1)$. But the + hypothesis $h$ cannot be consistent with both prefixes $f^x0$ and $f^x1$. + Running the process once with $z = 0$ and once with $z = 1$ yields two + functions starting with $f^x0$ and $f^x1$, respectively, such that $S$ + outputs the same hypothesis, $h$, on both prefixes and thus cannot be + consistent for both functions. +\end{enumerate} +This process is computable because $S$ is total. The construction does not +work if we only assume $S$ to be a CONS strategy for $V_{01}$, because we +need to be able to apply $S$ to prefixes not in $V_{01}$. + +The parameters $i$ and $j$ provide flexibility to find functions built by the +above process that are actually in $V_{01}$. To this end we will use +Smullyan's double fixed-point theorem.\ + +context + fixes s :: partial1 + assumes s_in_R1 [simp, intro]: "s \ \" +begin + +text \The function @{term prefixes} constructs prefixes according to the +aforementioned process.\ + +fun prefixes :: "nat \ nat \ nat \ nat \ nat list" where + "prefixes z i j 0 = [i]" +| "prefixes z i j (Suc x) = prefixes z i j x @ + [if x = 0 then j + else if s (list_encode (prefixes z i j x @ [0])) \ s (list_encode (prefixes z i j x)) + then 0 + else if s (list_encode (prefixes z i j x @ [1])) \ s (list_encode (prefixes z i j x)) + then 1 + else z]" + +lemma prefixes_length: "length (prefixes z i j x) = Suc x" + by (induction x) simp_all + +text \The functions @{term[names_short] "adverse z i j"} are the +functions constructed by @{term[names_short] "prefixes"}.\ + +definition adverse :: "nat \ nat \ nat \ nat \ nat option" where + "adverse z i j x \ Some (last (prefixes z i j x))" + +lemma init_adverse_eq_prefixes: "(adverse z i j) \ n = list_encode (prefixes z i j n)" +proof - + have "prefix (adverse z i j) n = prefixes z i j n" + proof (induction n) + case 0 + then show ?case using adverse_def prefixes_length prefixI' by fastforce + next + case (Suc n) + then show ?case using adverse_def by (simp add: prefix_Suc) + qed + then show ?thesis by (simp add: init_def) +qed + +lemma adverse_at_01: + "adverse z i j 0 \= i" + "adverse z i j 1 \= j" + by (auto simp add: adverse_def) + +text \Had we introduced ternary partial recursive functions, the +@{term[names_short] "adverse z"} functions would be among them.\ + +lemma adverse_in_R3: "\r. recfn 3 r \ total r \ (\i j x. eval r [i, j, x]) = adverse z" +proof - + obtain rs where rs: "recfn 1 rs" "total rs" "(\x. eval rs [x]) = s" + using R1E by auto + have s_total: "\x. s x \" by simp + + define f where "f = Cn 2 r_singleton_encode [Id 2 0]" + then have "recfn 2 f" by simp + have f: "\i j. eval f [i, j] \= list_encode [i]" + unfolding f_def by simp + + define ch1 where "ch1 = Cn 4 r_ifeq + [Cn 4 rs [Cn 4 r_snoc [Id 4 1, r_constn 3 1]], + Cn 4 rs [Id 4 1], + r_dummy 3 (r_const z), + r_constn 3 1]" + then have ch1: "recfn 4 ch1" "total ch1" + using Cn_total prim_recfn_total rs by auto + + define ch0 where "ch0 = Cn 4 r_ifeq + [Cn 4 rs [Cn 4 r_snoc [Id 4 1, r_constn 3 0]], + Cn 4 rs [Id 4 1], + ch1, + r_constn 3 0]" + then have ch0_total: "total ch0" "recfn 4 ch0" + using Cn_total prim_recfn_total rs ch1 by auto + + have "eval ch1 [l, v, i, j] \= (if s (e_snoc v 1) = s v then z else 1)" for l v i j + proof - + have "eval ch1 [l, v, i, j] = eval r_ifeq [the (s (e_snoc v 1)), the (s v), z, 1]" + unfolding ch1_def using rs by auto + then show ?thesis by (simp add: s_total option.expand) + qed + moreover have "eval ch0 [l, v, i, j] \= + (if s (e_snoc v 0) = s v then the (eval ch1 [l, v, i, j]) else 0)" for l v i j + proof - + have "eval ch0 [l, v, i, j] = + eval r_ifeq [the (s (e_snoc v 0)), the (s v), the (eval ch1 [l, v, i, j]), 0]" + unfolding ch0_def using rs ch1 by auto + then show ?thesis by (simp add: s_total option.expand) + qed + ultimately have ch0: "\l v i j. eval ch0 [l, v, i, j] \= + (if s (e_snoc v 0) \ s v then 0 + else if s (e_snoc v 1) \ s v then 1 else z)" + by simp + + define app where "app = Cn 4 r_ifz [Id 4 0, Id 4 3, ch0]" + then have "recfn 4 app" "total app" + using ch0_total totalI4 by auto + have "eval app [l, v, i, j] \= (if l = 0 then j else the (eval ch0 [l, v, i, j]))" for l v i j + unfolding app_def using ch0_total by simp + with ch0 have app: "\l v i j. eval app [l, v, i, j] \= + (if l = 0 then j + else if s (e_snoc v 0) \ s v then 0 + else if s (e_snoc v 1) \ s v then 1 else z)" + by simp + + define g where "g = Cn 4 r_snoc [Id 4 1, app]" + with app have g: "\l v i j. eval g [l, v, i, j] \= e_snoc v + (if l = 0 then j + else if s (e_snoc v 0) \ s v then 0 + else if s (e_snoc v 1) \ s v then 1 else z)" + using `recfn 4 app` by auto + from g_def have "recfn 4 g" "total g" + using `recfn 4 app` `total app` Cn_total Mn_free_imp_total by auto + + define b where "b = Pr 2 f g" + then have "recfn 3 b" + using `recfn 2 f` `recfn 4 g` by simp + have b: "eval b [x, i, j] \= list_encode (prefixes z i j x)" for x i j + proof (induction x) + case 0 + then show ?case + unfolding b_def using f `recfn 2 f` \recfn 4 g\ by simp + next + case (Suc x) + then have "eval b [Suc x, i, j] = eval g [x, the (eval b [x, i, j]), i, j]" + using b_def `recfn 3 b` by simp + also have "... \= + (let v = list_encode (prefixes z i j x) + in e_snoc v + (if x = 0 then j + else if s (e_snoc v 0) \ s v then 0 + else if s (e_snoc v 1) \ s v then 1 else z))" + using g Suc by simp + also have "... \= + (let v = list_encode (prefixes z i j x) + in e_snoc v + (if x = 0 then j + else if s (list_encode (prefixes z i j x @ [0])) \ s v then 0 + else if s (list_encode (prefixes z i j x @ [1])) \ s v then 1 else z))" + using list_decode_encode by presburger + finally show ?case by simp + qed + + define b' where "b' = Cn 3 b [Id 3 2, Id 3 0, Id 3 1]" + then have "recfn 3 b'" + using `recfn 3 b` by simp + with b have b': "\i j x. eval b' [i, j, x] \= list_encode (prefixes z i j x)" + using b'_def by simp + + define r where "r = Cn 3 r_last [b']" + then have "recfn 3 r" + using `recfn 3 b'` by simp + with b' have "\i j x. eval r [i, j, x] \= last (prefixes z i j x)" + using r_def prefixes_length by auto + moreover from this have "total r" + using totalI3 `recfn 3 r` by simp + ultimately have "(\i j x. eval r [i, j, x]) = adverse z" + unfolding adverse_def by simp + with `recfn 3 r` `total r` show ?thesis by auto +qed + +lemma adverse_in_R1: "adverse z i j \ \" +proof - + from adverse_in_R3 obtain r where + r: "recfn 3 r" "total r" "(\i j x. eval r [i, j, x]) = adverse z" + by blast + define rij where "rij = Cn 1 r [r_const i, r_const j, Id 1 0]" + then have "recfn 1 rij" "total rij" + using r(1,2) Cn_total Mn_free_imp_total by auto + from rij_def have "\x. eval rij [x] = eval r [i, j, x]" + using r(1) by auto + with r(3) have "\x. eval rij [x] = adverse z i j x" + by metis + with `recfn 1 rij` `total rij` show ?thesis by auto +qed + +text \Next we show that for every $z$ there are $i$, $j$ such that +@{term[names_short] "adverse z i j \ V\<^sub>0\<^sub>1"}. The first step is to show that for every +$z$, Gödel numbers for @{term[names_short] "adverse z i j"} can be computed +uniformly from $i$ and $j$.\ + +lemma phi_translate_adverse: "\f\\\<^sup>2.\i j. \ (the (f i j)) = adverse z i j" +proof - + obtain r where r: "recfn 3 r" "total r" "(\i j x. eval r [i, j, x]) = adverse z" + using adverse_in_R3 by blast + let ?p = "encode r" + define rf where "rf = Cn 2 (r_smn 1 2) [r_dummy 1 (r_const ?p), Id 2 0, Id 2 1]" + then have "recfn 2 rf" and "total rf" + using Mn_free_imp_total by simp_all + define f where "f \ \i j. eval rf [i, j]" + with `recfn 2 rf` `total rf` have "f \ \\<^sup>2" by auto + have rf: "eval rf [i, j] = eval (r_smn 1 2) [?p, i, j]" for i j + unfolding rf_def by simp + { + fix i j x + have "\ (the (f i j)) x = eval r_phi [the (f i j), x]" + using phi_def by simp + also have "... = eval r_phi [the (eval rf [i, j]), x]" + using f_def by simp + also have "... = eval (r_universal 1) [the (eval (r_smn 1 2) [?p, i, j]), x]" + using rf r_phi_def by simp + also have "... = eval (r_universal (2 + 1)) (?p # [i, j] @ [x])" + using smn_lemma[of 1 "[i, j]" 2 "[x]"] by simp + also have "... = eval (r_universal 3) [?p, i, j, x]" + by simp + also have "... = eval r [i, j, x]" + using r_universal r by simp + also have "... = adverse z i j x" + using r(3) by metis + finally have "\ (the (f i j)) x = adverse z i j x" . + } + with `f \ \\<^sup>2` show ?thesis by blast +qed + +text \The second, and final, step is to apply Smullyan's double +fixed-point theorem to show the existence of @{term[names_short] adverse} +functions in @{term "V\<^sub>0\<^sub>1"}.\ + +lemma adverse_in_V01: "\m n. adverse 0 m n \ V\<^sub>0\<^sub>1 \ adverse 1 m n \ V\<^sub>0\<^sub>1" +proof - + obtain f\<^sub>0 where f0: "f\<^sub>0 \ \\<^sup>2" "\i j. \ (the (f\<^sub>0 i j)) = adverse 0 i j" + using phi_translate_adverse[of 0] by auto + obtain f\<^sub>1 where f1: "f\<^sub>1 \ \\<^sup>2" "\i j. \ (the (f\<^sub>1 i j)) = adverse 1 i j" + using phi_translate_adverse[of 1] by auto + obtain m n where "\ m = \ (the (f\<^sub>0 m n))" and "\ n = \ (the (f\<^sub>1 m n))" + using smullyan_double_fixed_point[OF f0(1) f1(1)] by blast + with f0(2) f1(2) have "\ m = adverse 0 m n" and "\ n = adverse 1 m n" + by simp_all + moreover have "the (adverse 0 m n 0) = m" and "the (adverse 1 m n 1) = n" + using adverse_at_01 by simp_all + ultimately have + "\ (the (adverse 0 m n 0)) = adverse 0 m n" + "\ (the (adverse 1 m n 1)) = adverse 1 m n" + by simp_all + moreover have "adverse 0 m n \ \" and "adverse 1 m n \ \" + using adverse_in_R1 by simp_all + ultimately show ?thesis using V01_def by auto +qed + +text \Before we prove the main result of this section we need some +lemmas regarding the shape of the @{term[names_short] adverse} functions and +hypothesis changes of the strategy.\ + +lemma adverse_Suc: + assumes "x > 0" + shows "adverse z i j (Suc x) \= + (if s (e_snoc ((adverse z i j) \ x) 0) \ s ((adverse z i j) \ x) + then 0 + else if s (e_snoc ((adverse z i j) \ x) 1) \ s ((adverse z i j) \ x) + then 1 else z)" +proof - + have "adverse z i j (Suc x) \= + (if s (list_encode (prefixes z i j x @ [0])) \ s (list_encode (prefixes z i j x)) + then 0 + else if s (list_encode (prefixes z i j x @ [1])) \ s (list_encode (prefixes z i j x)) + then 1 else z)" + using assms adverse_def by simp + then show ?thesis by (simp add: init_adverse_eq_prefixes) +qed + +text \The process in the proof sketch (page~\pageref{s:v01_not_total}) +consists of steps (a), (b), and (c). The next abbreviation is true iff.\ step +(a) or (b) applies.\ + +abbreviation "hyp_change z i j x \ + s (e_snoc ((adverse z i j) \ x) 0) \ s ((adverse z i j) \ x) \ + s (e_snoc ((adverse z i j) \ x) 1) \ s ((adverse z i j) \ x)" + +text \If step (c) applies, the process appends $z$.\ + +lemma adverse_Suc_not_hyp_change: + assumes "x > 0" and "\ hyp_change z i j x" + shows "adverse z i j (Suc x) \= z" + using assms adverse_Suc by simp + +text \While (a) or (b) applies, the process appends a value that +forces $S$ to change its hypothesis.\ + +lemma while_hyp_change: + assumes "\x\n. x > 0 \ hyp_change z i j x" + shows "\x\Suc n. adverse z i j x = adverse z' i j x" + using assms +proof (induction n) + case 0 + then show ?case by (simp add: adverse_def le_Suc_eq) +next + case (Suc n) + then have "\x\n. x > 0 \ hyp_change z i j x" by simp + with Suc have "\x\Suc n. x > 0 \ adverse z i j x = adverse z' i j x" + by simp + moreover have "adverse z i j 0 = adverse z' i j 0" + using adverse_at_01 by simp + ultimately have zz': "\x\Suc n. adverse z i j x = adverse z' i j x" + by auto + moreover have "adverse z i j \ \" "adverse z' i j \ \" + using adverse_in_R1 by simp_all + ultimately have init_zz': "(adverse z i j) \ (Suc n) = (adverse z' i j) \ (Suc n)" + using init_eqI by blast + + have "adverse z i j (Suc (Suc n)) = adverse z' i j (Suc (Suc n))" + proof (cases "s (e_snoc ((adverse z i j) \ (Suc n)) 0) \ s ((adverse z i j) \ (Suc n))") + case True + then have "s (e_snoc ((adverse z' i j) \ (Suc n)) 0) \ s ((adverse z' i j) \ (Suc n))" + using init_zz' by simp + then have "adverse z' i j (Suc (Suc n)) \= 0" + by (simp add: adverse_Suc) + moreover have "adverse z i j (Suc (Suc n)) \= 0" + using True by (simp add: adverse_Suc) + ultimately show ?thesis by simp + next + case False + then have "s (e_snoc ((adverse z' i j) \ (Suc n)) 0) = s ((adverse z' i j) \ (Suc n))" + using init_zz' by simp + then have "adverse z' i j (Suc (Suc n)) \= 1" + using init_zz' Suc.prems adverse_Suc by (smt le_refl zero_less_Suc) + moreover have "adverse z i j (Suc (Suc n)) \= 1" + using False Suc.prems adverse_Suc by auto + ultimately show ?thesis by simp + qed + with zz' show ?case using le_SucE by blast +qed + +text \The next result corresponds to Case~1 from the proof sketch.\ + +lemma always_hyp_change_no_lim: + assumes "\x>0. hyp_change z i j x" + shows "\ learn_lim \ {adverse z i j} s" +proof (rule infinite_hyp_changes_not_Lim[of "adverse z i j"]) + show "adverse z i j \ {adverse z i j}" by simp + show "\n. \m\<^sub>1>n. \m\<^sub>2>n. s (adverse z i j \ m\<^sub>1) \ s (adverse z i j \ m\<^sub>2)" + proof + fix n + from assms obtain m\<^sub>1 where m1: "m\<^sub>1 > n" "hyp_change z i j m\<^sub>1" + by auto + have "s (adverse z i j \ m\<^sub>1) \ s (adverse z i j \ (Suc m\<^sub>1))" + proof (cases "s (e_snoc ((adverse z i j) \ m\<^sub>1) 0) \ s ((adverse z i j) \ m\<^sub>1)") + case True + then have "adverse z i j (Suc m\<^sub>1) \= 0" + using m1 adverse_Suc by simp + then have "(adverse z i j) \ (Suc m\<^sub>1) = e_snoc ((adverse z i j) \ m\<^sub>1) 0" + by (simp add: init_Suc_snoc) + with True show ?thesis by simp + next + case False + then have "adverse z i j (Suc m\<^sub>1) \= 1" + using m1 adverse_Suc by simp + then have "(adverse z i j) \ (Suc m\<^sub>1) = e_snoc ((adverse z i j) \ m\<^sub>1) 1" + by (simp add: init_Suc_snoc) + with False m1(2) show ?thesis by simp + qed + then show "\m\<^sub>1>n. \m\<^sub>2>n. s (adverse z i j \ m\<^sub>1) \ s (adverse z i j \ m\<^sub>2)" + using less_SucI m1(1) by blast + qed +qed + +text \The next result corresponds to Case~2 from the proof sketch.\ + +lemma no_hyp_change_no_cons: + assumes "x > 0" and "\ hyp_change z i j x" + shows "\ learn_cons \ {adverse 0 i j, adverse 1 i j} s" +proof - + let ?P = "\x. x > 0 \ \ hyp_change z i j x" + define xmin where "xmin = Least ?P" + with assms have xmin: + "?P xmin" + "\x. x < xmin \ \ ?P x" + using LeastI[of ?P] not_less_Least[of _ ?P] by simp_all + then have "xmin > 0" by simp + + have "\x\xmin - 1. x > 0 \ hyp_change z i j x" + using xmin by (metis One_nat_def Suc_pred le_imp_less_Suc) + then have + "\x\xmin. adverse z i j x = adverse 0 i j x" + "\x\xmin. adverse z i j x = adverse 1 i j x" + using while_hyp_change[of "xmin - 1" z i j 0] + using while_hyp_change[of "xmin - 1" z i j 1] + by simp_all + then have + init_z0: "(adverse z i j) \ xmin = (adverse 0 i j) \ xmin" and + init_z1: "(adverse z i j) \ xmin = (adverse 1 i j) \ xmin" + using adverse_in_R1 init_eqI by blast+ + then have + a0: "adverse 0 i j (Suc xmin) \= 0" and + a1: "adverse 1 i j (Suc xmin) \= 1" + using adverse_Suc_not_hyp_change xmin(1) init_z1 + by metis+ + then have + i0: "(adverse 0 i j) \ (Suc xmin) = e_snoc ((adverse z i j) \ xmin) 0" and + i1: "(adverse 1 i j) \ (Suc xmin) = e_snoc ((adverse z i j) \ xmin) 1" + using init_z0 init_z1 by (simp_all add: init_Suc_snoc) + moreover have + "s (e_snoc ((adverse z i j) \ xmin) 0) = s ((adverse z i j) \ xmin)" + "s (e_snoc ((adverse z i j) \ xmin) 1) = s ((adverse z i j) \ xmin)" + using xmin by simp_all + ultimately have + "s ((adverse 0 i j) \ (Suc xmin)) = s ((adverse z i j) \ xmin)" + "s ((adverse 1 i j) \ (Suc xmin)) = s ((adverse z i j) \ xmin)" + by simp_all + then have + "s ((adverse 0 i j) \ (Suc xmin)) = s ((adverse 1 i j) \ (Suc xmin))" + by simp + moreover have "(adverse 0 i j) \ (Suc xmin) \ (adverse 1 i j) \ (Suc xmin)" + using a0 a1 i0 i1 by (metis append1_eq_conv list_decode_encode zero_neq_one) + ultimately show "\ learn_cons \ {adverse 0 i j, adverse 1 i j} s" + using same_hyp_different_init_not_cons by blast +qed + +text \Combining the previous two lemmas shows that @{term +"V\<^sub>0\<^sub>1"} cannot be learned consistently in the limit by the total +strategy $S$.\ + +lemma V01_not_in_R_cons: "\ learn_cons \ V\<^sub>0\<^sub>1 s" +proof - + obtain m n where + mn0: "adverse 0 m n \ V\<^sub>0\<^sub>1" and + mn1: "adverse 1 m n \ V\<^sub>0\<^sub>1" + using adverse_in_V01 by auto + show "\ learn_cons \ V\<^sub>0\<^sub>1 s" + proof (cases "\x>0. hyp_change 0 m n x") + case True + then have "\ learn_lim \ {adverse 0 m n} s" + using always_hyp_change_no_lim by simp + with mn0 show ?thesis + using learn_cons_def learn_lim_closed_subseteq by auto + next + case False + then obtain x where x: "x > 0" "\ hyp_change 0 m n x" by auto + then have "\ learn_cons \ {adverse 0 m n, adverse 1 m n} s" + using no_hyp_change_no_cons[OF x] by simp + with mn0 mn1 show ?thesis using learn_cons_closed_subseteq by auto + qed +qed + +end + + +subsubsection \@{term "V\<^sub>0\<^sub>1"} is in CONS\ + +text \At first glance, consistently learning @{term "V\<^sub>0\<^sub>1"} looks fairly +easy. After all every @{term "f \ V\<^sub>0\<^sub>1"} provides a Gödel number of itself +either at argument 0 or 1. A strategy only has to figure out which one is +right. However, the strategy $S$ we are going to devise does not always +converge to $f(0)$ or $f(1)$. Instead it uses a technique called +``amalgamation''. The amalgamation of two Gödel numbers $i$ and $j$ is a +function whose value at $x$ is determined by simulating $\varphi_i(x)$ and +$\varphi_j(x)$ in parallel and outputting the value of the first one to halt. +If neither halts the value is undefined. There is a function +$a\in\mathcal{R}^2$ such that $\varphi_{a(i,j)}$ is the amalgamation of $i$ +and $j$. + +If @{term "f \ V\<^sub>0\<^sub>1"} then $\varphi_{a(f(0), f(1))}$ is +total because by definition of @{term "V\<^sub>0\<^sub>1"} we have +$\varphi_{f(0)} = f$ or $\varphi_{f(1)} = f$ and $f$ is total. + +Given a prefix $f^n$ of an @{term "f \ V\<^sub>0\<^sub>1"} the strategy +$S$ first computes $\varphi_{a(f(0), f(1))}(x)$ for $x = 0, \ldots, n$. For +the resulting prefix $\varphi^n_{a(f(0), f(1))}$ there are two cases: +\begin{enumerate} +\item[Case 1.] It differs from $f^n$, say at minimum index $x$. Then for + either $z = 0$ or $z = 1$ we have $\varphi_{f(z)}(x) \neq f(x)$ by + definition of amalgamation. This + implies $\varphi_{f(z)} \neq f$, and thus $\varphi_{f(1-z)} = f$ by + definition of @{term "V\<^sub>0\<^sub>1"}. We set $S(f^n) = f(1 - z)$. This + hypothesis is correct and hence consistent. +\item[Case 2.] It equals $f^n$. Then we set $S(f^n) = a(f(0), f(1))$. This + hypothesis is consistent by definition of this case. +\end{enumerate} + +In both cases the hypothesis is consistent. If Case~1 holds for some $n$, the +same $x$ and $z$ will be found also for all larger values of $n$. Therefore +$S$ converges to the correct hypothesis $f(1 - z)$. If Case~2 holds for all +$n$, then $S$ always outputs the same hypothesis $a(f(0), f(1))$ and thus +also converges. + +The above discussion tacitly assumes $n \geq 1$, such that both $f(0)$ and +$f(1)$ are available to $S$. For $n = 0$ the strategy outputs an arbitrary +consistent hypothesis.\ + +text \Amalgamation uses the concurrent simulation of functions.\ + +definition parallel :: "nat \ nat \ nat \ nat option" where + "parallel i j x \ eval r_parallel [i, j, x]" + +lemma r_parallel': "eval r_parallel [i, j, x] = parallel i j x" + using parallel_def by simp + +lemma r_parallel'': + shows "eval r_phi [i, x] \ \ eval r_phi [j, x] \ \ eval r_parallel [i, j, x] \" + and "eval r_phi [i, x] \ \ eval r_phi [j, x] \ \ + eval r_parallel [i, j, x] \= prod_encode (0, the (eval r_phi [i, x]))" + and "eval r_phi [j, x] \ \ eval r_phi [i, x] \ \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval r_phi [j, x]))" + and "eval r_phi [i, x] \ \ eval r_phi [j, x] \ \ + eval r_parallel [i, j, x] \= prod_encode (0, the (eval r_phi [i, x])) \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval r_phi [j, x]))" +proof - + let ?f = "Cn 1 r_phi [r_const i, Id 1 0]" + let ?g = "Cn 1 r_phi [r_const j, Id 1 0]" + have *: "\x. eval r_phi [i, x] = eval ?f [x]" "\x. eval r_phi [j, x] = eval ?g [x]" + by simp_all + show "eval r_phi [i, x] \ \ eval r_phi [j, x] \ \ eval r_parallel [i, j, x] \" + and "eval r_phi [i, x] \ \ eval r_phi [j, x] \ \ + eval r_parallel [i, j, x] \= prod_encode (0, the (eval r_phi [i, x]))" + and "eval r_phi [j, x] \ \ eval r_phi [i, x] \ \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval r_phi [j, x]))" + and "eval r_phi [i, x] \ \ eval r_phi [j, x] \ \ + eval r_parallel [i, j, x] \= prod_encode (0, the (eval r_phi [i, x])) \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval r_phi [j, x]))" + using r_parallel[OF *] by simp_all +qed + +lemma parallel: + "\ i x \ \ \ j x \ \ parallel i j x \" + "\ i x \ \ \ j x \ \ parallel i j x \= prod_encode (0, the (\ i x))" + "\ j x \ \ \ i x \ \ parallel i j x \= prod_encode (1, the (\ j x))" + "\ i x \ \ \ j x \ \ + parallel i j x \= prod_encode (0, the (\ i x)) \ + parallel i j x \= prod_encode (1, the (\ j x))" + using phi_def r_parallel'' r_parallel parallel_def by simp_all + +lemma parallel_converg_pdec1_0_or_1: + assumes "parallel i j x \" + shows "pdec1 (the (parallel i j x)) = 0 \ pdec1 (the (parallel i j x)) = 1" + using assms parallel[of i x j] parallel(3)[of j x i] + by (metis fst_eqD option.sel prod_encode_inverse) + +lemma parallel_converg_either: "(\ i x \ \ \ j x \) = (parallel i j x \)" + using parallel by (metis option.simps(3)) + +lemma parallel_0: + assumes "parallel i j x \= prod_encode (0, v)" + shows "\ i x \= v" + using parallel assms + by (smt option.collapse option.sel option.simps(3) prod.inject prod_encode_eq zero_neq_one) + +lemma parallel_1: + assumes "parallel i j x \= prod_encode (1, v)" + shows "\ j x \= v" + using parallel assms + by (smt option.collapse option.sel option.simps(3) prod.inject prod_encode_eq zero_neq_one) + +lemma parallel_converg_V01: + assumes "f \ V\<^sub>0\<^sub>1" + shows "parallel (the (f 0)) (the (f 1)) x \" +proof - + have "f \ \ \ (\ (the (f 0)) = f \ \ (the (f 1)) = f)" + using assms V01_def by auto + then have "\ (the (f 0)) \ \ \ \ (the (f 1)) \ \" + by auto + then have "\ (the (f 0)) x \ \ \ (the (f 1)) x \" + using R1_imp_total1 by auto + then show ?thesis using parallel_converg_either by simp +qed + +text \The amalgamation of two Gödel numbers can then be described +in terms of @{term "parallel"}.\ + +definition amalgamation :: "nat \ nat \ partial1" where + "amalgamation i j x \ + if parallel i j x \ then None else Some (pdec2 (the (parallel i j x)))" + +lemma amalgamation_diverg: "amalgamation i j x \ \ \ i x \ \ \ j x \" + using amalgamation_def parallel by (metis option.simps(3)) + +lemma amalgamation_total: + assumes "total1 (\ i) \ total1 (\ j)" + shows "total1 (amalgamation i j)" + using assms amalgamation_diverg[of i j] total_def by auto + +lemma amalgamation_V01_total: + assumes "f \ V\<^sub>0\<^sub>1" + shows "total1 (amalgamation (the (f 0)) (the (f 1)))" + using assms V01_def amalgamation_total R1_imp_total1 total1_def + by (metis (mono_tags, lifting) mem_Collect_eq) + +definition "r_amalgamation \ Cn 3 r_pdec2 [r_parallel]" + +lemma r_amalgamation_recfn: "recfn 3 r_amalgamation" + unfolding r_amalgamation_def by simp + +lemma r_amalgamation: "eval r_amalgamation [i, j, x] = amalgamation i j x" +proof (cases "parallel i j x \") + case True + then have "eval r_parallel [i, j, x] \" + by (simp add: r_parallel') + then have "eval r_amalgamation [i, j, x] \" + unfolding r_amalgamation_def by simp + moreover from True have "amalgamation i j x \" + using amalgamation_def by simp + ultimately show ?thesis by simp +next + case False + then have "eval r_parallel [i, j, x] \" + by (simp add: r_parallel') + then have "eval r_amalgamation [i, j, x] = eval r_pdec2 [the (eval r_parallel [i, j, x])]" + unfolding r_amalgamation_def by simp + also have "... \= pdec2 (the (eval r_parallel [i, j, x]))" + by simp + finally show ?thesis by (simp add: False amalgamation_def r_parallel') +qed + +text \The function @{term "amalgamate"} computes Gödel numbers of +amalgamations. It corresponds to the function $a$ from the proof sketch.\ + +definition amalgamate :: "nat \ nat \ nat" where + "amalgamate i j \ smn 1 (encode r_amalgamation) [i, j]" + +lemma amalgamate: "\ (amalgamate i j) = amalgamation i j" +proof + fix x + have "\ (amalgamate i j) x = eval r_phi [amalgamate i j, x]" + by (simp add: phi_def) + also have "... = eval r_phi [smn 1 (encode r_amalgamation) [i, j], x]" + using amalgamate_def by simp + also have "... = eval r_phi + [encode (Cn 1 (r_universal 3) + (r_constn 0 (encode r_amalgamation) # map (r_constn 0) [i, j] @ map (Id 1) [0])), x]" + using smn[of 1 "encode r_amalgamation" "[i, j]"] by (simp add: numeral_3_eq_3) + also have "... = eval r_phi + [encode (Cn 1 (r_universal 3) + (r_const (encode r_amalgamation) # [r_const i, r_const j, Id 1 0])), x]" + (is "... = eval r_phi [encode ?f, x]") + by (simp add: r_constn_def) + finally have "\ (amalgamate i j) x = eval r_phi + [encode (Cn 1 (r_universal 3) + (r_const (encode r_amalgamation) # [r_const i, r_const j, Id 1 0])), x]" . + then have "\ (amalgamate i j) x = eval (r_universal 3) [encode r_amalgamation, i, j, x]" + unfolding r_phi_def using r_universal[of ?f 1] r_amalgamation_recfn by simp + then show "\ (amalgamate i j) x = amalgamation i j x" + using r_amalgamation by (simp add: r_amalgamation_recfn r_universal) +qed + +lemma amalgamation_in_P1: "amalgamation i j \ \

" + using amalgamate by (metis P2_proj_P1 phi_in_P2) + +lemma amalgamation_V01_R1: + assumes "f \ V\<^sub>0\<^sub>1" + shows "amalgamation (the (f 0)) (the (f 1)) \ \" + using assms amalgamation_V01_total amalgamation_in_P1 + by (simp add: P1_total_imp_R1) + +definition "r_amalgamate \ + Cn 2 (r_smn 1 2) [r_dummy 1 (r_const (encode r_amalgamation)), Id 2 0, Id 2 1]" + +lemma r_amalgamate_recfn: "recfn 2 r_amalgamate" + unfolding r_amalgamate_def by simp + +lemma r_amalgamate: "eval r_amalgamate [i, j] \= amalgamate i j" +proof - + let ?p = "encode r_amalgamation" + have rs21: "eval (r_smn 1 2) [?p, i, j] \= smn 1 ?p [i, j]" + using r_smn by simp + moreover have "eval r_amalgamate [i, j] = eval (r_smn 1 2) [?p, i, j]" + unfolding r_amalgamate_def by auto + ultimately have "eval r_amalgamate [i, j] \= smn 1 ?p [i, j]" + by simp + then show ?thesis using amalgamate_def by simp +qed + +text \The strategy $S$ distinguishes the two cases from the proof +sketch with the help of the next function, which checks if a hypothesis +$\varphi_i$ is inconsistent with a prefix $e$. If so, it returns the least $x +< |e|$ witnessing the inconsistency; otherwise it returns the length $|e|$. +If $\varphi_i$ diverges for some $x < |e|$, so does the function.\ + +definition inconsist :: partial2 where + "inconsist i e \ + (if \x i x \ then None + else if \x i x \\ e_nth e x + then Some (LEAST x. x < e_length e \ \ i x \\ e_nth e x) + else Some (e_length e))" + +lemma inconsist_converg: + assumes "inconsist i e \" + shows "inconsist i e = + (if \x i x \\ e_nth e x + then Some (LEAST x. x < e_length e \ \ i x \\ e_nth e x) + else Some (e_length e))" + and "\x i x \" + using inconsist_def assms by (presburger, meson) + +lemma inconsist_bounded: + assumes "inconsist i e \" + shows "the (inconsist i e) \ e_length e" +proof (cases "\x i x \\ e_nth e x") + case True + then show ?thesis + using inconsist_converg[OF assms] + by (smt Least_le dual_order.strict_implies_order dual_order.strict_trans2 option.sel) +next + case False + then show ?thesis using inconsist_converg[OF assms] by auto +qed + +lemma inconsist_consistent: + assumes "inconsist i e \" + shows "inconsist i e \= e_length e \ (\x i x \= e_nth e x)" +proof + show "\x i x \= e_nth e x" if "inconsist i e \= e_length e" + proof (cases "\x i x \\ e_nth e x") + case True + then show ?thesis + using that inconsist_converg[OF assms] + by (metis (mono_tags, lifting) not_less_Least option.inject) + next + case False + then show ?thesis + using that inconsist_converg[OF assms] by simp + qed + show "\x i x \= e_nth e x \ inconsist i e \= e_length e" + unfolding inconsist_def using assms by auto +qed + +lemma inconsist_converg_eq: + assumes "inconsist i e \= e_length e" + shows "\x i x \= e_nth e x" + using assms inconsist_consistent by auto + +lemma inconsist_converg_less: + assumes "inconsist i e \" and "the (inconsist i e) < e_length e" + shows "\x i x \\ e_nth e x" + and "inconsist i e \= (LEAST x. x < e_length e \ \ i x \\ e_nth e x)" +proof - + show "\x i x \\ e_nth e x" + using assms by (metis (no_types, lifting) inconsist_converg(1) nat_neq_iff option.sel) + then show "inconsist i e \= (LEAST x. x < e_length e \ \ i x \\ e_nth e x)" + using assms inconsist_converg by presburger +qed + +lemma least_bounded_Suc: + assumes "\x. x < upper \ P x" + shows "(LEAST x. x < upper \ P x) = (LEAST x. x < Suc upper \ P x)" +proof - + let ?Q = "\x. x < upper \ P x" + let ?x = "Least ?Q" + from assms have "?x < upper \ P ?x" + using LeastI_ex[of ?Q] by simp + then have 1: "?x < Suc upper \ P ?x" by simp + from assms have 2: "\y P y" + using Least_le[of ?Q] not_less_Least by fastforce + have "(LEAST x. x < Suc upper \ P x) = ?x" + proof (rule Least_equality) + show "?x < Suc upper \ P ?x" using 1 2 by blast + show "\y. y < Suc upper \ P y \ ?x \ y" + using 1 2 leI by blast + qed + then show ?thesis .. +qed + +lemma least_bounded_gr: + fixes P :: "nat \ bool" and m :: nat + assumes "\x. x < upper \ P x" + shows "(LEAST x. x < upper \ P x) = (LEAST x. x < upper + m \ P x)" +proof (induction m) + case 0 + then show ?case by simp +next + case (Suc m) + moreover have "\x. x < upper + m \ P x" + using assms trans_less_add1 by blast + ultimately show ?case using least_bounded_Suc by simp +qed + +lemma inconsist_init_converg_less: + assumes "f \ \" + and "\ i \ \" + and "inconsist i (f \ n) \" + and "the (inconsist i (f \ n)) < Suc n" + shows "inconsist i (f \ (n + m)) = inconsist i (f \ n)" +proof - + have phi_i_total: "\ i x \" for x + using assms by simp + moreover have f_nth: "f x \= e_nth (f \ n) x" if "x < Suc n" for x n + using that assms(1) by simp + ultimately have "(\ i x \ f x) = (\ i x \\ e_nth (f \ n) x)" if "x < Suc n" for x n + using that by simp + then have cond: "(x < Suc n \ \ i x \ f x) = + (x < e_length (f \ n) \ \ i x \\ e_nth (f \ n) x)" for x n + using length_init by metis + then have + 1: "\x i x \ f x" and + 2: "inconsist i (f \ n) \= (LEAST x. x < Suc n \ \ i x \ f x)" + using assms(3,4) inconsist_converg_less[of i "f \ n"] by simp_all + then have 3: "\x i x \ f x" + using not_add_less1 by fastforce + then have "\x i x \\ e_nth (f \ (n + m)) x" + using cond by blast + then have "\x (n + m)). \ i x \\ e_nth (f \ (n + m)) x" + by simp + moreover have 4: "inconsist i (f \ (n + m)) \" + using assms(2) R1_imp_total1 inconsist_def by simp + ultimately have "inconsist i (f \ (n + m)) \= + (LEAST x. x < e_length (f \ (n + m)) \ \ i x \\ e_nth (f \ (n + m)) x)" + using inconsist_converg[OF 4] by simp + then have 5: "inconsist i (f \ (n + m)) \= (LEAST x. x < Suc (n + m) \ \ i x \ f x)" + using cond[of _ "n + m"] by simp + then have "(LEAST x. x < Suc n \ \ i x \ f x) = + (LEAST x. x < Suc n + m \ \ i x \ f x)" + using least_bounded_gr[where ?upper="Suc n"] 1 3 by simp + then show ?thesis using 2 5 by simp +qed + +definition "r_inconsist \ + let + f = Cn 2 r_length [Id 2 1]; + g = Cn 4 r_ifless + [Id 4 1, + Cn 4 r_length [Id 4 3], + Id 4 1, + Cn 4 r_ifeq + [Cn 4 r_phi [Id 4 2, Id 4 0], + Cn 4 r_nth [Id 4 3, Id 4 0], + Id 4 1, + Id 4 0]] + in Cn 2 (Pr 2 f g) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]" + +lemma r_inconsist_recfn: "recfn 2 r_inconsist" + unfolding r_inconsist_def by simp + +lemma r_inconsist: "eval r_inconsist [i, e] = inconsist i e" +proof - + define f where "f = Cn 2 r_length [Id 2 1]" + define len where "len = Cn 4 r_length [Id 4 3]" + define nth where "nth = Cn 4 r_nth [Id 4 3, Id 4 0]" + define ph where "ph = Cn 4 r_phi [Id 4 2, Id 4 0]" + define g where + "g = Cn 4 r_ifless [Id 4 1, len, Id 4 1, Cn 4 r_ifeq [ph, nth, Id 4 1, Id 4 0]]" + have "recfn 2 f" + unfolding f_def by simp + have f: "eval f [i, e] \= e_length e" + unfolding f_def by simp + have "recfn 4 len" + unfolding len_def by simp + have len: "eval len [j, v, i, e] \= e_length e" for j v + unfolding len_def by simp + have "recfn 4 nth" + unfolding nth_def by simp + have nth: "eval nth [j, v, i, e] \= e_nth e j" for j v + unfolding nth_def by simp + have "recfn 4 ph" + unfolding ph_def by simp + have ph: "eval ph [j, v, i, e] = \ i j" for j v + unfolding ph_def using phi_def by simp + have "recfn 4 g" + unfolding g_def using `recfn 4 nth` `recfn 4 ph` `recfn 4 len` by simp + have g_diverg: "eval g [j, v, i, e] \" if "eval ph [j, v, i, e] \" for j v + unfolding g_def using that `recfn 4 nth` `recfn 4 ph` `recfn 4 len` by simp + have g_converg: "eval g [j, v, i, e] \= + (if v < e_length e then v else if \ i j \= e_nth e j then v else j)" + if "eval ph [j, v, i, e] \" for j v + unfolding g_def using that `recfn 4 nth` `recfn 4 ph` `recfn 4 len` len nth ph + by auto + define h where "h \ Pr 2 f g" + then have "recfn 3 h" + by (simp add: \recfn 2 f\ \recfn 4 g\) + + let ?invariant = "\j i e. + (if \x i x \ then None + else if \x i x \\ e_nth e x + then Some (LEAST x. x < j \ \ i x \\ e_nth e x) + else Some (e_length e))" + + have "eval h [j, i, e] = ?invariant j i e" if "j \ e_length e" for j + using that + proof (induction j) + case 0 + then show ?case unfolding h_def using `recfn 2 f` f `recfn 4 g` by simp + next + case (Suc j) + then have j_less: "j < e_length e" by simp + then have j_le: "j \ e_length e" by simp + show ?case + proof (cases "eval h [j, i, e] \") + case True + then have "\x i x \" + using j_le Suc.IH by (metis option.simps(3)) + then have "\x i x \" + using less_SucI by blast + moreover have h: "eval h [Suc j, i, e] \" + using True h_def `recfn 3 h` by simp + ultimately show ?thesis by simp + next + case False + with Suc.IH j_le have h_j: "eval h [j, i, e] = + (if \x i x \\ e_nth e x + then Some (LEAST x. x < j \ \ i x \\ e_nth e x) + else Some (e_length e))" + by presburger + then have the_h_j: "the (eval h [j, i, e]) = + (if \x i x \\ e_nth e x + then LEAST x. x < j \ \ i x \\ e_nth e x + else e_length e)" + (is "_ = ?v") + by auto + have h_Suc: "eval h [Suc j, i, e] = eval g [j, the (eval h [j, i, e]), i, e]" + using False h_def `recfn 4 g` `recfn 2 f` by auto + show ?thesis + proof (cases "\ i j \") + case True + with ph g_diverg h_Suc show ?thesis by auto + next + case False + with h_Suc have "eval h [Suc j, i, e] \= + (if ?v < e_length e then ?v + else if \ i j \= e_nth e j then ?v else j)" + (is "_ \= ?lhs") + using g_converg ph the_h_j by simp + moreover have "?invariant (Suc j) i e \= + (if \x i x \\ e_nth e x + then LEAST x. x < Suc j \ \ i x \\ e_nth e x + else e_length e)" + (is "_ \= ?rhs") + proof - + from False have "\ i j \" by simp + moreover have "\ (\x i x \)" + by (metis (no_types, lifting) Suc.IH h_j j_le option.simps(3)) + ultimately have "\ (\x i x \)" + using less_Suc_eq by auto + then show ?thesis by auto + qed + moreover have "?lhs = ?rhs" + proof (cases "?v < e_length e") + case True + then have + ex_j: "\x i x \\ e_nth e x" and + v_eq: "?v = (LEAST x. x < j \ \ i x \\ e_nth e x)" + by presburger+ + with True have "?lhs = ?v" by simp + from ex_j have "\x i x \\ e_nth e x" + using less_SucI by blast + then have "?rhs = (LEAST x. x < Suc j \ \ i x \\ e_nth e x)" by simp + with True v_eq ex_j show ?thesis + using least_bounded_Suc[of j "\x. \ i x \\ e_nth e x"] by simp + next + case False + then have not_ex: "\ (\x i x \\ e_nth e x)" + using Least_le[of "\x. x < j \ \ i x \\ e_nth e x"] j_le + by (smt leD le_less_linear le_trans) + then have "?v = e_length e" by argo + with False have lhs: "?lhs = (if \ i j \= e_nth e j then e_length e else j)" + by simp + show ?thesis + proof (cases "\ i j \= e_nth e j") + case True + then have "\ (\x i x \\ e_nth e x)" + using less_SucE not_ex by blast + then have "?rhs = e_length e" by argo + moreover from True have "?lhs = e_length e" + using lhs by simp + ultimately show ?thesis by simp + next case False + then have "\ i j \\ e_nth e j" + using `\ i j \` by simp + with not_ex have "(LEAST x. x \ i x \\ e_nth e x) = j" + using LeastI[of "\x. x \ i x \\ e_nth e x" j] less_Suc_eq + by blast + then have "?rhs = j" + using \\ i j \\ e_nth e j\ by (meson lessI) + moreover from False lhs have "?lhs = j" by simp + ultimately show ?thesis by simp + qed + qed + ultimately show ?thesis by simp + qed + qed + qed + then have "eval h [e_length e, i, e] = ?invariant (e_length e) i e" + by auto + then have "eval h [e_length e, i, e] = inconsist i e" + using inconsist_def by simp + moreover have "eval (Cn 2 (Pr 2 f g) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]) [i, e] = + eval h [e_length e, i, e]" + using `recfn 4 g` `recfn 2 f` h_def by auto + ultimately show ?thesis + unfolding r_inconsist_def by (simp add: f_def g_def len_def nth_def ph_def) +qed + +lemma inconsist_for_total: + assumes "total1 (\ i)" + shows "inconsist i e \= + (if \x i x \\ e_nth e x + then LEAST x. x < e_length e \ \ i x \\ e_nth e x + else e_length e)" + unfolding inconsist_def using assms total1_def by (auto; blast) + +lemma inconsist_for_V01: + assumes "f \ V\<^sub>0\<^sub>1" and "k = amalgamate (the (f 0)) (the (f 1))" + shows "inconsist k e \= + (if \x k x \\ e_nth e x + then LEAST x. x < e_length e \ \ k x \\ e_nth e x + else e_length e)" +proof - + have "\ k \ \" + using amalgamation_V01_R1[OF assms(1)] assms(2) amalgamate by simp + then have "total1 (\ k)" by simp + with inconsist_for_total[of k] show ?thesis by simp +qed + +text \The next function computes Gödel numbers of functions consistent +with a given prefix. The strategy will use these as consistent auxiliary +hypotheses when receiving a prefix of length one.\ + +definition "r_auxhyp \ Cn 1 (r_smn 1 1) [r_const (encode r_prenum), Id 1 0]" + +lemma r_auxhyp_prim: "prim_recfn 1 r_auxhyp" + unfolding r_auxhyp_def by simp + +lemma r_auxhyp: "\ (the (eval r_auxhyp [e])) = prenum e" +proof + fix x + let ?p = "encode r_prenum" + let ?p = "encode r_prenum" + have "eval r_auxhyp [e] = eval (r_smn 1 1) [?p, e]" + unfolding r_auxhyp_def by simp + then have "eval r_auxhyp [e] \= smn 1 ?p [e]" + by (simp add: r_smn) + also have "... \= encode (Cn 1 (r_universal (1 + length [e])) + (r_constn (1 - 1) ?p # + map (r_constn (1 - 1)) [e] @ map (recf.Id 1) [0..<1]))" + using smn[of 1 ?p "[e]"] by simp + also have "... \= encode (Cn 1 (r_universal (1 + 1)) + (r_constn 0 ?p # map (r_constn 0) [e] @ [Id 1 0]))" + by simp + also have "... \= encode (Cn 1 (r_universal 2) + (r_constn 0 ?p # map (r_constn 0) [e] @ [Id 1 0]))" + by (metis one_add_one) + also have "... \= encode (Cn 1 (r_universal 2) [r_constn 0 ?p, r_constn 0 e, Id 1 0])" + by simp + also have "... \= encode (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0])" + using r_constn_def by simp + finally have "eval r_auxhyp [e] \= + encode (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0])" . + moreover have "\ (the (eval r_auxhyp [e])) x = eval r_phi [the (eval r_auxhyp [e]), x]" + by (simp add: phi_def) + ultimately have "\ (the (eval r_auxhyp [e])) x = + eval r_phi [encode (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0]), x]" + (is "_ = eval r_phi [encode ?f, x]") + by simp + then have "\ (the (eval r_auxhyp [e])) x = + eval (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0]) [x]" + using r_phi_def r_universal[of ?f 1 "[x]"] by simp + then have "\ (the (eval r_auxhyp [e])) x = eval (r_universal 2) [?p, e, x]" + by simp + then have "\ (the (eval r_auxhyp [e])) x = eval r_prenum [e, x]" + using r_universal by simp + then show "\ (the (eval r_auxhyp [e])) x = prenum e x" by simp +qed + +definition auxhyp :: partial1 where + "auxhyp e \ eval r_auxhyp [e]" + +lemma auxhyp_prenum: "\ (the (auxhyp e)) = prenum e" + using auxhyp_def r_auxhyp by metis + +lemma auxhyp_in_R1: "auxhyp \ \" + using auxhyp_def Mn_free_imp_total R1I r_auxhyp_prim by metis + +text \Now we can define our consistent learning strategy for @{term "V\<^sub>0\<^sub>1"}.\ + +definition "r_sv01 \ + let + at0 = Cn 1 r_nth [Id 1 0, Z]; + at1 = Cn 1 r_nth [Id 1 0, r_const 1]; + m = Cn 1 r_amalgamate [at0, at1]; + c = Cn 1 r_inconsist [m, Id 1 0]; + p = Cn 1 r_pdec1 [Cn 1 r_parallel [at0, at1, c]]; + g = Cn 1 r_ifeq [c, r_length, m, Cn 1 r_ifz [p, at1, at0]] + in Cn 1 (r_lifz r_auxhyp g) [Cn 1 r_eq [r_length, r_const 1], Id 1 0]" + +lemma r_sv01_recfn: "recfn 1 r_sv01" + unfolding r_sv01_def using r_auxhyp_prim r_inconsist_recfn r_amalgamate_recfn + by (simp add: Let_def) + +definition sv01 :: partial1 ("s\<^bsub>01\<^esub>")where + "sv01 e \ eval r_sv01 [e]" + +lemma sv01_in_P1: "s\<^bsub>01\<^esub> \ \

" + using sv01_def r_sv01_recfn P1I by presburger + +text \We are interested in the behavior of @{term "s\<^bsub>01\<^esub>"} only on +prefixes of functions in @{term "V\<^sub>0\<^sub>1"}. This behavior is linked +to the amalgamation of $f(0)$ and $f(1)$, where $f$ is the function +to be learned.\ + +abbreviation amalg01 :: "partial1 \ nat" where + "amalg01 f \ amalgamate (the (f 0)) (the (f 1))" + +lemma sv01: + assumes "f \ V\<^sub>0\<^sub>1" + shows "s\<^bsub>01\<^esub> (f \ 0) = auxhyp (f \ 0)" + and "n \ 0 \ + inconsist (amalg01 f) (f \ n) \= Suc n \ + s\<^bsub>01\<^esub> (f \ n) \= amalg01 f" + and "n \ 0 \ + the (inconsist (amalg01 f) (f \ n)) < Suc n \ + pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist (amalg01 f) (f \ n))))) = 0 \ + s\<^bsub>01\<^esub> (f \ n) = f 1" + and "n \ 0 \ + the (inconsist (amalg01 f) (f \ n)) < Suc n \ + pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist (amalg01 f) (f \ n))))) \ 0 \ + s\<^bsub>01\<^esub> (f \ n) = f 0" +proof - + have f_total: "\x. f x \" + using assms V01_def R1_imp_total1 by blast + define at0 where "at0 = Cn 1 r_nth [Id 1 0, Z]" + define at1 where "at1 = Cn 1 r_nth [Id 1 0, r_const 1]" + define m where "m = Cn 1 r_amalgamate [at0, at1]" + define c where "c = Cn 1 r_inconsist [m, Id 1 0]" + define p where "p = Cn 1 r_pdec1 [Cn 1 r_parallel [at0, at1, c]]" + define g where "g = Cn 1 r_ifeq [c, r_length, m, Cn 1 r_ifz [p, at1, at0]]" + have "recfn 1 g" + unfolding g_def p_def c_def m_def at1_def at0_def + using r_auxhyp_prim r_inconsist_recfn r_amalgamate_recfn + by simp + have "eval (Cn 1 r_eq [r_length, r_const 1]) [f \ 0] \= 0" + by simp + then have "eval r_sv01 [f \ 0] = eval r_auxhyp [f \ 0]" + unfolding r_sv01_def using `recfn 1 g` c_def g_def m_def p_def r_auxhyp_prim + by (auto simp add: Let_def) + then show "s\<^bsub>01\<^esub> (f \ 0) = auxhyp (f \ 0)" + by (simp add: auxhyp_def sv01_def) + + have sv01: "s\<^bsub>01\<^esub> (f \ n) = eval g [f \ n]" if "n \ 0" + proof - + have *: "eval (Cn 1 r_eq [r_length, r_const 1]) [f \ n] \\ 0" + (is "?r_eq \\ 0") + using that by simp + moreover have "recfn 2 (r_lifz r_auxhyp g)" + using `recfn 1 g` r_auxhyp_prim by simp + moreover have "eval r_sv01 [f \ n] = + eval (Cn 1 (r_lifz r_auxhyp g) [Cn 1 r_eq [r_length, r_const 1], Id 1 0]) [f \ n]" + using r_sv01_def by (metis at0_def at1_def c_def g_def m_def p_def) + ultimately have "eval r_sv01 [f \ n] = eval (r_lifz r_auxhyp g) [the ?r_eq, f \ n]" + by simp + then have "eval r_sv01 [f \ n] = eval g [f \ n]" + using "*" \recfn 1 g\ r_auxhyp_prim by auto + then show ?thesis by (simp add: sv01_def that) + qed + + have "recfn 1 at0" + unfolding at0_def by simp + have at0: "eval at0 [f \ n] \= the (f 0)" + unfolding at0_def by simp + have "recfn 1 at1" + unfolding at1_def by simp + have at1: "n \ 0 \ eval at1 [f \ n] \= the (f 1)" + unfolding at1_def by simp + have "recfn 1 m" + unfolding m_def at0_def at1_def using r_amalgamate_recfn by simp + have m: "n \ 0 \ eval m [f \ n] \= amalg01 f" + (is "_ \ _ \= ?m") + unfolding m_def at0_def at1_def + using at0 at1 amalgamate r_amalgamate r_amalgamate_recfn by simp + then have c: "n \ 0 \ eval c [f \ n] = inconsist (amalg01 f) (f \ n)" + (is "_ \ _ = ?c") + unfolding c_def using r_inconsist_recfn `recfn 1 m` r_inconsist by auto + then have c_converg: "n \ 0 \ eval c [f \ n] \" + using inconsist_for_V01[OF assms] by simp + have "recfn 1 c" + unfolding c_def using `recfn 1 m` r_inconsist_recfn by simp + + have par: "n \ 0 \ + eval (Cn 1 r_parallel [at0, at1, c]) [f \ n] = parallel (the (f 0)) (the (f 1)) (the ?c)" + (is "_ \ _ = ?par") + using at0 at1 c c_converg m r_parallel' `recfn 1 at0` `recfn 1 at1` `recfn 1 c` + by simp + with parallel_converg_V01[OF assms] have + par_converg: "n \ 0 \ eval (Cn 1 r_parallel [at0, at1, c]) [f \ n] \" + by simp + then have p_converg: "n \ 0 \ eval p [f \ n] \" + unfolding p_def using at0 at1 c_converg `recfn 1 at0` `recfn 1 at1` `recfn 1 c` + by simp + have p: "n \ 0 \ eval p [f \ n] \= pdec1 (the ?par)" + unfolding p_def + using at0 at1 c_converg `recfn 1 at0` `recfn 1 at1` `recfn 1 c` par par_converg + by simp + have "recfn 1 p" + unfolding p_def using `recfn 1 at0` `recfn 1 at1` `recfn 1 m` `recfn 1 c` + by simp + + let ?r = "Cn 1 r_ifz [p, at1, at0]" + have r: "n \ 0 \ eval ?r [f \ n] = (if pdec1 (the ?par) = 0 then f 1 else f 0)" + using at0 at1 c_converg `recfn 1 at0` `recfn 1 at1` `recfn 1 c` + `recfn 1 m` `recfn 1 p` p f_total + by fastforce + + have g: "n \ 0 \ + eval g [f \ n] \= + (if the ?c = e_length (f \ n) + then ?m else the (eval (Cn 1 r_ifz [p, at1, at0]) [f \ n]))" + unfolding g_def + using `recfn 1 p` `recfn 1 at0` `recfn 1 at1` `recfn 1 c` `recfn 1 m` + p_converg at1 at0 c c_converg m + by simp + { + assume "n \ 0" and "?c \= Suc n" + moreover have "e_length (f \ n) = Suc n" by simp + ultimately have "eval g [f \ n] \= ?m" using g by simp + then show "s\<^bsub>01\<^esub> (f \ n) \= amalg01 f" + using sv01[OF `n \ 0`] by simp + next + assume "n \ 0" and "the ?c < Suc n" and "pdec1 (the ?par) = 0" + with g r f_total have "eval g [f \ n] = f 1" by simp + then show "s\<^bsub>01\<^esub> (f \ n) = f 1" + using sv01[OF `n \ 0`] by simp + next + assume "n \ 0" and "the ?c < Suc n" and "pdec1 (the ?par) \ 0" + with g r f_total have "eval g [f \ n] = f 0" by simp + then show "s\<^bsub>01\<^esub> (f \ n) = f 0" + using sv01[OF `n \ 0`] by simp + } +qed + +text \Part of the correctness of @{term sv01} is convergence on +prefixes of functions in @{term "V\<^sub>0\<^sub>1"}.\ + +lemma sv01_converg_V01: + assumes "f \ V\<^sub>0\<^sub>1" + shows "s\<^bsub>01\<^esub> (f \ n) \" +proof (cases "n = 0") + case True + then show ?thesis + using assms sv01 R1_imp_total1 auxhyp_in_R1 by simp +next + case n_gr_0: False + show ?thesis + proof (cases "inconsist (amalg01 f) (f \ n) \= Suc n") + case True + then show ?thesis + using n_gr_0 assms sv01 by simp + next + case False + then have "the (inconsist (amalg01 f) (f \ n)) < Suc n" + using assms inconsist_bounded inconsist_for_V01 length_init + by (metis (no_types, lifting) le_neq_implies_less option.collapse option.simps(3)) + then show ?thesis + using n_gr_0 assms sv01 R1_imp_total1 total1E V01_def + by (metis (no_types, lifting) mem_Collect_eq) + qed +qed + +text \Another part of the correctness of @{term sv01} is its hypotheses +being consistent on prefixes of functions in @{term "V\<^sub>0\<^sub>1"}.\ + +lemma sv01_consistent_V01: + assumes "f \ V\<^sub>0\<^sub>1" + shows "\x\n. \ (the (s\<^bsub>01\<^esub> (f \ n))) x = f x" +proof (cases "n = 0") + case True + then have "s\<^bsub>01\<^esub> (f \ n) = auxhyp (f \ n)" + using sv01[OF assms] by simp + then have "\ (the (s\<^bsub>01\<^esub> (f \ n))) = prenum (f \ n)" + using auxhyp_prenum by simp + then show ?thesis + using R1_imp_total1 total1E assms by (simp add: V01_def) +next + case n_gr_0: False + let ?m = "amalg01 f" + let ?e = "f \ n" + let ?c = "the (inconsist ?m ?e)" + have c: "inconsist ?m ?e \" + using assms inconsist_for_V01 by blast + show ?thesis + proof (cases "inconsist ?m ?e \= Suc n") + case True + then show ?thesis + using assms n_gr_0 sv01 R1_imp_total1 total1E V01_def is_init_of_def + inconsist_consistent not_initial_imp_not_eq length_init inconsist_converg_eq + by (metis (no_types, lifting) le_imp_less_Suc mem_Collect_eq option.sel) + next + case False + then have less: "the (inconsist ?m ?e) < Suc n" + using c assms inconsist_bounded inconsist_for_V01 length_init + by (metis le_neq_implies_less option.collapse) + then have "the (inconsist ?m ?e) < e_length ?e" + by auto + then have + "\x ?m x \\ e_nth ?e x" + "inconsist ?m ?e \= (LEAST x. x < e_length ?e \ \ ?m x \\ e_nth ?e x)" + (is "_ \= Least ?P") + using inconsist_converg_less[OF c] by simp_all + then have "?P ?c" and "\x. x < ?c \ \ ?P x" + using LeastI_ex[of ?P] not_less_Least[of _ ?P] by (auto simp del: e_nth) + then have "\ ?m ?c \ f ?c" by auto + then have "amalgamation (the (f 0)) (the (f 1)) ?c \ f ?c" + using amalgamate by simp + then have *: "Some (pdec2 (the (parallel (the (f 0)) (the (f 1)) ?c))) \ f ?c" + using amalgamation_def by (metis assms parallel_converg_V01) + let ?p = "parallel (the (f 0)) (the (f 1)) ?c" + show ?thesis + proof (cases "pdec1 (the ?p) = 0") + case True + then have "\ (the (f 0)) ?c \= pdec2 (the ?p)" + using assms parallel_0 parallel_converg_V01 + by (metis option.collapse prod.collapse prod_decode_inverse) + then have "\ (the (f 0)) ?c \ f ?c" + using * by simp + then have "\ (the (f 0)) \ f" by auto + then have "\ (the (f 1)) = f" + using assms V01_def by auto + moreover have "s\<^bsub>01\<^esub> (f \ n) = f 1" + using True less n_gr_0 sv01 assms by simp + ultimately show ?thesis by simp + next + case False + then have "pdec1 (the ?p) = 1" + by (meson assms parallel_converg_V01 parallel_converg_pdec1_0_or_1) + then have "\ (the (f 1)) ?c \= pdec2 (the ?p)" + using assms parallel_1 parallel_converg_V01 + by (metis option.collapse prod.collapse prod_decode_inverse) + then have "\ (the (f 1)) ?c \ f ?c" + using * by simp + then have "\ (the (f 1)) \ f" by auto + then have "\ (the (f 0)) = f" + using assms V01_def by auto + moreover from False less n_gr_0 sv01 assms have "s\<^bsub>01\<^esub> (f \ n) = f 0" + by simp + ultimately show ?thesis by simp + qed + qed +qed + +text \The final part of the correctness is @{term "sv01"} converging +for all functions in @{term "V\<^sub>0\<^sub>1"}.\ + +lemma sv01_limit_V01: + assumes "f \ V\<^sub>0\<^sub>1" + shows "\i. \\<^sup>\n. s\<^bsub>01\<^esub> (f \ n) \= i" +proof (cases "\n>0. s\<^bsub>01\<^esub> (f \ n) \= amalgamate (the (f 0)) (the (f 1))") + case True + then show ?thesis by (meson less_le_trans zero_less_one) +next + case False + then obtain n\<^sub>0 where n0: + "n\<^sub>0 \ 0" + "s\<^bsub>01\<^esub> (f \ n\<^sub>0) \\ amalg01 f" + using \f \ V\<^sub>0\<^sub>1\ sv01_converg_V01 by blast + then have *: "the (inconsist (amalg01 f) (f \ n\<^sub>0)) < Suc n\<^sub>0" + (is "the (inconsist ?m (f \ n\<^sub>0)) < Suc n\<^sub>0") + using assms `n\<^sub>0 \ 0` sv01(2) inconsist_bounded inconsist_for_V01 length_init + by (metis (no_types, lifting) le_neq_implies_less option.collapse option.simps(3)) + moreover have "f \ \" + using assms V01_def by auto + moreover have "\ ?m \ \" + using amalgamate amalgamation_V01_R1 assms by auto + moreover have "inconsist ?m (f \ n\<^sub>0) \" + using inconsist_for_V01 assms by blast + ultimately have **: "inconsist ?m (f \ (n\<^sub>0 + m)) = inconsist ?m (f \ n\<^sub>0)" for m + using inconsist_init_converg_less[of f ?m] by simp + then have "the (inconsist ?m (f \ (n\<^sub>0 + m))) < Suc n\<^sub>0 + m" for m + using * by auto + moreover have + "pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist ?m (f \ (n\<^sub>0 + m)))))) = + pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist ?m (f \ n\<^sub>0)))))" + for m + using ** by auto + moreover have "n\<^sub>0 + m \ 0" for m + using `n\<^sub>0 \ 0` by simp + ultimately have "s\<^bsub>01\<^esub> (f \ (n\<^sub>0 + m)) = s\<^bsub>01\<^esub> (f \ n\<^sub>0)" for m + using assms sv01 * \n\<^sub>0 \ 0\ by (metis add_Suc) + moreover define i where "i = s\<^bsub>01\<^esub> (f \ n\<^sub>0)" + ultimately have "\n\n\<^sub>0. s\<^bsub>01\<^esub> (f \ n) = i" + using nat_le_iff_add by auto + then have "\n\n\<^sub>0. s\<^bsub>01\<^esub> (f \ n) \= the i" + using n0(2) by simp + then show ?thesis by auto +qed + +lemma V01_learn_cons: "learn_cons \ V\<^sub>0\<^sub>1 s\<^bsub>01\<^esub>" +proof (rule learn_consI2) + show "environment \ V\<^sub>0\<^sub>1 s\<^bsub>01\<^esub>" + by (simp add: Collect_mono V01_def phi_in_P2 sv01_in_P1 sv01_converg_V01) + show "\f n. f \ V\<^sub>0\<^sub>1 \ \k\n. \ (the (s\<^bsub>01\<^esub> (f \ n))) k = f k" + using sv01_consistent_V01 . + show "\i n\<^sub>0. \n\n\<^sub>0. s\<^bsub>01\<^esub> (f \ n) \= i" if "f \ V\<^sub>0\<^sub>1" for f + using sv01_limit_V01 that by simp +qed + +corollary V01_in_CONS: "V\<^sub>0\<^sub>1 \ CONS" + using V01_learn_cons CONS_def by auto + +text \Now we can show the main result of this section, namely that +there is a consistently learnable class that cannot be learned consistently +by a total strategy. In other words, there is no Lemma~R for CONS.\ + +lemma no_lemma_R_for_CONS: "\U. U \ CONS \ (\ (\s. s \ \ \ learn_cons \ U s))" + using V01_in_CONS V01_not_in_R_cons by auto + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/Partial_Recursive.thy b/thys/Inductive_Inference/Partial_Recursive.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/Partial_Recursive.thy @@ -0,0 +1,1914 @@ +chapter \Partial recursive functions\ + +theory Partial_Recursive + imports Main "HOL-Library.Nat_Bijection" +begin + +text \This chapter lays the foundation for Chapter~\ref{c:iirf}. +Essentially it develops recursion theory up to the point of certain +fixed-point theorems. This in turn requires standard results such as the +existence of a universal function and the $s$-$m$-$n$ theorem. Besides these, +the chapter contains some results, mostly regarding decidability and the +Kleene normal form, that are not strictly needed later. They are included as +relatively low-hanging fruits to round off the chapter. + +The formalization of partial recursive functions is very much inspired by the +Universal Turing Machine AFP entry by Xu +et~al.~\cite{Universal_Turing_Machine-AFP}. It models partial recursive +functions as algorithms whose semantics is given by an evaluation function. +This works well for most of this chapter. For the next chapter, however, it +is beneficial to regard partial recursive functions as ``proper'' partial +functions. To that end, Section~\ref{s:alternative} introduces more +conventional and convenient notation for the common special cases of unary +and binary partial recursive functions. + +Especially for the nontrivial proofs I consulted the classical textbook by +Rogers~\cite{Rogers87}, which also partially explains my preferring the +traditional term ``recursive'' to the more modern ``computable''.\ + + +section \Basic definitions\ + + +subsection \Partial recursive functions\ + +text \To represent partial recursive functions we start from the same +datatype as Xu et~al.~\cite{Universal_Turing_Machine-AFP}, more specifically +from Urban's version of the formalization. In fact the datatype @{text recf} +and the function @{text arity} below have been copied verbatim from it.\ + +datatype recf = + Z +| S +| Id nat nat +| Cn nat recf "recf list" +| Pr nat recf recf +| Mn nat recf + +fun arity :: "recf \ nat" where + "arity Z = 1" +| "arity S = 1" +| "arity (Id m n) = m" +| "arity (Cn n f gs) = n" +| "arity (Pr n f g) = Suc n" +| "arity (Mn n f) = n" + +text \Already we deviate from Xu et~al.\ in that we define a +well-formedness predicate for partial recursive functions. Well-formedness +essentially means arity constraints are respected when combining @{typ +recf}s.\ + +fun wellf :: "recf \ bool" where + "wellf Z = True" +| "wellf S = True" +| "wellf (Id m n) = (n < m)" +| "wellf (Cn n f gs) = + (n > 0 \ (\g \ set gs. arity g = n \ wellf g) \ arity f = length gs \ wellf f)" +| "wellf (Pr n f g) = + (arity g = Suc (Suc n) \ arity f = n \ wellf f \ wellf g)" +| "wellf (Mn n f) = (n > 0 \ arity f = Suc n \ wellf f)" + +lemma wellf_arity_nonzero: "wellf f \ arity f > 0" + by (induction f rule: arity.induct) simp_all + +lemma wellf_Pr_arity_greater_1: "wellf (Pr n f g) \ arity (Pr n f g) > 1" + using wellf_arity_nonzero by auto + +text \For the most part of this chapter this is the meaning of ``$f$ +is an $n$-ary partial recursive function'':\ + +abbreviation recfn :: "nat \ recf \ bool" where + "recfn n f \ wellf f \ arity f = n" + +text \Some abbreviations for working with @{typ "nat option"}:\ + +abbreviation divergent :: "nat option \ bool" ("_ \" [50] 50) where + "x \ \ x = None" + +abbreviation convergent :: "nat option \ bool" ("_ \" [50] 50) where + "x \ \ x \ None" + +abbreviation convergent_eq :: "nat option \ nat \ bool" (infix "\=" 50) where + "x \= y \ x = Some y" + +abbreviation convergent_neq :: "nat option \ nat \ bool" (infix "\\" 50) where + "x \\ y \ x \ \ x \ Some y" + +text \In prose the terms ``halt'', ``terminate'', ``converge'', and +``defined'' will be used interchangeably; likewise for ``not halt'', +``diverge'', and ``undefined''. In names of lemmas, the abbreviations @{text +converg} and @{text diverg} will be used consistently.\ + +text \Our second major deviation from Xu +et~al.~\cite{Universal_Turing_Machine-AFP} is that we model the semantics of +a @{typ recf} by combining the value and the termination of a function into +one evaluation function with codomain @{typ "nat option"}, rather than +separating both aspects into an evaluation function with codomain @{typ nat} +and a termination predicate. + +The value of a well-formed partial recursive function applied to a +correctly-sized list of arguments:\ + +fun eval_wellf :: "recf \ nat list \ nat option" where + "eval_wellf Z xs \= 0" +| "eval_wellf S xs \= Suc (xs ! 0)" +| "eval_wellf (Id m n) xs \= xs ! n" +| "eval_wellf (Cn n f gs) xs = + (if \g \ set gs. eval_wellf g xs \ + then eval_wellf f (map (\g. the (eval_wellf g xs)) gs) + else None)" +| "eval_wellf (Pr n f g) [] = undefined" +| "eval_wellf (Pr n f g) (0 # xs) = eval_wellf f xs" +| "eval_wellf (Pr n f g) (Suc x # xs) = + Option.bind (eval_wellf (Pr n f g) (x # xs)) (\v. eval_wellf g (x # v # xs))" +| "eval_wellf (Mn n f) xs = + (let E = \z. eval_wellf f (z # xs) + in if \z. E z \= 0 \ (\y) + then Some (LEAST z. E z \= 0 \ (\y)) + else None)" + +text \We define a function value only if the @{typ recf} is well-formed +and its arity matches the number of arguments.\ + +definition eval :: "recf \ nat list \ nat option" where + "recfn (length xs) f \ eval f xs \ eval_wellf f xs" + +lemma eval_Z [simp]: "eval Z [x] \= 0" + by (simp add: eval_def) + +lemma eval_Z' [simp]: "length xs = 1 \ eval Z xs \= 0" + by (simp add: eval_def) + +lemma eval_S [simp]: "eval S [x] \= Suc x" + by (simp add: eval_def) + +lemma eval_S' [simp]: "length xs = 1 \ eval S xs \= Suc (hd xs)" + using eval_def hd_conv_nth[of xs] by fastforce + +lemma eval_Id [simp]: + assumes "n < m" and "m = length xs" + shows "eval (Id m n) xs \= xs ! n" + using assms by (simp add: eval_def) + +lemma eval_Cn [simp]: + assumes "recfn (length xs) (Cn n f gs)" + shows "eval (Cn n f gs) xs = + (if \g\set gs. eval g xs \ + then eval f (map (\g. the (eval g xs)) gs) + else None)" +proof - + have "eval (Cn n f gs) xs = eval_wellf (Cn n f gs) xs" + using assms eval_def by blast + moreover have "\g. g \ set gs \ eval_wellf g xs = eval g xs" + using assms eval_def by simp + ultimately have "eval (Cn n f gs) xs = + (if \g\set gs. eval g xs \ + then eval_wellf f (map (\g. the (eval g xs)) gs) + else None)" + using map_eq_conv[of "\g. the (eval_wellf g xs)" gs "\g. the (eval g xs)"] + by (auto, metis) + moreover have "\ys. length ys = length gs \ eval f ys = eval_wellf f ys" + using assms eval_def by simp + ultimately show ?thesis by auto +qed + +lemma eval_Pr_0 [simp]: + assumes "recfn (Suc n) (Pr n f g)" and "n = length xs" + shows "eval (Pr n f g) (0 # xs) = eval f xs" + using assms by (simp add: eval_def) + +lemma eval_Pr_diverg_Suc [simp]: + assumes "recfn (Suc n) (Pr n f g)" + and "n = length xs" + and "eval (Pr n f g) (x # xs) \" + shows "eval (Pr n f g) (Suc x # xs) \" + using assms by (simp add: eval_def) + +lemma eval_Pr_converg_Suc [simp]: + assumes "recfn (Suc n) (Pr n f g)" + and "n = length xs" + and "eval (Pr n f g) (x # xs) \" + shows "eval (Pr n f g) (Suc x # xs) = eval g (x # the (eval (Pr n f g) (x # xs)) # xs)" + using assms eval_def by auto + +lemma eval_Pr_diverg_add: + assumes "recfn (Suc n) (Pr n f g)" + and "n = length xs" + and "eval (Pr n f g) (x # xs) \" + shows "eval (Pr n f g) ((x + y) # xs) \" + using assms by (induction y) auto + +lemma eval_Pr_converg_le: + assumes "recfn (Suc n) (Pr n f g)" + and "n = length xs" + and "eval (Pr n f g) (x # xs) \" + and "y \ x" + shows "eval (Pr n f g) (y # xs) \" + using assms eval_Pr_diverg_add le_Suc_ex by metis + +lemma eval_Pr_Suc_converg: + assumes "recfn (Suc n) (Pr n f g)" + and "n = length xs" + and "eval (Pr n f g) (Suc x # xs) \" + shows "eval g (x # (the (eval (Pr n f g) (x # xs))) # xs) \" + and "eval (Pr n f g) (Suc x # xs) = eval g (x # the (eval (Pr n f g) (x # xs)) # xs)" + using eval_Pr_converg_Suc[of n f g xs x, OF assms(1,2)] + eval_Pr_converg_le[of n f g xs "Suc x" x, OF assms] assms(3) + by simp_all + +lemma eval_Mn [simp]: + assumes "recfn (length xs) (Mn n f)" + shows "eval (Mn n f) xs = + (if (\z. eval f (z # xs) \= 0 \ (\y)) + then Some (LEAST z. eval f (z # xs) \= 0 \ (\y)) + else None)" + using assms eval_def by auto + +text \For $\mu$-recursion, the condition @{term "(\y)"} +inside @{text LEAST} in the definition of @{term eval_wellf} is redundant.\ + +lemma eval_wellf_Mn: + "eval_wellf (Mn n f) xs = + (if (\z. eval_wellf f (z # xs) \= 0 \ (\y)) + then Some (LEAST z. eval_wellf f (z # xs) \= 0) + else None)" +proof - + let ?P = "\z. eval_wellf f (z # xs) \= 0 \ (\y)" + { + assume "\z. ?P z" + moreover define z where "z = Least ?P" + ultimately have "?P z" + using LeastI[of ?P] by blast + have "(LEAST z. eval_wellf f (z # xs) \= 0) = z" + proof (rule Least_equality) + show "eval_wellf f (z # xs) \= 0" + using `?P z` by simp + show "z \ y" if "eval_wellf f (y # xs) \= 0" for y + proof (rule ccontr) + assume "\ z \ y" + then have "y < z" by simp + moreover from this have "?P y" + using that `?P z` by simp + ultimately show False + using that not_less_Least z_def by blast + qed + qed + } + then show ?thesis by simp +qed + +lemma eval_Mn': + assumes "recfn (length xs) (Mn n f)" + shows "eval (Mn n f) xs = + (if (\z. eval f (z # xs) \= 0 \ (\y)) + then Some (LEAST z. eval f (z # xs) \= 0) + else None)" + using assms eval_def eval_wellf_Mn by auto + +text \Proving that $\mu$-recursion converges is easier if one does not +have to deal with @{text LEAST} directly.\ + +lemma eval_Mn_convergI: + assumes "recfn (length xs) (Mn n f)" + and "eval f (z # xs) \= 0" + and "\y. y < z \ eval f (y # xs) \\ 0" + shows "eval (Mn n f) xs \= z" +proof - + let ?P = "\z. eval f (z # xs) \= 0 \ (\y)" + have "z = Least ?P" + using Least_equality[of ?P z] assms(2,3) not_le_imp_less by blast + moreover have "?P z" using assms(2,3) by simp + ultimately show "eval (Mn n f) xs \= z" + using eval_Mn[OF assms(1)] by meson +qed + +text \Similarly, reasoning from a $\mu$-recursive function is +simplified somewhat by the next lemma.\ + +lemma eval_Mn_convergE: + assumes "recfn (length xs) (Mn n f)" and "eval (Mn n f) xs \= z" + shows "z = (LEAST z. eval f (z # xs) \= 0 \ (\y))" + and "eval f (z # xs) \= 0" + and "\y. y < z \ eval f (y # xs) \\ 0" +proof - + let ?P = "\z. eval f (z # xs) \= 0 \ (\y)" + show "z = Least ?P" + using assms eval_Mn[OF assms(1)] + by (metis (no_types, lifting) option.inject option.simps(3)) + moreover have "\z. ?P z" + using assms eval_Mn[OF assms(1)] by (metis option.distinct(1)) + ultimately have "?P z" + using LeastI[of ?P] by blast + then have "eval f (z # xs) \= 0 \ (\y)" + by simp + then show "eval f (z # xs) \= 0" by simp + show "\y. y < z \ eval f (y # xs) \\ 0" + using not_less_Least[of _ ?P] `z = Least ?P` `?P z` less_trans by blast +qed + +lemma eval_Mn_diverg: + assumes "recfn (length xs) (Mn n f)" + shows "\ (\z. eval f (z # xs) \= 0 \ (\y)) \ eval (Mn n f) xs \" + using assms eval_Mn[OF assms(1)] by simp + + +subsection \Extensional equality\ + +definition exteq :: "recf \ recf \ bool" (infix "\" 55) where + "f \ g \ arity f = arity g \ (\xs. length xs = arity f \ eval f xs = eval g xs)" + +lemma exteq_refl: "f \ f" + using exteq_def by simp + +lemma exteq_sym: "f \ g \ g \ f" + using exteq_def by simp + +lemma exteq_trans: "f \ g \ g \ h \ f \ h" + using exteq_def by simp + +lemma exteqI: + assumes "arity f = arity g" and "\xs. length xs = arity f \ eval f xs = eval g xs" + shows "f \ g" + using assms exteq_def by simp + +lemma exteqI1: + assumes "arity f = 1" and "arity g = 1" and "\x. eval f [x] = eval g [x]" + shows "f \ g" + using assms exteqI by (metis One_nat_def Suc_length_conv length_0_conv) + +text \For every partial recursive function @{term f} there are +infinitely many extensionally equal ones, for example, those that wrap @{term +f} arbitrarily often in the identity function.\ + +fun wrap_Id :: "recf \ nat \ recf" where + "wrap_Id f 0 = f" +| "wrap_Id f (Suc n) = Cn (arity f) (Id 1 0) [wrap_Id f n]" + +lemma recfn_wrap_Id: "recfn a f \ recfn a (wrap_Id f n)" + using wellf_arity_nonzero by (induction n) auto + +lemma exteq_wrap_Id: "recfn a f \ f \ wrap_Id f n" +proof (induction n) + case 0 + then show ?case by (simp add: exteq_refl) +next + case (Suc n) + have "wrap_Id f n \ wrap_Id f (Suc n) " + proof (rule exteqI) + show "arity (wrap_Id f n) = arity (wrap_Id f (Suc n))" + using Suc by (simp add: recfn_wrap_Id) + show "eval (wrap_Id f n) xs = eval (wrap_Id f (Suc n)) xs" + if "length xs = arity (wrap_Id f n)" for xs + proof - + have "recfn (length xs) (Cn (arity f) (Id 1 0) [wrap_Id f n])" + using that Suc recfn_wrap_Id by (metis wrap_Id.simps(2)) + then show "eval (wrap_Id f n) xs = eval (wrap_Id f (Suc n)) xs" + by auto + qed + qed + then show ?case using Suc exteq_trans by fast +qed + +fun depth :: "recf \ nat" where + "depth Z = 0" +| "depth S = 0" +| "depth (Id m n) = 0" +| "depth (Cn n f gs) = Suc (max (depth f) (Max (set (map depth gs))))" +| "depth (Pr n f g) = Suc (max (depth f) (depth g))" +| "depth (Mn n f) = Suc (depth f)" + +lemma depth_wrap_Id: "recfn a f \ depth (wrap_Id f n) = depth f + n" + by (induction n) simp_all + +lemma wrap_Id_injective: + assumes "recfn a f" and "wrap_Id f n\<^sub>1 = wrap_Id f n\<^sub>2" + shows "n\<^sub>1 = n\<^sub>2" + using assms by (metis add_left_cancel depth_wrap_Id) + +lemma exteq_infinite: + assumes "recfn a f" + shows "infinite {g. recfn a g \ g \ f}" (is "infinite ?R") +proof - + have "inj (wrap_Id f)" + using wrap_Id_injective `recfn a f` by (meson inj_onI) + then have "infinite (range (wrap_Id f))" + using finite_imageD by blast + moreover have "range (wrap_Id f) \ ?R" + using assms exteq_sym exteq_wrap_Id recfn_wrap_Id by blast + ultimately show ?thesis by (simp add: infinite_super) +qed + + +subsection \Primitive recursive and total functions\ + +fun Mn_free :: "recf \ bool" where + "Mn_free Z = True" +| "Mn_free S = True" +| "Mn_free (Id m n) = True" +| "Mn_free (Cn n f gs) = ((\g \ set gs. Mn_free g) \ Mn_free f)" +| "Mn_free (Pr n f g) = (Mn_free f \ Mn_free g)" +| "Mn_free (Mn n f) = False" + +text \This is our notion of $n$-ary primitive recursive function:\ + +abbreviation prim_recfn :: "nat \ recf \ bool" where + "prim_recfn n f \ recfn n f \ Mn_free f" + +definition total :: "recf \ bool" where + "total f \ \xs. length xs = arity f \ eval f xs \" + +lemma totalI [intro]: + assumes "\xs. length xs = arity f \ eval f xs \" + shows "total f" + using assms total_def by simp + +lemma totalE [simp]: + assumes "total f" and "recfn n f" and "length xs = n" + shows "eval f xs \" + using assms total_def by simp + +lemma totalI1 : + assumes "recfn 1 f" and "\x. eval f [x] \" + shows "total f" + using assms totalI[of f] by (metis One_nat_def length_0_conv length_Suc_conv) + +lemma totalI2: + assumes "recfn 2 f" and "\x y. eval f [x, y] \" + shows "total f" + using assms totalI[of f] by (smt length_0_conv length_Suc_conv numeral_2_eq_2) + +lemma totalI3: + assumes "recfn 3 f" and "\x y z. eval f [x, y, z] \" + shows "total f" + using assms totalI[of f] by (smt length_0_conv length_Suc_conv numeral_3_eq_3) + +lemma totalI4: + assumes "recfn 4 f" and "\w x y z. eval f [w, x, y, z] \" + shows "total f" +proof (rule totalI[of f]) + fix xs :: "nat list" + assume "length xs = arity f" + then have "length xs = Suc (Suc (Suc (Suc 0)))" + using assms(1) by simp + then obtain w x y z where "xs = [w, x, y, z]" + by (smt Suc_length_conv length_0_conv) + then show "eval f xs \" using assms(2) by simp +qed + +lemma Mn_free_imp_total [intro]: + assumes "wellf f" and "Mn_free f" + shows "total f" + using assms +proof (induction f rule: Mn_free.induct) + case (5 n f g) + have "eval (Pr n f g) (x # xs) \" if "length (x # xs) = arity (Pr n f g)" for x xs + using 5 that by (induction x) auto + then show ?case by (metis arity.simps(5) length_Suc_conv totalI) +qed (auto simp add: total_def eval_def) + +lemma prim_recfn_total: "prim_recfn n f \ total f" + using Mn_free_imp_total by simp + +lemma eval_Pr_prim_Suc: + assumes "h = Pr n f g" and "prim_recfn (Suc n) h" and "length xs = n" + shows "eval h (Suc x # xs) = eval g (x # the (eval h (x # xs)) # xs)" + using assms eval_Pr_converg_Suc prim_recfn_total by simp + +lemma Cn_total: + assumes "\g\set gs. total g" and "total f" and "recfn n (Cn n f gs)" + shows "total (Cn n f gs)" + using assms by (simp add: totalI) + +lemma Pr_total: + assumes "total f" and "total g" and "recfn (Suc n) (Pr n f g)" + shows "total (Pr n f g)" +proof - + have "eval (Pr n f g) (x # xs) \" if "length xs = n" for x xs + using that assms by (induction x) auto + then show ?thesis + using assms(3) totalI by (metis Suc_length_conv arity.simps(5)) +qed + +lemma eval_Mn_total: + assumes "recfn (length xs) (Mn n f)" and "total f" + shows "eval (Mn n f) xs = + (if (\z. eval f (z # xs) \= 0) + then Some (LEAST z. eval f (z # xs) \= 0) + else None)" + using assms by auto + + +section \Simple functions\ + +text \This section, too, bears some similarity to Urban's formalization +in Xu et~al.~\cite{Universal_Turing_Machine-AFP}, but is more minimalistic in +scope. + +As a general naming rule, instances of @{typ recf} and functions +returning such instances get names starting with @{text r_}. Typically, for +an @{text r_xyz} there will be a lemma @{text r_xyz_recfn} or @{text +r_xyz_prim} establishing its (primitive) recursiveness, arity, and +well-formedness. Moreover there will be a lemma @{text r_xyz} describing its +semantics, for which we will sometimes introduce an Isabelle function @{text +xyz}.\ + + +subsection \Manipulating parameters\ + +text \Appending dummy parameters:\ + +definition r_dummy :: "nat \ recf \ recf" where + "r_dummy n f \ Cn (arity f + n) f (map (\i. Id (arity f + n) i) [0.. prim_recfn (a + n) (r_dummy n f)" + using wellf_arity_nonzero by (auto simp add: r_dummy_def) + +lemma r_dummy_recfn [simp]: + "recfn a f \ recfn (a + n) (r_dummy n f)" + using wellf_arity_nonzero by (auto simp add: r_dummy_def) + +lemma r_dummy [simp]: + "r_dummy n f = Cn (arity f + n) f (map (\i. Id (arity f + n) i) [0..= xs ! i" if "i < arity f" for i + using that assms by (simp add: nth_append) + ultimately have "map (\g. the (eval_wellf g (xs @ ys))) ?gs = xs" + by (metis (no_types, lifting) assms(1) length_map nth_equalityI nth_map option.sel) + moreover have "\g \ set ?gs. eval_wellf g (xs @ ys) \" + using * by simp + moreover have "recfn (length (xs @ ys)) ?r" + using assms r_dummy_recfn by fastforce + ultimately show ?thesis + by (auto simp add: assms eval_def) +qed + +text \Shrinking a binary function to a unary one is useful when we want +to define a unary function via the @{term Pr} operation, which can only +construct @{typ recf}s of arity two or higher.\ + +definition r_shrink :: "recf \ recf" where + "r_shrink f \ Cn 1 f [Id 1 0, Id 1 0]" + +lemma r_shrink_prim [simp]: "prim_recfn 2 f \ prim_recfn 1 (r_shrink f)" + by (simp add: r_shrink_def) + +lemma r_shrink_recfn [simp]: "recfn 2 f \ recfn 1 (r_shrink f)" + by (simp add: r_shrink_def) + +lemma r_shrink [simp]: "recfn 2 f \ eval (r_shrink f) [x] = eval f [x, x]" + by (simp add: r_shrink_def) + +definition r_swap :: "recf \ recf" where + "r_swap f \ Cn 2 f [Id 2 1, Id 2 0]" + +lemma r_swap_recfn [simp]: "recfn 2 f \ recfn 2 (r_swap f)" + by (simp add: r_swap_def) + +lemma r_swap_prim [simp]: "prim_recfn 2 f \ prim_recfn 2 (r_swap f)" + by (simp add: r_swap_def) + +lemma r_swap [simp]: "recfn 2 f \ eval (r_swap f) [x, y] = eval f [y, x]" + by (simp add: r_swap_def) + +text \Prepending one dummy parameter:\ + +definition r_shift :: "recf \ recf" where + "r_shift f \ Cn (Suc (arity f)) f (map (\i. Id (Suc (arity f)) (Suc i)) [0.. prim_recfn (Suc a) (r_shift f)" + by (simp add: r_shift_def) + +lemma r_shift_recfn [simp]: "recfn a f \ recfn (Suc a) (r_shift f)" + by (simp add: r_shift_def) + +lemma r_shift [simp]: + assumes "recfn (length xs) f" + shows "eval (r_shift f) (x # xs) = eval f xs" +proof - + let ?r = "r_shift f" + let ?gs = "map (\i. Id (Suc (arity f)) (Suc i)) [0..= xs ! i" if "i < arity f" for i + using assms nth_append that by simp + ultimately have "map (\g. the (eval g (x # xs))) ?gs = xs" + by (metis (no_types, lifting) assms length_map nth_equalityI nth_map option.sel) + moreover have "\g \ set ?gs. eval g (x # xs) \ None" + using * by simp + ultimately show ?thesis using r_shift_def assms by simp +qed + + +subsection \Arithmetic and logic\ + +text \The unary constants:\ + +fun r_const :: "nat \ recf" where + "r_const 0 = Z" +| "r_const (Suc c) = Cn 1 S [r_const c]" + +lemma r_const_prim [simp]: "prim_recfn 1 (r_const c)" + by (induction c) (simp_all) + +lemma r_const [simp]: "eval (r_const c) [x] \= c" + by (induction c) simp_all + +text \Constants of higher arities:\ + +definition "r_constn n c \ if n = 0 then r_const c else r_dummy n (r_const c)" + +lemma r_constn_prim [simp]: "prim_recfn (Suc n) (r_constn n c)" + unfolding r_constn_def by simp + +lemma r_constn [simp]: "length xs = Suc n \ eval (r_constn n c) xs \= c" + unfolding r_constn_def + by simp (metis length_0_conv length_Suc_conv r_const) + +text \We introduce addition, subtraction, and multiplication, but +interestingly enough we can make do without division.\ + +definition "r_add \ Pr 1 (Id 1 0) (Cn 3 S [Id 3 1])" + +lemma r_add_prim [simp]: "prim_recfn 2 r_add" + by (simp add: r_add_def) + +lemma r_add [simp]: "eval r_add [a, b] \= a + b" + unfolding r_add_def by (induction a) simp_all + +definition "r_mul \ Pr 1 Z (Cn 3 r_add [Id 3 1, Id 3 2])" + +lemma r_mul_prim [simp]: "prim_recfn 2 r_mul" + unfolding r_mul_def by simp + +lemma r_mul [simp]: "eval r_mul [a, b] \= a * b" + unfolding r_mul_def by (induction a) simp_all + +definition "r_dec \ Cn 1 (Pr 1 Z (Id 3 0)) [Id 1 0, Id 1 0]" + +lemma r_dec_prim [simp]: "prim_recfn 1 r_dec" + by (simp add: r_dec_def) + +lemma r_dec [simp]: "eval r_dec [a] \= a - 1" +proof - + have "eval (Pr 1 Z (Id 3 0)) [x, y] \= x - 1" for x y + by (induction x) simp_all + then show ?thesis by (simp add: r_dec_def) +qed + +definition "r_sub \ r_swap (Pr 1 (Id 1 0) (Cn 3 r_dec [Id 3 1]))" + +lemma r_sub_prim [simp]: "prim_recfn 2 r_sub" + unfolding r_sub_def by simp + +lemma r_sub [simp]: "eval r_sub [a, b] \= a - b" +proof - + have "eval (Pr 1 (Id 1 0) (Cn 3 r_dec [Id 3 1])) [x, y] \= y - x" for x y + by (induction x) simp_all + then show ?thesis unfolding r_sub_def by simp +qed + +definition "r_sign \ r_shrink (Pr 1 Z (r_constn 2 1))" + +lemma r_sign_prim [simp]: "prim_recfn 1 r_sign" + unfolding r_sign_def by simp + +lemma r_sign [simp]: "eval r_sign [x] \= (if x = 0 then 0 else 1)" +proof - + have "eval (Pr 1 Z (r_constn 2 1)) [x, y] \= (if x = 0 then 0 else 1)" for x y + by (induction x) simp_all + then show ?thesis unfolding r_sign_def by simp +qed + +text \In the logical functions, true will be represented by zero, and +false will be represented by non-zero as argument and by one as +result.\ + +definition "r_not \ Cn 1 r_sub [r_const 1, r_sign]" + +lemma r_not_prim [simp]: "prim_recfn 1 r_not" + unfolding r_not_def by simp + +lemma r_not [simp]: "eval r_not [x] \= (if x = 0 then 1 else 0)" + unfolding r_not_def by simp + +definition "r_nand \ Cn 2 r_not [r_add]" + +lemma r_nand_prim [simp]: "prim_recfn 2 r_nand" + unfolding r_nand_def by simp + +lemma r_nand [simp]: "eval r_nand [x, y] \= (if x = 0 \ y = 0 then 1 else 0)" + unfolding r_nand_def by simp + +definition "r_and \ Cn 2 r_not [r_nand]" + +lemma r_and_prim [simp]: "prim_recfn 2 r_and" + unfolding r_and_def by simp + +lemma r_and [simp]: "eval r_and [x, y] \= (if x = 0 \ y = 0 then 0 else 1)" + unfolding r_and_def by simp + +definition "r_or \ Cn 2 r_sign [r_mul]" + +lemma r_or_prim [simp]: "prim_recfn 2 r_or" + unfolding r_or_def by simp + +lemma r_or [simp]: "eval r_or [x, y] \= (if x = 0 \ y = 0 then 0 else 1)" + unfolding r_or_def by simp + + +subsection \Comparison and conditions\ + +definition "r_ifz \ + let ifzero = (Cn 3 r_mul [r_dummy 2 r_not, Id 3 1]); + ifnzero = (Cn 3 r_mul [r_dummy 2 r_sign, Id 3 2]) + in Cn 3 r_add [ifzero, ifnzero]" + +lemma r_ifz_prim [simp]: "prim_recfn 3 r_ifz" + unfolding r_ifz_def by simp + +lemma r_ifz [simp]: "eval r_ifz [cond, val0, val1] \= (if cond = 0 then val0 else val1)" + unfolding r_ifz_def by (simp add: Let_def) + +definition "r_eq \ Cn 2 r_sign [Cn 2 r_add [r_sub, r_swap r_sub]]" + +lemma r_eq_prim [simp]: "prim_recfn 2 r_eq" + unfolding r_eq_def by simp + +lemma r_eq [simp]: "eval r_eq [x, y] \= (if x = y then 0 else 1)" + unfolding r_eq_def by simp + +definition "r_ifeq \ Cn 4 r_ifz [r_dummy 2 r_eq, Id 4 2, Id 4 3]" + +lemma r_ifeq_prim [simp]: "prim_recfn 4 r_ifeq" + unfolding r_ifeq_def by simp + +lemma r_ifeq [simp]: "eval r_ifeq [a, b, v\<^sub>0, v\<^sub>1] \= (if a = b then v\<^sub>0 else v\<^sub>1)" + unfolding r_ifeq_def using r_dummy_append[of r_eq "[a, b]" "[v\<^sub>0, v\<^sub>1]" 2] + by simp + +definition "r_neq \ Cn 2 r_not [r_eq]" + +lemma r_neq_prim [simp]: "prim_recfn 2 r_neq" + unfolding r_neq_def by simp + +lemma r_neq [simp]: "eval r_neq [x, y] \= (if x = y then 1 else 0)" + unfolding r_neq_def by simp + +definition "r_ifle \ Cn 4 r_ifz [r_dummy 2 r_sub, Id 4 2, Id 4 3]" + +lemma r_ifle_prim [simp]: "prim_recfn 4 r_ifle" + unfolding r_ifle_def by simp + +lemma r_ifle [simp]: "eval r_ifle [a, b, v\<^sub>0, v\<^sub>1] \= (if a \ b then v\<^sub>0 else v\<^sub>1)" + unfolding r_ifle_def using r_dummy_append[of r_sub "[a, b]" "[v\<^sub>0, v\<^sub>1]" 2] + by simp + +definition "r_ifless \ Cn 4 r_ifle [Id 4 1, Id 4 0, Id 4 3, Id 4 2]" + +lemma r_ifless_prim [simp]: "prim_recfn 4 r_ifless" + unfolding r_ifless_def by simp + +lemma r_ifless [simp]: "eval r_ifless [a, b, v\<^sub>0, v\<^sub>1] \= (if a < b then v\<^sub>0 else v\<^sub>1)" + unfolding r_ifless_def by simp + +definition "r_less \ Cn 2 r_ifle [Id 2 1, Id 2 0, r_constn 1 1, r_constn 1 0]" + +lemma r_less_prim [simp]: "prim_recfn 2 r_less" + unfolding r_less_def by simp + +lemma r_less [simp]: "eval r_less [x, y] \= (if x < y then 0 else 1)" + unfolding r_less_def by simp + +definition "r_le \ Cn 2 r_ifle [Id 2 0, Id 2 1, r_constn 1 0, r_constn 1 1]" + +lemma r_le_prim [simp]: "prim_recfn 2 r_le" + unfolding r_le_def by simp + +lemma r_le [simp]: "eval r_le [x, y] \= (if x \ y then 0 else 1)" + unfolding r_le_def by simp + +text \Arguments are evaluated eagerly. Therefore @{term "r_ifz"}, etc. +cannot be combined with a diverging function to implement a conditionally +diverging function in the naive way. The following function implements a +special case needed in the next section. A \hyperlink{p:r_lifz}{general lazy +version} of @{term "r_ifz"} will be introduced later with the help of a +universal function.\ + +definition "r_ifeq_else_diverg \ + Cn 3 r_add [Id 3 2, Mn 3 (Cn 4 r_add [Id 4 0, Cn 4 r_eq [Id 4 1, Id 4 2]])]" + +lemma r_ifeq_else_diverg_recfn [simp]: "recfn 3 r_ifeq_else_diverg" + unfolding r_ifeq_else_diverg_def by simp + +lemma r_ifeq_else_diverg [simp]: + "eval r_ifeq_else_diverg [a, b, v] = (if a = b then Some v else None)" + unfolding r_ifeq_else_diverg_def by simp + + +section \The halting problem\label{s:halting}\ + +text \Decidability will be treated more thoroughly in +Section~\ref{s:decidable}. But the halting problem is prominent enough to +deserve an early mention.\ + +definition decidable :: "nat set \ bool" where + "decidable X \ \f. recfn 1 f \ (\x. eval f [x] \= (if x \ X then 1 else 0))" + +text \No matter how partial recursive functions are encoded as natural +numbers, the set of all codes of functions halting on their own code is +undecidable.\ + +theorem halting_problem_undecidable: + fixes code :: "nat \ recf" + assumes "\f. recfn 1 f \ \i. code i = f" + shows "\ decidable {x. eval (code x) [x] \}" (is "\ decidable ?K") +proof + assume "decidable ?K" + then obtain f where "recfn 1 f" and f: "\x. eval f [x] \= (if x \ ?K then 1 else 0)" + using decidable_def by auto + define g where "g \ Cn 1 r_ifeq_else_diverg [f, Z, Z]" + then have "recfn 1 g" + using `recfn 1 f` r_ifeq_else_diverg_recfn by simp + with assms obtain i where i: "code i = g" by auto + from g_def have "eval g [x] = (if x \ ?K then Some 0 else None)" for x + using r_ifeq_else_diverg_recfn `recfn 1 f` f by simp + then have "eval g [i] \ \ i \ ?K" by simp + also have "... \ eval (code i) [i] \" by simp + also have "... \ eval g [i] \" + using i by simp + finally have "eval g [i] \ \ eval g [i] \" . + then show False by auto +qed + + +section \Encoding tuples and lists\ + +text \This section is based on the Cantor encoding for pairs. Tuples +are encoded by repeated application of the pairing function, lists by pairing +their length with the code for a tuple. Thus tuples have a fixed length that +must be known when decoding, whereas lists are dynamically sized and know +their current length.\ + + +subsection \Pairs and tuples\ + + +subsubsection \The Cantor pairing function\ + +definition "r_triangle \ r_shrink (Pr 1 Z (r_dummy 1 (Cn 2 S [r_add])))" + +lemma r_triangle_prim: "prim_recfn 1 r_triangle" + unfolding r_triangle_def by simp + +lemma r_triangle: "eval r_triangle [n] \= Sum {0..n}" +proof - + let ?r = "r_dummy 1 (Cn 2 S [r_add])" + have "eval ?r [x, y, z] \= Suc (x + y)" for x y z + using r_dummy_append[of "Cn 2 S [r_add]" "[x, y]" "[z]" 1] by simp + then have "eval (Pr 1 Z ?r) [x, y] \= Sum {0..x}" for x y + by (induction x) simp_all + then show ?thesis unfolding r_triangle_def by simp +qed + +lemma r_triangle_eq_triangle [simp]: "eval r_triangle [n] \= triangle n" + using r_triangle gauss_sum_nat triangle_def by simp + +definition "r_prod_encode \ Cn 2 r_add [Cn 2 r_triangle [r_add], Id 2 0]" + +lemma r_prod_encode_prim [simp]: "prim_recfn 2 r_prod_encode" + unfolding r_prod_encode_def using r_triangle_prim by simp + +lemma r_prod_encode [simp]: "eval r_prod_encode [m, n] \= prod_encode (m, n)" + unfolding r_prod_encode_def prod_encode_def using r_triangle_prim by simp + +text \These abbreviations are just two more things borrowed from +Xu~et~al.~\cite{Universal_Turing_Machine-AFP}.\ + +abbreviation "pdec1 z \ fst (prod_decode z)" + +abbreviation "pdec2 z \ snd (prod_decode z)" + +lemma pdec1_le: "pdec1 i \ i" + by (metis le_prod_encode_1 prod.collapse prod_decode_inverse) + +lemma pdec2_le: "pdec2 i \ i" + by (metis le_prod_encode_2 prod.collapse prod_decode_inverse) + +lemma pdec_less: "pdec2 i < Suc i" + using pdec2_le by (simp add: le_imp_less_Suc) + +lemma pdec1_zero: "pdec1 0 = 0" + using pdec1_le by auto + +definition "r_maxletr \ + Pr 1 Z (Cn 3 r_ifle [r_dummy 2 (Cn 1 r_triangle [S]), Id 3 2, Cn 3 S [Id 3 0], Id 3 1])" + +lemma r_maxletr_prim: "prim_recfn 2 r_maxletr" + unfolding r_maxletr_def using r_triangle_prim by simp + +lemma not_Suc_Greatest_not_Suc: + assumes "\ P (Suc x)" and "\x. P x" + shows "(GREATEST y. y \ x \ P y) = (GREATEST y. y \ Suc x \ P y)" + using assms by (metis le_SucI le_Suc_eq) + +lemma r_maxletr: "eval r_maxletr [x\<^sub>0, x\<^sub>1] \= (GREATEST y. y \ x\<^sub>0 \ triangle y \ x\<^sub>1)" +proof - + let ?g = "Cn 3 r_ifle [r_dummy 2 (Cn 1 r_triangle [S]), Id 3 2, Cn 3 S [Id 3 0], Id 3 1]" + have greatest: + "(if triangle (Suc x\<^sub>0) \ x\<^sub>1 then Suc x\<^sub>0 else (GREATEST y. y \ x\<^sub>0 \ triangle y \ x\<^sub>1)) = + (GREATEST y. y \ Suc x\<^sub>0 \ triangle y \ x\<^sub>1)" + for x\<^sub>0 x\<^sub>1 + proof (cases "triangle (Suc x\<^sub>0) \ x\<^sub>1") + case True + then show ?thesis + using Greatest_equality[of "\y. y \ Suc x\<^sub>0 \ triangle y \ x\<^sub>1"] by fastforce + next + case False + then show ?thesis + using not_Suc_Greatest_not_Suc[of "\y. triangle y \ x\<^sub>1" x\<^sub>0] by fastforce + qed + show ?thesis + unfolding r_maxletr_def using r_triangle_prim + proof (induction x\<^sub>0) + case 0 + then show ?case + using Greatest_equality[of "\y. y \ 0 \ triangle y \ x\<^sub>1" 0] by simp + next + case (Suc x\<^sub>0) + then show ?case using greatest by simp + qed +qed + +definition "r_maxlt \ r_shrink r_maxletr" + +lemma r_maxlt_prim: "prim_recfn 1 r_maxlt" + unfolding r_maxlt_def using r_maxletr_prim by simp + +lemma r_maxlt: "eval r_maxlt [e] \= (GREATEST y. triangle y \ e)" +proof - + have "y \ triangle y" for y + by (induction y) auto + then have "triangle y \ e \ y \ e" for y e + using order_trans by blast + then have "(GREATEST y. y \ e \ triangle y \ e) = (GREATEST y. triangle y \ e)" + by metis + moreover have "eval r_maxlt [e] \= (GREATEST y. y \ e \ triangle y \ e)" + using r_maxletr r_shrink r_maxlt_def r_maxletr_prim by fastforce + ultimately show ?thesis by simp +qed + +definition "pdec1' e \ e - triangle (GREATEST y. triangle y \ e)" + +definition "pdec2' e \ (GREATEST y. triangle y \ e) - pdec1' e" + +lemma max_triangle_bound: "triangle z \ e \ z \ e" + by (metis Suc_pred add_leD2 less_Suc_eq triangle_Suc zero_le zero_less_Suc) + +lemma triangle_greatest_le: "triangle (GREATEST y. triangle y \ e) \ e" + using max_triangle_bound GreatestI_nat[of "\y. triangle y \ e" 0 e] by simp + +lemma prod_encode_pdec': "prod_encode (pdec1' e, pdec2' e) = e" +proof - + let ?P = "\y. triangle y \ e" + let ?y = "GREATEST y. ?P y" + have "pdec1' e \ ?y" + proof (rule ccontr) + assume "\ pdec1' e \ ?y" + then have "e - triangle ?y > ?y" + using pdec1'_def by simp + then have "?P (Suc ?y)" by simp + moreover have "\z. ?P z \ z \ e" + using max_triangle_bound by simp + ultimately have "Suc ?y \ ?y" + using Greatest_le_nat[of ?P "Suc ?y" e] by blast + then show False by simp + qed + then have "pdec1' e + pdec2' e = ?y" + using pdec1'_def pdec2'_def by simp + then have "prod_encode (pdec1' e, pdec2' e) = triangle ?y + pdec1' e" + by (simp add: prod_encode_def) + then show ?thesis using pdec1'_def triangle_greatest_le by simp +qed + +lemma pdec': + "pdec1' e = pdec1 e" + "pdec2' e = pdec2 e" + using prod_encode_pdec' prod_encode_inverse by (metis fst_conv, metis snd_conv) + +definition "r_pdec1 \ Cn 1 r_sub [Id 1 0, Cn 1 r_triangle [r_maxlt]]" + +lemma r_pdec1_prim [simp]: "prim_recfn 1 r_pdec1" + unfolding r_pdec1_def using r_triangle_prim r_maxlt_prim by simp + +lemma r_pdec1 [simp]: "eval r_pdec1 [e] \= pdec1 e" + unfolding r_pdec1_def using r_triangle_prim r_maxlt_prim pdec' pdec1'_def + by (simp add: r_maxlt) + +definition "r_pdec2 \ Cn 1 r_sub [r_maxlt, r_pdec1]" + +lemma r_pdec2_prim [simp]: "prim_recfn 1 r_pdec2" + unfolding r_pdec2_def using r_maxlt_prim by simp + +lemma r_pdec2 [simp]: "eval r_pdec2 [e] \= pdec2 e" + unfolding r_pdec2_def using r_maxlt_prim r_maxlt pdec' pdec2'_def by simp + +abbreviation "pdec12 i \ pdec1 (pdec2 i)" +abbreviation "pdec22 i \ pdec2 (pdec2 i)" +abbreviation "pdec122 i \ pdec1 (pdec22 i)" +abbreviation "pdec222 i \ pdec2 (pdec22 i)" + +definition "r_pdec12 \ Cn 1 r_pdec1 [r_pdec2]" + +lemma r_pdec12_prim [simp]: "prim_recfn 1 r_pdec12" + unfolding r_pdec12_def by simp + +lemma r_pdec12 [simp]: "eval r_pdec12 [e] \= pdec12 e" + unfolding r_pdec12_def by simp + +definition "r_pdec22 \ Cn 1 r_pdec2 [r_pdec2]" + +lemma r_pdec22_prim [simp]: "prim_recfn 1 r_pdec22" + unfolding r_pdec22_def by simp + +lemma r_pdec22 [simp]: "eval r_pdec22 [e] \= pdec22 e" + unfolding r_pdec22_def by simp + +definition "r_pdec122 \ Cn 1 r_pdec1 [r_pdec22]" + +lemma r_pdec122_prim [simp]: "prim_recfn 1 r_pdec122" + unfolding r_pdec122_def by simp + +lemma r_pdec122 [simp]: "eval r_pdec122 [e] \= pdec122 e" + unfolding r_pdec122_def by simp + +definition "r_pdec222 \ Cn 1 r_pdec2 [r_pdec22]" + +lemma r_pdec222_prim [simp]: "prim_recfn 1 r_pdec222" + unfolding r_pdec222_def by simp + +lemma r_pdec222 [simp]: "eval r_pdec222 [e] \= pdec222 e" + unfolding r_pdec222_def by simp + + +subsubsection \The Cantor tuple function\ + +text \The empty tuple gets no code, whereas singletons are encoded by their +only element and other tuples by recursively applying the pairing function. +This yields, for every $n$, the function @{term "tuple_encode n"}, which is a +bijection between the natural numbers and the lists of length $(n + 1)$.\ + +fun tuple_encode :: "nat \ nat list \ nat" where + "tuple_encode n [] = undefined" +| "tuple_encode 0 (x # xs) = x" +| "tuple_encode (Suc n) (x # xs) = prod_encode (x, tuple_encode n xs)" + +lemma tuple_encode_prod_encode: "tuple_encode 1 [x, y] = prod_encode (x, y)" + by simp + +fun tuple_decode where + "tuple_decode 0 i = [i]" +| "tuple_decode (Suc n) i = pdec1 i # tuple_decode n (pdec2 i)" + +lemma tuple_encode_decode [simp]: + "tuple_encode (length xs - 1) (tuple_decode (length xs - 1) i) = i" +proof (induction "length xs - 1" arbitrary: xs i) + case 0 + then show ?case by simp +next + case (Suc n) + then have "length xs - 1 > 0" by simp + with Suc have *: "tuple_encode n (tuple_decode n j) = j" for j + by (metis diff_Suc_1 length_tl) + from Suc have "tuple_decode (Suc n) i = pdec1 i # tuple_decode n (pdec2 i)" + using tuple_decode.simps(2) by blast + then have "tuple_encode (Suc n) (tuple_decode (Suc n) i) = + tuple_encode (Suc n) (pdec1 i # tuple_decode n (pdec2 i))" + using Suc by simp + also have "... = prod_encode (pdec1 i, tuple_encode n (tuple_decode n (pdec2 i)))" + by simp + also have "... = prod_encode (pdec1 i, pdec2 i)" + using Suc * by simp + also have "... = i" by simp + finally have "tuple_encode (Suc n) (tuple_decode (Suc n) i) = i" . + then show ?case by (simp add: Suc.hyps(2)) +qed + +lemma tuple_encode_decode' [simp]: "tuple_encode n (tuple_decode n i) = i" + using tuple_encode_decode by (metis Ex_list_of_length diff_Suc_1 length_Cons) + +lemma tuple_decode_encode: + assumes "length xs > 0" + shows "tuple_decode (length xs - 1) (tuple_encode (length xs - 1) xs) = xs" + using assms +proof (induction "length xs - 1" arbitrary: xs) + case 0 + moreover from this have "length xs = 1" by linarith + ultimately show ?case + by (metis One_nat_def length_0_conv length_Suc_conv tuple_decode.simps(1) + tuple_encode.simps(2)) +next + case (Suc n) + let ?t = "tl xs" + let ?i = "tuple_encode (Suc n) xs" + have "length ?t > 0" and "length ?t - 1 = n" + using Suc by simp_all + then have "tuple_decode n (tuple_encode n ?t) = ?t" + using Suc by blast + moreover have "?i = prod_encode (hd xs, tuple_encode n ?t)" + using Suc by (metis hd_Cons_tl length_greater_0_conv tuple_encode.simps(3)) + moreover have "tuple_decode (Suc n) ?i = pdec1 ?i # tuple_decode n (pdec2 ?i)" + using tuple_decode.simps(2) by blast + ultimately have "tuple_decode (Suc n) ?i = xs" + using Suc.prems by simp + then show ?case by (simp add: Suc.hyps(2)) +qed + +lemma tuple_decode_encode' [simp]: + assumes "length xs = Suc n" + shows "tuple_decode n (tuple_encode n xs) = xs" + using assms tuple_decode_encode by (metis diff_Suc_1 zero_less_Suc) + +lemma tuple_decode_length [simp]: "length (tuple_decode n i) = Suc n" + by (induction n arbitrary: i) simp_all + +lemma tuple_decode_nonzero: + assumes "n > 0" + shows "tuple_decode n i = pdec1 i # tuple_decode (n - 1) (pdec2 i)" + using assms by (metis One_nat_def Suc_pred tuple_decode.simps(2)) + +text \The tuple encoding functions are primitive recursive.\ + +fun r_tuple_encode :: "nat \ recf" where + "r_tuple_encode 0 = Id 1 0" +| "r_tuple_encode (Suc n) = + Cn (Suc (Suc n)) r_prod_encode [Id (Suc (Suc n)) 0, r_shift (r_tuple_encode n)]" + +lemma r_tuple_encode_prim [simp]: "prim_recfn (Suc n) (r_tuple_encode n)" + by (induction n) simp_all + +lemma r_tuple_encode: + assumes "length xs = Suc n" + shows "eval (r_tuple_encode n) xs \= tuple_encode n xs" + using assms +proof (induction n arbitrary: xs) + case 0 + then show ?case + by (metis One_nat_def eval_Id length_Suc_conv nth_Cons_0 + r_tuple_encode.simps(1) tuple_encode.simps(2) zero_less_one) +next + case (Suc n) + then obtain y ys where y_ys: "y # ys = xs" + by (metis length_Suc_conv) + with Suc have "eval (r_tuple_encode n) ys \= tuple_encode n ys" + by auto + with y_ys have "eval (r_shift (r_tuple_encode n)) xs \= tuple_encode n ys" + using Suc.prems r_shift_prim r_tuple_encode_prim by auto + moreover have "eval (Id (Suc (Suc n)) 0) xs \= y" + using y_ys Suc.prems by auto + ultimately have "eval (r_tuple_encode (Suc n)) xs \= prod_encode (y, tuple_encode n ys)" + using Suc.prems by simp + then show ?case using y_ys by auto +qed + + +subsubsection \Functions on encoded tuples\ + +text \The function for accessing the $n$-th element of a tuple returns +$0$ for out-of-bounds access.\ + +definition e_tuple_nth :: "nat \ nat \ nat \ nat" where + "e_tuple_nth a i n \ if n \ a then (tuple_decode a i) ! n else 0" + +lemma e_tuple_nth_le [simp]: "n \ a \ e_tuple_nth a i n = (tuple_decode a i) ! n" + using e_tuple_nth_def by simp + +lemma e_tuple_nth_gr [simp]: "n > a \ e_tuple_nth a i n = 0" + using e_tuple_nth_def by simp + +lemma tuple_decode_pdec2: "tuple_decode a (pdec2 es) = tl (tuple_decode (Suc a) es)" + by simp + +fun iterate :: "nat \ ('a \ 'a) \ ('a \ 'a)" where + "iterate 0 f = id" +| "iterate (Suc n) f = f \ (iterate n f)" + +lemma iterate_additive: + assumes "iterate t\<^sub>1 f x = y" and "iterate t\<^sub>2 f y = z" + shows "iterate (t\<^sub>1 + t\<^sub>2) f x = z" + using assms by (induction t\<^sub>2 arbitrary: z) auto + +lemma iterate_additive': "iterate (t\<^sub>1 + t\<^sub>2) f x = iterate t\<^sub>2 f (iterate t\<^sub>1 f x)" + using iterate_additive by metis + +lemma e_tuple_nth_elementary: + assumes "k \ a" + shows "e_tuple_nth a i k = (if a = k then (iterate k pdec2 i) else (pdec1 (iterate k pdec2 i)))" +proof - + have *: "tuple_decode (a - k) (iterate k pdec2 i) = drop k (tuple_decode a i)" + using assms + by (induction k) (simp, simp add: Suc_diff_Suc tuple_decode_pdec2 drop_Suc tl_drop) + show ?thesis + proof (cases "a = k") + case True + then have "tuple_decode 0 (iterate k pdec2 i) = drop k (tuple_decode a i)" + using assms * by simp + moreover from this have "drop k (tuple_decode a i) = [tuple_decode a i ! k]" + using assms True by (metis nth_via_drop tuple_decode.simps(1)) + ultimately show ?thesis using True by simp + next + case False + with assms have "a - k > 0" by simp + with * have "tuple_decode (a - k) (iterate k pdec2 i) = drop k (tuple_decode a i)" + by simp + then have "pdec1 (iterate k pdec2 i) = hd (drop k (tuple_decode a i))" + using tuple_decode_nonzero `a - k > 0` by (metis list.sel(1)) + with `a - k > 0` have "pdec1 (iterate k pdec2 i) = (tuple_decode a i) ! k" + by (simp add: hd_drop_conv_nth) + with False assms show ?thesis by simp + qed +qed + +definition "r_nth_inbounds \ + let r = Pr 1 (Id 1 0) (Cn 3 r_pdec2 [Id 3 1]) + in Cn 3 r_ifeq + [Id 3 0, + Id 3 2, + Cn 3 r [Id 3 2, Id 3 1], + Cn 3 r_pdec1 [Cn 3 r [Id 3 2, Id 3 1]]]" + +lemma r_nth_inbounds_prim: "prim_recfn 3 r_nth_inbounds" + unfolding r_nth_inbounds_def by (simp add: Let_def) + +lemma r_nth_inbounds: + "k \ a \ eval r_nth_inbounds [a, i, k] \= e_tuple_nth a i k" + "eval r_nth_inbounds [a, i, k] \" +proof - + let ?r = "Pr 1 (Id 1 0) (Cn 3 r_pdec2 [Id 3 1])" + let ?h = "Cn 3 ?r [Id 3 2, Id 3 1]" + have "eval ?r [k, i] \= iterate k pdec2 i" for k i + using r_pdec2_prim by (induction k) (simp_all) + then have "eval ?h [a, i, k] \= iterate k pdec2 i" + using r_pdec2_prim by simp + then have "eval r_nth_inbounds [a, i, k] \= + (if a = k then iterate k pdec2 i else pdec1 (iterate k pdec2 i))" + unfolding r_nth_inbounds_def by (simp add: Let_def) + then show "k \ a \ eval r_nth_inbounds [a, i, k] \= e_tuple_nth a i k" + and "eval r_nth_inbounds [a, i, k] \" + using e_tuple_nth_elementary by simp_all +qed + +definition "r_tuple_nth \ + Cn 3 r_ifle [Id 3 2, Id 3 0, r_nth_inbounds, r_constn 2 0]" + +lemma r_tuple_nth_prim: "prim_recfn 3 r_tuple_nth" + unfolding r_tuple_nth_def using r_nth_inbounds_prim by simp + +lemma r_tuple_nth [simp]: "eval r_tuple_nth [a, i, k] \= e_tuple_nth a i k" + unfolding r_tuple_nth_def using r_nth_inbounds_prim r_nth_inbounds by simp + + +subsection \Lists\ + + +subsubsection \Encoding and decoding\ + +text \Lists are encoded by pairing the length of the list with the code +for the tuple made up of the list's elements. Then all these codes are +incremented in order to make room for the empty list +(cf.~Rogers~\cite[p.~71]{Rogers87}).\ + +fun list_encode :: "nat list \ nat" where + "list_encode [] = 0" +| "list_encode (x # xs) = Suc (prod_encode (length xs, tuple_encode (length xs) (x # xs)))" + +lemma list_encode_0 [simp]: "list_encode xs = 0 \ xs = []" + using list_encode.elims by blast + +lemma list_encode_1: "list_encode [0] = 1" + by (simp add: prod_encode_def) + +fun list_decode :: "nat \ nat list" where + "list_decode 0 = []" +| "list_decode (Suc n) = tuple_decode (pdec1 n) (pdec2 n)" + +lemma list_encode_decode [simp]: "list_encode (list_decode n) = n" +proof (cases n) + case 0 + then show ?thesis by simp +next + case (Suc k) + then have *: "list_decode n = tuple_decode (pdec1 k) (pdec2 k)" (is "_ = ?t") + by simp + then obtain x xs where xxs: "x # xs = ?t" + by (metis tuple_decode.elims) + then have "list_encode ?t = list_encode (x # xs)" by simp + then have 1: "list_encode ?t = Suc (prod_encode (length xs, tuple_encode (length xs) (x # xs)))" + by simp + have 2: "length xs = length ?t - 1" + using xxs by (metis length_tl list.sel(3)) + then have 3: "length xs = pdec1 k" + using * by simp + then have "tuple_encode (length ?t - 1) ?t = pdec2 k" + using 2 tuple_encode_decode by metis + then have "list_encode ?t = Suc (prod_encode (pdec1 k, pdec2 k))" + using 1 2 3 xxs by simp + with * Suc show ?thesis by simp +qed + +lemma list_decode_encode [simp]: "list_decode (list_encode xs) = xs" +proof (cases xs) + case Nil + then show ?thesis by simp +next + case (Cons y ys) + then have "list_encode xs = + Suc (prod_encode (length ys, tuple_encode (length ys) xs))" + (is "_ = Suc ?i") + by simp + then have "list_decode (Suc ?i) = tuple_decode (pdec1 ?i) (pdec2 ?i)" by simp + moreover have "pdec1 ?i = length ys" by simp + moreover have "pdec2 ?i = tuple_encode (length ys) xs" by simp + ultimately have "list_decode (Suc ?i) = + tuple_decode (length ys) (tuple_encode (length ys) xs)" + by simp + moreover have "length ys = length xs - 1" + using Cons by simp + ultimately have "list_decode (Suc ?i) = + tuple_decode (length xs - 1) (tuple_encode (length xs - 1) xs)" + by simp + then show ?thesis using Cons by simp +qed + +abbreviation singleton_encode :: "nat \ nat" where + "singleton_encode x \ list_encode [x]" + +lemma list_decode_singleton: "list_decode (singleton_encode x) = [x]" + by simp + +definition "r_singleton_encode \ Cn 1 S [Cn 1 r_prod_encode [Z, Id 1 0]]" + +lemma r_singleton_encode_prim [simp]: "prim_recfn 1 r_singleton_encode" + unfolding r_singleton_encode_def by simp + +lemma r_singleton_encode [simp]: "eval r_singleton_encode [x] \= singleton_encode x" + unfolding r_singleton_encode_def by simp + +definition r_list_encode :: "nat \ recf" where + "r_list_encode n \ Cn (Suc n) S [Cn (Suc n) r_prod_encode [r_constn n n, r_tuple_encode n]]" + +lemma r_list_encode_prim [simp]: "prim_recfn (Suc n) (r_list_encode n)" + unfolding r_list_encode_def by simp + +lemma r_list_encode: + assumes "length xs = Suc n" + shows "eval (r_list_encode n) xs \= list_encode xs" +proof - + have "eval (r_tuple_encode n) xs \" + by (simp add: assms r_tuple_encode) + then have "eval (Cn (Suc n) r_prod_encode [r_constn n n, r_tuple_encode n]) xs \" + using assms by simp + then have "eval (r_list_encode n) xs = + eval S [the (eval (Cn (Suc n) r_prod_encode [r_constn n n, r_tuple_encode n]) xs)]" + unfolding r_list_encode_def using assms r_tuple_encode by simp + moreover from assms obtain y ys where "xs = y # ys" + by (meson length_Suc_conv) + ultimately show ?thesis + unfolding r_list_encode_def using assms r_tuple_encode by simp +qed + + +subsubsection \Functions on encoded lists\ + +text \The functions in this section mimic those on type @{typ "nat +list"}. Their names are prefixed by @{text e_} and the names of the +corresponding @{typ recf}s by @{text r_}.\ + +abbreviation e_tl :: "nat \ nat" where + "e_tl e \ list_encode (tl (list_decode e))" + +text \In order to turn @{term e_tl} into a partial recursive function +we first represent it in a more elementary way.\ + +lemma e_tl_elementary: + "e_tl e = + (if e = 0 then 0 + else if pdec1 (e - 1) = 0 then 0 + else Suc (prod_encode (pdec1 (e - 1) - 1, pdec22 (e - 1))))" +proof (cases e) + case 0 + then show ?thesis by simp +next + case Suc_d: (Suc d) + then show ?thesis + proof (cases "pdec1 d") + case 0 + then show ?thesis using Suc_d by simp + next + case (Suc a) + have *: "list_decode e = tuple_decode (pdec1 d) (pdec2 d)" + using Suc_d by simp + with Suc obtain x xs where xxs: "list_decode e = x # xs" by simp + then have **: "e_tl e = list_encode xs" by simp + have "list_decode (Suc (prod_encode (pdec1 (e - 1) - 1, pdec22 (e - 1)))) = + tuple_decode (pdec1 (e - 1) - 1) (pdec22 (e - 1))" + (is "?lhs = _") + by simp + also have "... = tuple_decode a (pdec22 (e - 1))" + using Suc Suc_d by simp + also have "... = tl (tuple_decode (Suc a) (pdec2 (e - 1)))" + using tuple_decode_pdec2 Suc by presburger + also have "... = tl (tuple_decode (pdec1 (e - 1)) (pdec2 (e - 1)))" + using Suc Suc_d by auto + also have "... = tl (list_decode e)" + using * Suc_d by simp + also have "... = xs" + using xxs by simp + finally have "?lhs = xs" . + then have "list_encode ?lhs = list_encode xs" by simp + then have "Suc (prod_encode (pdec1 (e - 1) - 1, pdec22 (e - 1))) = list_encode xs" + using list_encode_decode by metis + then show ?thesis using ** Suc_d Suc by simp + qed +qed + +definition "r_tl \ + let r = Cn 1 r_pdec1 [r_dec] + in Cn 1 r_ifz + [Id 1 0, + Z, + Cn 1 r_ifz + [r, Z, Cn 1 S [Cn 1 r_prod_encode [Cn 1 r_dec [r], Cn 1 r_pdec22 [r_dec]]]]]" + +lemma r_tl_prim [simp]: "prim_recfn 1 r_tl" + unfolding r_tl_def by (simp add: Let_def) + +lemma r_tl [simp]: "eval r_tl [e] \= e_tl e" + unfolding r_tl_def using e_tl_elementary by (simp add: Let_def) + +text \We define the head of the empty encoded list to be zero.\ + +definition e_hd :: "nat \ nat" where + "e_hd e \ if e = 0 then 0 else hd (list_decode e)" + +lemma e_hd [simp]: + assumes "list_decode e = x # xs" + shows "e_hd e = x" + using e_hd_def assms by auto + +lemma e_hd_0 [simp]: "e_hd 0 = 0" + using e_hd_def by simp + +lemma e_hd_neq_0 [simp]: + assumes "e \ 0" + shows "e_hd e = hd (list_decode e)" + using e_hd_def assms by simp + +definition "r_hd \ + Cn 1 r_ifz [Cn 1 r_pdec1 [r_dec], Cn 1 r_pdec2 [r_dec], Cn 1 r_pdec12 [r_dec]]" + +lemma r_hd_prim [simp]: "prim_recfn 1 r_hd" + unfolding r_hd_def by simp + +lemma r_hd [simp]: "eval r_hd [e] \= e_hd e" +proof - + have "e_hd e = (if pdec1 (e - 1) = 0 then pdec2 (e - 1) else pdec12 (e - 1))" + proof (cases e) + case 0 + then show ?thesis using pdec1_zero pdec2_le by auto + next + case (Suc d) + then show ?thesis by (cases "pdec1 d") (simp_all add: pdec1_zero) + qed + then show ?thesis unfolding r_hd_def by simp +qed + +abbreviation e_length :: "nat \ nat" where + "e_length e \ length (list_decode e)" + +lemma e_length_0: "e_length e = 0 \ e = 0" + by (metis list_encode.simps(1) length_0_conv list_encode_decode) + +definition "r_length \ Cn 1 r_ifz [Id 1 0, Z, Cn 1 S [Cn 1 r_pdec1 [r_dec]]]" + +lemma r_length_prim [simp]: "prim_recfn 1 r_length" + unfolding r_length_def by simp + +lemma r_length [simp]: "eval r_length [e] \= e_length e" + unfolding r_length_def by (cases e) simp_all + +text \Accessing an encoded list out of bounds yields zero.\ + +definition e_nth :: "nat \ nat \ nat" where + "e_nth e n \ if e = 0 then 0 else e_tuple_nth (pdec1 (e - 1)) (pdec2 (e - 1)) n" + +lemma e_nth [simp]: + "e_nth e n = (if n < e_length e then (list_decode e) ! n else 0)" + by (cases e) (simp_all add: e_nth_def e_tuple_nth_def) + +lemma e_hd_nth0: "e_hd e = e_nth e 0" + by (simp add: e_hd_def e_length_0 hd_conv_nth) + +definition "r_nth \ + Cn 2 r_ifz + [Id 2 0, + r_constn 1 0, + Cn 2 r_tuple_nth + [Cn 2 r_pdec1 [r_dummy 1 r_dec], Cn 2 r_pdec2 [r_dummy 1 r_dec], Id 2 1]]" + +lemma r_nth_prim [simp]: "prim_recfn 2 r_nth" + unfolding r_nth_def using r_tuple_nth_prim by simp + +lemma r_nth [simp]: "eval r_nth [e, n] \= e_nth e n" + unfolding r_nth_def e_nth_def using r_tuple_nth_prim by simp + +definition "r_rev_aux \ + Pr 1 r_hd (Cn 3 r_prod_encode [Cn 3 r_nth [Id 3 2, Cn 3 S [Id 3 0]], Id 3 1])" + +lemma r_rev_aux_prim: "prim_recfn 2 r_rev_aux" + unfolding r_rev_aux_def by simp + +lemma r_rev_aux: + assumes "list_decode e = xs" and "length xs > 0" and "i < length xs" + shows "eval r_rev_aux [i, e] \= tuple_encode i (rev (take (Suc i) xs))" + using assms(3) +proof (induction i) + case 0 + then show ?case + unfolding r_rev_aux_def using assms e_hd_def r_hd by (auto simp add: take_Suc) +next + case (Suc i) + let ?g = "Cn 3 r_prod_encode [Cn 3 r_nth [Id 3 2, Cn 3 S [Id 3 0]], Id 3 1]" + from Suc have "eval r_rev_aux [Suc i, e] = eval ?g [i, the (eval r_rev_aux [i, e]), e]" + unfolding r_rev_aux_def by simp + also have "... \= prod_encode (xs ! (Suc i), tuple_encode i (rev (take (Suc i) xs)))" + using Suc by (simp add: assms(1)) + finally show ?case by (simp add: Suc.prems take_Suc_conv_app_nth) +qed + +corollary r_rev_aux_full: + assumes "list_decode e = xs" and "length xs > 0" + shows "eval r_rev_aux [length xs - 1, e] \= tuple_encode (length xs - 1) (rev xs)" + using r_rev_aux assms by simp + +lemma r_rev_aux_total: "eval r_rev_aux [i, e] \" + using r_rev_aux_prim totalE by fastforce + +definition "r_rev \ + Cn 1 r_ifz + [Id 1 0, + Z, + Cn 1 S + [Cn 1 r_prod_encode + [Cn 1 r_dec [r_length], Cn 1 r_rev_aux [Cn 1 r_dec [r_length], Id 1 0]]]]" + +lemma r_rev_prim [simp]: "prim_recfn 1 r_rev" + unfolding r_rev_def using r_rev_aux_prim by simp + +lemma r_rev [simp]: "eval r_rev [e] \= list_encode (rev (list_decode e))" +proof - + let ?d = "Cn 1 r_dec [r_length]" + let ?a = "Cn 1 r_rev_aux [?d, Id 1 0]" + let ?p = "Cn 1 r_prod_encode [?d, ?a]" + let ?s = "Cn 1 S [?p]" + have eval_a: "eval ?a [e] = eval r_rev_aux [e_length e - 1, e]" + using r_rev_aux_prim by simp + then have "eval ?s [e] \" + using r_rev_aux_prim by (simp add: r_rev_aux_total) + then have *: "eval r_rev [e] \= (if e = 0 then 0 else the (eval ?s [e]))" + using r_rev_aux_prim by (simp add: r_rev_def) + show ?thesis + proof (cases "e = 0") + case True + then show ?thesis using * by simp + next + case False + then obtain xs where xs: "xs = list_decode e" "length xs > 0" + using e_length_0 by auto + then have len: "length xs = e_length e" by simp + with eval_a have "eval ?a [e] = eval r_rev_aux [length xs - 1, e]" + by simp + then have "eval ?a [e] \= tuple_encode (length xs - 1) (rev xs)" + using xs r_rev_aux_full by simp + then have "eval ?s [e] \= + Suc (prod_encode (length xs - 1, tuple_encode (length xs - 1) (rev xs)))" + using len r_rev_aux_prim by simp + then have "eval ?s [e] \= + Suc (prod_encode + (length (rev xs) - 1, tuple_encode (length (rev xs) - 1) (rev xs)))" + by simp + moreover have "length (rev xs) > 0" + using xs by simp + ultimately have "eval ?s [e] \= list_encode (rev xs)" + by (metis list_encode.elims diff_Suc_1 length_Cons length_greater_0_conv) + then show ?thesis using xs * by simp + qed +qed + +abbreviation e_cons :: "nat \ nat \ nat" where + "e_cons e es \ list_encode (e # list_decode es)" + +lemma e_cons_elementary: + "e_cons e es = + (if es = 0 then Suc (prod_encode (0, e)) + else Suc (prod_encode (e_length es, prod_encode (e, pdec2 (es - 1)))))" +proof (cases "es = 0") + case True + then show ?thesis by simp +next + case False + then have "e_length es = Suc (pdec1 (es - 1))" + by (metis list_decode.elims diff_Suc_1 tuple_decode_length) + moreover have "es = e_tl (list_encode (e # list_decode es))" + by (metis list.sel(3) list_decode_encode list_encode_decode) + ultimately show ?thesis + using False e_tl_elementary + by (metis list_decode.simps(2) diff_Suc_1 list_encode_decode prod.sel(1) + prod_encode_inverse snd_conv tuple_decode.simps(2)) +qed + +definition "r_cons_else \ + Cn 2 S + [Cn 2 r_prod_encode + [Cn 2 r_length + [Id 2 1], Cn 2 r_prod_encode [Id 2 0, Cn 2 r_pdec2 [Cn 2 r_dec [Id 2 1]]]]]" + +lemma r_cons_else_prim: "prim_recfn 2 r_cons_else" + unfolding r_cons_else_def by simp + +lemma r_cons_else: + "eval r_cons_else [e, es] \= + Suc (prod_encode (e_length es, prod_encode (e, pdec2 (es - 1))))" + unfolding r_cons_else_def by simp + +definition "r_cons \ + Cn 2 r_ifz + [Id 2 1, Cn 2 S [Cn 2 r_prod_encode [r_constn 1 0, Id 2 0]], r_cons_else]" + +lemma r_cons_prim [simp]: "prim_recfn 2 r_cons" + unfolding r_cons_def using r_cons_else_prim by simp + +lemma r_cons [simp]: "eval r_cons [e, es] \= e_cons e es" + unfolding r_cons_def using r_cons_else_prim r_cons_else e_cons_elementary by simp + +abbreviation e_snoc :: "nat \ nat \ nat" where + "e_snoc es e \ list_encode (list_decode es @ [e])" + +lemma e_nth_snoc_small [simp]: + assumes "n < e_length b" + shows "e_nth (e_snoc b z) n = e_nth b n" + using assms by (simp add: nth_append) + +lemma e_hd_snoc [simp]: + assumes "e_length b > 0" + shows "e_hd (e_snoc b x) = e_hd b" +proof - + from assms have "b \ 0" + using less_imp_neq by force + then have hd: "e_hd b = hd (list_decode b)" by simp + have "e_length (e_snoc b x) > 0" by simp + then have "e_snoc b x \ 0" + using not_gr_zero by fastforce + then have "e_hd (e_snoc b x) = hd (list_decode (e_snoc b x))" by simp + with assms hd show ?thesis by simp +qed + +definition "r_snoc \ Cn 2 r_rev [Cn 2 r_cons [Id 2 1, Cn 2 r_rev [Id 2 0]]]" + +lemma r_snoc_prim [simp]: "prim_recfn 2 r_snoc" + unfolding r_snoc_def by simp + +lemma r_snoc [simp]: "eval r_snoc [es, e] \= e_snoc es e" + unfolding r_snoc_def by simp + +abbreviation e_butlast :: "nat \ nat" where + "e_butlast e \ list_encode (butlast (list_decode e))" + +abbreviation e_take :: "nat \ nat \ nat" where + "e_take n x \ list_encode (take n (list_decode x))" + +definition "r_take \ + Cn 2 r_ifle + [Id 2 0, Cn 2 r_length [Id 2 1], + Pr 1 Z (Cn 3 r_snoc [Id 3 1, Cn 3 r_nth [Id 3 2, Id 3 0]]), + Id 2 1]" + +lemma r_take_prim [simp]: "prim_recfn 2 r_take" + unfolding r_take_def by simp_all + +lemma r_take: + assumes "x = list_encode es" + shows "eval r_take [n, x] \= list_encode (take n es)" +proof - + let ?g = "Cn 3 r_snoc [Id 3 1, Cn 3 r_nth [Id 3 2, Id 3 0]]" + let ?h = "Pr 1 Z ?g" + have "total ?h" using Mn_free_imp_total by simp + have "m \ length es \ eval ?h [m, x] \= list_encode (take m es)" for m + proof (induction m) + case 0 + then show ?case using assms r_take_def by (simp add: r_take_def) + next + case (Suc m) + then have "m < length es" by simp + then have "eval ?h [Suc m, x] = eval ?g [m, the (eval ?h [m, x]), x]" + using Suc r_take_def by simp + also have "... = eval ?g [m, list_encode (take m es), x]" + using Suc by simp + also have "... \= e_snoc (list_encode (take m es)) (es ! m)" + by (simp add: \m < length es\ assms) + also have "... \= list_encode ((take m es) @ [es ! m])" + using list_decode_encode by simp + also have "... \= list_encode (take (Suc m) es)" + by (simp add: \m < length es\ take_Suc_conv_app_nth) + finally show ?case . + qed + moreover have "eval (Id 2 1) [m, x] \= list_encode (take m es)" if "m > length es" for m + using that assms by simp + moreover have "eval r_take [m, x] \= + (if m \ e_length x then the (eval ?h [m, x]) else the (eval (Id 2 1) [m, x]))" + for m + unfolding r_take_def using `total ?h` by simp + ultimately show ?thesis unfolding r_take_def by fastforce +qed + +corollary r_take' [simp]: "eval r_take [n, x] \= e_take n x" + by (simp add: r_take) + +definition "r_last \ Cn 1 r_hd [r_rev]" + +lemma r_last_prim [simp]: "prim_recfn 1 r_last" + unfolding r_last_def by simp + +lemma r_last [simp]: + assumes "e = list_encode xs" and "length xs > 0" + shows "eval r_last [e] \= last xs" +proof - + from assms(2) have "length (rev xs) > 0" by simp + then have "list_encode (rev xs) > 0" + by (metis gr0I list.size(3) list_encode_0) + moreover have "eval r_last [e] = eval r_hd [the (eval r_rev [e])]" + unfolding r_last_def by simp + ultimately show ?thesis using assms hd_rev by auto +qed + +definition "r_update_aux \ + let + f = r_constn 2 0; + g = Cn 5 r_snoc + [Id 5 1, Cn 5 r_ifeq [Id 5 0, Id 5 3, Id 5 4, Cn 5 r_nth [Id 5 2, Id 5 0]]] + in Pr 3 f g" + +lemma r_update_aux_recfn: "recfn 4 r_update_aux" + unfolding r_update_aux_def by simp + +lemma r_update_aux: + assumes "n \ e_length b" + shows "eval r_update_aux [n, b, j, v] \= list_encode ((take n (list_decode b))[j:=v])" + using assms +proof (induction n) + case 0 + then show ?case unfolding r_update_aux_def by simp +next + case (Suc n) + then have n: "n < e_length b" + by simp + let ?a = "Cn 5 r_nth [Id 5 2, Id 5 0]" + let ?b = "Cn 5 r_ifeq [Id 5 0, Id 5 3, Id 5 4, ?a]" + define g where "g \ Cn 5 r_snoc [Id 5 1, ?b]" + then have g: "eval g [n, r, b, j, v] \= e_snoc r (if n = j then v else e_nth b n)" for r + by simp + + have "Pr 3 (r_constn 2 0) g = r_update_aux" + using r_update_aux_def g_def by simp + then have "eval r_update_aux [Suc n, b, j, v] = + eval g [n, the (eval r_update_aux [n, b, j, v]), b, j, v]" + using r_update_aux_recfn Suc n eval_Pr_converg_Suc + by (metis arity.simps(5) length_Cons list.size(3) nat_less_le + numeral_3_eq_3 option.simps(3)) + then have *: "eval r_update_aux [Suc n, b, j, v] \= e_snoc + (list_encode ((take n (list_decode b))[j:=v])) + (if n = j then v else e_nth b n)" + using g Suc by simp + + consider (j_eq_n) "j = n" | (j_less_n) "j < n" | (j_gt_n) "j > n" + by linarith + then show ?case + proof (cases) + case j_eq_n + moreover from this have "(take (Suc n) (list_decode b))[j:=v] = + (take n (list_decode b))[j:=v] @ [v]" + using n + by (metis length_list_update nth_list_update_eq take_Suc_conv_app_nth take_update_swap) + ultimately show ?thesis using * by simp + next + case j_less_n + moreover from this have "(take (Suc n) (list_decode b))[j:=v] = + (take n (list_decode b))[j:=v] @ [(list_decode b) ! n]" + using n + by (simp add: le_eq_less_or_eq list_update_append min_absorb2 take_Suc_conv_app_nth) + ultimately show ?thesis using * by auto + next + case j_gt_n + moreover from this have "(take (Suc n) (list_decode b))[j:=v] = + (take n (list_decode b))[j:=v] @ [(list_decode b) ! n]" + using n take_Suc_conv_app_nth by auto + ultimately show ?thesis using * by auto + qed +qed + +abbreviation e_update :: "nat \ nat \ nat \ nat" where + "e_update b j v \ list_encode ((list_decode b)[j:=v])" + +definition "r_update \ + Cn 3 r_update_aux [Cn 3 r_length [Id 3 0], Id 3 0, Id 3 1, Id 3 2]" + +lemma r_update_recfn [simp]: "recfn 3 r_update" + unfolding r_update_def using r_update_aux_recfn by simp + +lemma r_update [simp]: "eval r_update [b, j, v] \= e_update b j v" + unfolding r_update_def using r_update_aux r_update_aux_recfn by simp + +lemma e_length_update [simp]: "e_length (e_update b k v) = e_length b" + by simp + +definition e_append :: "nat \ nat \ nat" where + "e_append xs ys \ list_encode (list_decode xs @ list_decode ys)" + +lemma e_length_append: "e_length (e_append xs ys) = e_length xs + e_length ys" + using e_append_def by simp + +lemma e_nth_append_small: + assumes "n < e_length xs" + shows "e_nth (e_append xs ys) n = e_nth xs n" + using e_append_def assms by (simp add: nth_append) + +lemma e_nth_append_big: + assumes "n \ e_length xs" + shows "e_nth (e_append xs ys) n = e_nth ys (n - e_length xs)" + using e_append_def assms e_nth by (simp add: less_diff_conv2 nth_append) + +definition "r_append \ + let + f = Id 2 0; + g = Cn 4 r_snoc [Id 4 1, Cn 4 r_nth [Id 4 3, Id 4 0]] + in Cn 2 (Pr 2 f g) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]" + +lemma r_append_prim [simp]: "prim_recfn 2 r_append" + unfolding r_append_def by simp + +lemma r_append [simp]: "eval r_append [a, b] \= e_append a b" +proof - + define g where "g = Cn 4 r_snoc [Id 4 1, Cn 4 r_nth [Id 4 3, Id 4 0]]" + then have g: "eval g [j, r, a, b] \= e_snoc r (e_nth b j)" for j r + by simp + let ?h = "Pr 2 (Id 2 0) g" + have "eval ?h [n, a, b] \= list_encode (list_decode a @ (take n (list_decode b)))" + if "n \ e_length b" for n + using that g g_def by (induction n) (simp_all add: take_Suc_conv_app_nth) + then show ?thesis + unfolding r_append_def g_def e_append_def by simp +qed + +definition e_append_zeros :: "nat \ nat \ nat" where + "e_append_zeros b z \ e_append b (list_encode (replicate z 0))" + +lemma e_append_zeros_length: "e_length (e_append_zeros b z) = e_length b + z" + using e_append_def e_append_zeros_def by simp + +lemma e_nth_append_zeros: "e_nth (e_append_zeros b z) i = e_nth b i" + using e_append_zeros_def e_nth_append_small e_nth_append_big by auto + +lemma e_nth_append_zeros_big: + assumes "i \ e_length b" + shows "e_nth (e_append_zeros b z) i = 0" + unfolding e_append_zeros_def + using e_nth_append_big[of b i "list_encode (replicate z 0)", OF assms(1)] + by simp + +definition "r_append_zeros \ + r_swap (Pr 1 (Id 1 0) (Cn 3 r_snoc [Id 3 1, r_constn 2 0]))" + +lemma r_append_zeros_prim [simp]: "prim_recfn 2 r_append_zeros" + unfolding r_append_zeros_def by simp + +lemma r_append_zeros: "eval r_append_zeros [b, z] \= e_append_zeros b z" +proof - + let ?r = "Pr 1 (Id 1 0) (Cn 3 r_snoc [Id 3 1, r_constn 2 0])" + have "eval ?r [z, b] \= e_append_zeros b z" + using e_append_zeros_def e_append_def + by (induction z) (simp_all add: replicate_append_same) + then show ?thesis by (simp add: r_append_zeros_def) +qed + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/R1_BC.thy b/thys/Inductive_Inference/R1_BC.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/R1_BC.thy @@ -0,0 +1,534 @@ +section \@{term "\"} is not in BC\label{s:r1_bc}\ + +theory R1_BC + imports Lemma_R + CP_FIN_NUM (* for V0 *) +begin + +text \We show that @{term "U\<^sub>0 \ V\<^sub>0"} is not in BC, +which implies @{term "\ \ BC"}. + +The proof is by contradiction. Assume there is a strategy $S$ learning @{term +"U\<^sub>0 \ V\<^sub>0"} behaviorally correct in the limit with respect to our +standard Gödel numbering $\varphi$. Thanks to Lemma~R for BC we can assume +$S$ to be total. Then we construct a function in @{term "U\<^sub>0 \ V\<^sub>0"} for +which $S$ fails. + +As usual, there is a computable process building prefixes of functions +$\psi_j$. For every $j$ it starts with the singleton prefix $b = [j]$ and +computes the next prefix from a given prefix $b$ as follows: + +\begin{enumerate} +\item Simulate $\varphi_{S(b0^k)}(|b| + k)$ for increasing $k$ for an + increasing number of steps. +\item Once a $k$ with $\varphi_{S(b0^k)}(|b| + k) = 0$ is found, extend the + prefix by $0^k1$. +\end{enumerate} + +There is always such a $k$ because by assumption $S$ learns $b0^\infty \in +U_0$ and thus outputs a hypothesis for $b0^\infty$ on almost all of its +prefixes. Therefore for almost all prefixes of the form $b0^k$, we have +$\varphi_{S(b0^k)} = b0^\infty$ and hence $\varphi_{S(b0^k)}(|b| + k) = 0$. +But Step~2 constructs $\psi_j$ such that $\psi_j(|b| + k) = 1$. Therefore $S$ +does not hypothesize $\psi_j$ on the prefix $b0^k$ of $\psi_j$. And since the +process runs forever, $S$ outputs infinitely many incorrect hypotheses for +$\psi_j$ and thus does not learn $\psi_j$. + +Applying Kleene's fixed-point theorem to @{term "\ \ \\<^sup>2"} +yields a $j$ with $\varphi_j = \psi_j$ and thus $\psi_j \in V_0$. But $S$ +does not learn any $\psi_j$, contradicting our assumption. + +The result @{prop "\ \ BC"} can be obtained more directly by +running the process with the empty prefix, thereby constructing only one +function instead of a numbering. This function is in @{term R1}, and $S$ +fails to learn it by the same reasoning as above. The stronger statement +about @{term "U\<^sub>0 \ V\<^sub>0"} will be exploited in +Section~\ref{s:union}. + +In the following locale the assumption that $S$ learns @{term "U\<^sub>0"} +suffices for analyzing the process. However, in order to arrive at the +desired contradiction this assumption is too weak because the functions built +by the process are not in @{term "U\<^sub>0"}.\ + +locale r1_bc = + fixes s :: partial1 + assumes s_in_R1: "s \ \" and s_learn_U0: "learn_bc \ U\<^sub>0 s" +begin + +lemma s_learn_prenum: "\b. learn_bc \ {prenum b} s" + using s_learn_U0 U0_altdef learn_bc_closed_subseteq by blast + +text \A @{typ recf} for the strategy:\ + +definition r_s :: recf where + "r_s \ SOME rs. recfn 1 rs \ total rs \ s = (\x. eval rs [x])" + +lemma r_s_recfn [simp]: "recfn 1 r_s" + and r_s_total: "\x. eval r_s [x] \" + and eval_r_s: "\x. s x = eval r_s [x]" + using r_s_def R1_SOME[OF s_in_R1, of r_s] by simp_all + +text \We begin with the function that finds the $k$ from Step~1 of the +construction of $\psi$.\ + +definition "r_find_k \ + let k = Cn 2 r_pdec1 [Id 2 0]; + r = Cn 2 r_result1 + [Cn 2 r_pdec2 [Id 2 0], + Cn 2 r_s [Cn 2 r_append_zeros [Id 2 1, k]], + Cn 2 r_add [Cn 2 r_length [Id 2 1], k]] + in Cn 1 r_pdec1 [Mn 1 (Cn 2 r_eq [r, r_constn 1 1])]" + +lemma r_find_k_recfn [simp]: "recfn 1 r_find_k" + unfolding r_find_k_def by (simp add: Let_def) + +text \There is always a suitable $k$, since the strategy learns +$b0^\infty$ for all $b$.\ + +lemma learn_bc_prenum_eventually_zero: + "\k. \ (the (s (e_append_zeros b k))) (e_length b + k) \= 0" +proof - + let ?f = "prenum b" + have "\n\e_length b. \ (the (s (?f \ n))) = ?f" + using learn_bcE s_learn_prenum by (meson le_cases singletonI) + then obtain n where n: "n \ e_length b" "\ (the (s (?f \ n))) = ?f" + by auto + define k where "k = Suc n - e_length b" + let ?e = "e_append_zeros b k" + have len: "e_length ?e = Suc n" + using k_def n e_append_zeros_length by simp + have "?f \ n = ?e" + proof - + have "e_length ?e > 0" + using len n(1) by simp + moreover have "?f x \= e_nth ?e x" for x + proof (cases "x < e_length b") + case True + then show ?thesis using e_nth_append_zeros by simp + next + case False + then have "?f x \= 0" by simp + moreover from False have "e_nth ?e x = 0" + using e_nth_append_zeros_big by simp + ultimately show ?thesis by simp + qed + ultimately show ?thesis using initI[of "?e"] len by simp + qed + with n(2) have "\ (the (s ?e)) = ?f" by simp + then have "\ (the (s ?e)) (e_length ?e) \= 0" + using len n(1) by auto + then show ?thesis using e_append_zeros_length by auto +qed + +lemma if_eq_eq: "(if v = 1 then (0 :: nat) else 1) = 0 \ v = 1" + by presburger + +lemma r_find_k: + shows "eval r_find_k [b] \" + and "let k = the (eval r_find_k [b]) + in \ (the (s (e_append_zeros b k))) (e_length b + k) \= 0" +proof - + let ?k = "Cn 2 r_pdec1 [Id 2 0]" + let ?argt = "Cn 2 r_pdec2 [Id 2 0]" + let ?argi = "Cn 2 r_s [Cn 2 r_append_zeros [Id 2 1, ?k]]" + let ?argx = "Cn 2 r_add [Cn 2 r_length [Id 2 1], ?k]" + let ?r = "Cn 2 r_result1 [?argt, ?argi, ?argx]" + define f where "f \ + let k = Cn 2 r_pdec1 [Id 2 0]; + r = Cn 2 r_result1 + [Cn 2 r_pdec2 [Id 2 0], + Cn 2 r_s [Cn 2 r_append_zeros [Id 2 1, k]], + Cn 2 r_add [Cn 2 r_length [Id 2 1], k]] + in Cn 2 r_eq [r, r_constn 1 1]" + then have "recfn 2 f" by (simp add: Let_def) + have "total r_s" + by (simp add: r_s_total totalI1) + then have "total f" + unfolding f_def using Cn_total Mn_free_imp_total by (simp add: Let_def) + + have "eval ?argi [z, b] = s (e_append_zeros b (pdec1 z))" for z + using r_append_zeros \recfn 2 f\ eval_r_s by auto + then have "eval ?argi [z, b] \= the (s (e_append_zeros b (pdec1 z)))" for z + using eval_r_s r_s_total by simp + moreover have "recfn 2 ?r" using \recfn 2 f\ by auto + ultimately have r: "eval ?r [z, b] = + eval r_result1 [pdec2 z, the (s (e_append_zeros b (pdec1 z))), e_length b + pdec1 z]" + for z + by simp + then have f: "eval f [z, b] \= (if the (eval ?r [z, b]) = 1 then 0 else 1)" for z + using f_def `recfn 2 f` prim_recfn_total by (auto simp add: Let_def) + + have "\k. \ (the (s (e_append_zeros b k))) (e_length b + k) \= 0" + using s_learn_prenum learn_bc_prenum_eventually_zero by auto + then obtain k where "\ (the (s (e_append_zeros b k))) (e_length b + k) \= 0" + by auto + then obtain t where "eval r_result1 [t, the (s (e_append_zeros b k)), e_length b + k] \= Suc 0" + using r_result1_converg_phi(1) by blast + then have t: "eval r_result1 [t, the (s (e_append_zeros b k)), e_length b + k] \= Suc 0" + by simp + + let ?z = "prod_encode (k, t)" + have "eval ?r [?z, b] \= Suc 0" + using t r by (metis fst_conv prod_encode_inverse snd_conv) + with f have fzb: "eval f [?z, b] \= 0" by simp + moreover have "eval (Mn 1 f) [b] = + (if (\z. eval f ([z, b]) \= 0) + then Some (LEAST z. eval f [z, b] \= 0) + else None)" + using eval_Mn_total[of 1 f "[b]"] `total f` `recfn 2 f` by simp + ultimately have mn1f: "eval (Mn 1 f) [b] \= (LEAST z. eval f [z, b] \= 0)" + by auto + with fzb have "eval f [the (eval (Mn 1 f) [b]), b] \= 0" (is "eval f [?zz, b] \= 0") + using \total f\ \recfn 2 f\ LeastI_ex[of "%z. eval f [z, b] \= 0"] by auto + moreover have "eval f [?zz, b] \= (if the (eval ?r [?zz, b]) = 1 then 0 else 1)" + using f by simp + ultimately have "(if the (eval ?r [?zz, b]) = 1 then (0 :: nat) else 1) = 0" by auto + then have "the (eval ?r [?zz, b]) = 1" + using if_eq_eq[of "the (eval ?r [?zz, b])"] by simp + then have + "eval r_result1 + [pdec2 ?zz, the (s (e_append_zeros b (pdec1 ?zz))), e_length b + pdec1 ?zz] \= + 1" + using r r_result1_total r_result1_prim totalE + by (metis length_Cons list.size(3) numeral_3_eq_3 option.collapse) + then have *: "\ (the (s (e_append_zeros b (pdec1 ?zz)))) (e_length b + pdec1 ?zz) \= 0" + by (simp add: r_result1_some_phi) + + define Mn1f where "Mn1f = Mn 1 f" + then have "eval Mn1f [b] \= ?zz" + using mn1f by auto + moreover have "recfn 1 (Cn 1 r_pdec1 [Mn1f])" + using `recfn 2 f` Mn1f_def by simp + ultimately have "eval (Cn 1 r_pdec1 [Mn1f]) [b] = eval r_pdec1 [the (eval (Mn1f) [b])]" + by auto + then have "eval (Cn 1 r_pdec1 [Mn1f]) [b] = eval r_pdec1 [?zz]" + using Mn1f_def by blast + then have 1: "eval (Cn 1 r_pdec1 [Mn1f]) [b] \= pdec1 ?zz" + by simp + moreover have "recfn 1 (Cn 1 S [Cn 1 r_pdec1 [Mn1f]])" + using `recfn 2 f` Mn1f_def by simp + ultimately have "eval (Cn 1 S [Cn 1 r_pdec1 [Mn1f]]) [b] = + eval S [the (eval (Cn 1 r_pdec1 [Mn1f]) [b])]" + by simp + then have "eval (Cn 1 S [Cn 1 r_pdec1 [Mn1f]]) [b] = eval S [pdec1 ?zz]" + using 1 by simp + then have "eval (Cn 1 S [Cn 1 r_pdec1 [Mn1f]]) [b] \= Suc (pdec1 ?zz)" + by simp + moreover have "eval r_find_k [b] = eval (Cn 1 r_pdec1 [Mn1f]) [b]" + unfolding r_find_k_def Mn1f_def f_def by metis + ultimately have r_find_ksb: "eval r_find_k [b] \= pdec1 ?zz" + using 1 by simp + then show "eval r_find_k [b] \" by simp_all + + from r_find_ksb have "the (eval r_find_k [b]) = pdec1 ?zz" + by simp + moreover have "\ (the (s (e_append_zeros b (pdec1 ?zz)))) (e_length b + pdec1 ?zz) \= 0" + using * by simp + ultimately show "let k = the (eval r_find_k [b]) + in \ (the (s (e_append_zeros b k))) (e_length b + k) \= 0" + by simp +qed + +lemma r_find_k_total: "total r_find_k" + by (simp add: s_learn_prenum r_find_k(1) totalI1) + +text \The following function represents one iteration of the +process.\ + +abbreviation "r_next \ + Cn 3 r_snoc [Cn 3 r_append_zeros [Id 3 1, Cn 3 r_find_k [Id 3 1]], r_constn 2 1]" + +text \Using @{term r_next} we define the function @{term r_prefixes} +that computes the prefix after every iteration of the process.\ + +definition r_prefixes :: recf where + "r_prefixes \ Pr 1 r_singleton_encode r_next" + +lemma r_prefixes_recfn: "recfn 2 r_prefixes" + unfolding r_prefixes_def by simp + +lemma r_prefixes_total: "total r_prefixes" +proof - + have "recfn 3 r_next" by simp + then have "total r_next" + using `recfn 3 r_next` r_find_k_total Cn_total Mn_free_imp_total by auto + then show ?thesis + by (simp add: Mn_free_imp_total Pr_total r_prefixes_def) +qed + +lemma r_prefixes_0: "eval r_prefixes [0, j] \= list_encode [j]" + unfolding r_prefixes_def by simp + +lemma r_prefixes_Suc: + "eval r_prefixes [Suc n, j] \= + (let b = the (eval r_prefixes [n, j]) + in e_snoc (e_append_zeros b (the (eval r_find_k [b]))) 1)" +proof - + have "recfn 3 r_next" by simp + then have "total r_next" + using `recfn 3 r_next` r_find_k_total Cn_total Mn_free_imp_total by auto + have eval_next: "eval r_next [t, v, j] \= + e_snoc (e_append_zeros v (the (eval r_find_k [v]))) 1" + for t v j + using r_find_k_total `recfn 3 r_next` r_append_zeros by simp + then have "eval r_prefixes [Suc n, j] = eval r_next [n, the (eval r_prefixes [n, j]), j]" + using r_prefixes_total by (simp add: r_prefixes_def) + then show "eval r_prefixes [Suc n, j] \= + (let b = the (eval r_prefixes [n, j]) + in e_snoc (e_append_zeros b (the (eval r_find_k [b]))) 1)" + using eval_next by metis +qed + +text \Since @{term r_prefixes} is total, we can get away with +introducing a total function.\ + +definition prefixes :: "nat \ nat \ nat" where + "prefixes j t \ the (eval r_prefixes [t, j])" + +lemma prefixes_Suc: + "prefixes j (Suc t) = + e_snoc (e_append_zeros (prefixes j t) (the (eval r_find_k [prefixes j t]))) 1" + unfolding prefixes_def using r_prefixes_Suc by (simp_all add: Let_def) + +lemma prefixes_Suc_length: + "e_length (prefixes j (Suc t)) = + Suc (e_length (prefixes j t) + the (eval r_find_k [prefixes j t]))" + using e_append_zeros_length prefixes_Suc by simp + +lemma prefixes_length_mono: "e_length (prefixes j t) < e_length (prefixes j (Suc t))" + using prefixes_Suc_length by simp + +lemma prefixes_length_mono': "e_length (prefixes j t) \ e_length (prefixes j (t + d))" +proof (induction d) + case 0 + then show ?case by simp +next + case (Suc d) + then show ?case using prefixes_length_mono le_less_trans by fastforce +qed + +lemma prefixes_length_lower_bound: "e_length (prefixes j t) \ Suc t" +proof (induction t) + case 0 + then show ?case by (simp add: prefixes_def r_prefixes_0) +next + case (Suc t) + moreover have "Suc (e_length (prefixes j t)) \ e_length (prefixes j (Suc t))" + using prefixes_length_mono by (simp add: Suc_leI) + ultimately show ?case by simp +qed + +lemma prefixes_Suc_nth: + assumes "x < e_length (prefixes j t)" + shows "e_nth (prefixes j t) x = e_nth (prefixes j (Suc t)) x" +proof - + define k where "k = the (eval r_find_k [prefixes j t])" + let ?u = "e_append_zeros (prefixes j t) k" + have "prefixes j (Suc t) = + e_snoc (e_append_zeros (prefixes j t) (the (eval r_find_k [prefixes j t]))) 1" + using prefixes_Suc by simp + with k_def have "prefixes j (Suc t) = e_snoc ?u 1" + by simp + then have "e_nth (prefixes j (Suc t)) x = e_nth (e_snoc ?u 1) x" + by simp + moreover have "x < e_length ?u" + using assms e_append_zeros_length by auto + ultimately have "e_nth (prefixes j (Suc t)) x = e_nth ?u x" + using e_nth_snoc_small by simp + moreover have "e_nth ?u x = e_nth (prefixes j t) x" + using assms e_nth_append_zeros by simp + ultimately show "e_nth (prefixes j t) x = e_nth (prefixes j (Suc t)) x" + by simp +qed + +lemma prefixes_Suc_last: "e_nth (prefixes j (Suc t)) (e_length (prefixes j (Suc t)) - 1) = 1" + using prefixes_Suc by simp + +lemma prefixes_le_nth: + assumes "x < e_length (prefixes j t)" + shows "e_nth (prefixes j t) x = e_nth (prefixes j (t + d)) x" +proof (induction d) + case 0 + then show ?case by simp +next + case (Suc d) + have "x < e_length (prefixes j (t + d))" + using s_learn_prenum assms prefixes_length_mono' + by (simp add: less_eq_Suc_le order_trans_rules(23)) + then have "e_nth (prefixes j (t + d)) x = e_nth (prefixes j (t + Suc d)) x" + using prefixes_Suc_nth by simp + with Suc show ?case by simp +qed + +text \The numbering $\psi$ is defined via @{term[names_short] prefixes}.\ + +definition psi :: partial2 ("\") where + "\ j x \ Some (e_nth (prefixes j (Suc x)) x)" + +lemma psi_in_R2: "\ \ \\<^sup>2" +proof + define r where "r \ Cn 2 r_nth [Cn 2 r_prefixes [Cn 2 S [Id 2 1], Id 2 0], Id 2 1]" + then have "recfn 2 r" + using r_prefixes_recfn by simp + then have "eval r [j, x] \= e_nth (prefixes j (Suc x)) x" for j x + unfolding r_def prefixes_def using r_prefixes_total r_prefixes_recfn e_nth by simp + then have "eval r [j, x] = \ j x" for j x + unfolding psi_def by simp + then show "\ \ \

\<^sup>2" + using `recfn 2 r` by auto + show "total2 \" + unfolding psi_def by auto +qed + +lemma psi_eq_nth_prefixes: + assumes "x < e_length (prefixes j t)" + shows "\ j x \= e_nth (prefixes j t) x" +proof (cases "Suc x < t") + case True + have "x \ e_length (prefixes j x)" + using prefixes_length_lower_bound by (simp add: Suc_leD) + also have "... < e_length (prefixes j (Suc x))" + using prefixes_length_mono s_learn_prenum by simp + finally have "x < e_length (prefixes j (Suc x))" . + with True have "e_nth (prefixes j (Suc x)) x = e_nth (prefixes j t) x" + using prefixes_le_nth[of x j "Suc x" "t - Suc x"] by simp + then show ?thesis using psi_def by simp +next + case False + then have "e_nth (prefixes j (Suc x)) x = e_nth (prefixes j t) x" + using prefixes_le_nth[of x j t "Suc x - t"] assms by simp + then show ?thesis using psi_def by simp +qed + +lemma psi_at_0: "\ j 0 \= j" + using psi_eq_nth_prefixes[of 0 j 0] prefixes_length_lower_bound[of 0 j] + by (simp add: prefixes_def r_prefixes_0) + +text \The prefixes output by the process @{term[names_short] "prefixes j"} are +indeed prefixes of $\psi_j$.\ + +lemma prefixes_init_psi: "\ j \ (e_length (prefixes j (Suc t)) - 1) = prefixes j (Suc t)" +proof (rule initI[of "prefixes j (Suc t)"]) + let ?e = "prefixes j (Suc t)" + show "e_length ?e > 0" + using prefixes_length_lower_bound[of "Suc t" j] by auto + show "\x. x < e_length ?e \ \ j x \= e_nth ?e x" + using prefixes_Suc_nth psi_eq_nth_prefixes by simp +qed + +text \Every prefix of $\psi_j$ generated by the process +@{term[names_short] "prefixes j"} (except for the initial one) is of the form +$b0^k1$. But $k$ is chosen such that $\varphi_{S(b0^k)}(|b|+k) = 0 \neq 1 = +b0^k1_{|b|+k}$. Therefore the hypothesis $S(b0^k)$ is incorrect for +$\psi_j$.\ + +lemma hyp_wrong_at_last: + "\ (the (s (e_butlast (prefixes j (Suc t))))) (e_length (prefixes j (Suc t)) - 1) \ + \ j (e_length (prefixes j (Suc t)) - 1)" + (is "?lhs \ ?rhs") +proof - + let ?b = "prefixes j t" + let ?k = "the (eval r_find_k [?b])" + let ?x = "e_length (prefixes j (Suc t)) - 1" + have "e_butlast (prefixes j (Suc t)) = e_append_zeros ?b ?k" + using s_learn_prenum prefixes_Suc by simp + then have "?lhs = \ (the (s (e_append_zeros ?b ?k))) ?x" + by simp + moreover have "?x = e_length ?b + ?k" + using prefixes_Suc_length by simp + ultimately have "?lhs = \ (the (s (e_append_zeros ?b ?k))) (e_length ?b + ?k)" + by simp + then have "?lhs \= 0" + using r_find_k(2) r_s_total s_learn_prenum by metis + moreover have "?x < e_length (prefixes j (Suc t))" + using prefixes_length_lower_bound le_less_trans linorder_not_le s_learn_prenum + by fastforce + ultimately have "?rhs \= e_nth (prefixes j (Suc t)) ?x" + using psi_eq_nth_prefixes[of ?x j "Suc t"] by simp + moreover have "e_nth (prefixes j (Suc t)) ?x = 1" + using prefixes_Suc prefixes_Suc_last by simp + ultimately have "?rhs \= 1" by simp + with `?lhs \= 0` show ?thesis by simp +qed + +corollary hyp_wrong: "\ (the (s (e_butlast (prefixes j (Suc t))))) \ \ j" + using hyp_wrong_at_last[of j t] by auto + +text \For all $j$, the strategy $S$ outputs infinitely many wrong hypotheses for +$\psi_j$\ + +lemma infinite_hyp_wrong: "\m>n. \ (the (s (\ j \ m))) \ \ j" +proof - + let ?b = "prefixes j (Suc (Suc n))" + let ?bb = "e_butlast ?b" + have len_b: "e_length ?b > Suc (Suc n)" + using prefixes_length_lower_bound by (simp add: Suc_le_lessD) + then have len_bb: "e_length ?bb > Suc n" by simp + define m where "m = e_length ?bb - 1" + with len_bb have "m > n" by simp + have "\ j \ m = ?bb" + proof - + have "\ j \ (e_length ?b - 1) = ?b" + using prefixes_init_psi by simp + then have "\ j \ (e_length ?b - 2) = ?bb" + using init_butlast_init psi_in_R2 R2_proj_R1 R1_imp_total1 len_bb length_init + by (metis Suc_1 diff_diff_left length_butlast length_greater_0_conv + list.size(3) list_decode_encode not_less0 plus_1_eq_Suc) + then show ?thesis by (metis diff_Suc_1 length_init m_def) + qed + moreover have "\ (the (s ?bb)) \ \ j" + using hyp_wrong by simp + ultimately have "\ (the (s (\ j \ m))) \ \ j" + by simp + with `m > n` show ?thesis by auto +qed + +lemma U0_V0_not_learn_bc: "\ learn_bc \ (U\<^sub>0 \ V\<^sub>0) s" +proof - + obtain j where j: "\ j = \ j" + using R2_imp_P2 kleene_fixed_point psi_in_R2 by blast + moreover have "\m>n. \ (the (s ((\ j) \ m))) \ \ j" for n + using infinite_hyp_wrong[of _ j] by simp + ultimately have "\ learn_bc \ {\ j} s" + using infinite_hyp_wrong_not_BC by simp + moreover have "\ j \ V\<^sub>0" + proof - + have "\ j \ \" (is "?f \ \") + using psi_in_R2 by simp + moreover have "\ (the (?f 0)) = ?f" + using j psi_at_0[of j] by simp + ultimately show ?thesis by (simp add: V0_def) + qed + ultimately show "\ learn_bc \ (U\<^sub>0 \ V\<^sub>0) s" + using learn_bc_closed_subseteq by auto +qed + +end + +lemma U0_V0_not_in_BC: "U\<^sub>0 \ V\<^sub>0 \ BC" +proof + assume in_BC: "U\<^sub>0 \ V\<^sub>0 \ BC" + then have "U\<^sub>0 \ V\<^sub>0 \ BC_wrt \" + using BC_wrt_phi_eq_BC by simp + then obtain s where "learn_bc \ (U\<^sub>0 \ V\<^sub>0) s" + using BC_wrt_def by auto + then obtain s' where s': "s' \ \" "learn_bc \ (U\<^sub>0 \ V\<^sub>0) s'" + using lemma_R_for_BC_simple by blast + then have learn_U0: "learn_bc \ U\<^sub>0 s'" + using learn_bc_closed_subseteq[of \ "U\<^sub>0 \ V\<^sub>0" "s'"] by simp + then interpret r1_bc s' + by (simp add: r1_bc_def s'(1)) + have "\ learn_bc \ (U\<^sub>0 \ V\<^sub>0) s'" + using learn_bc_closed_subseteq U0_V0_not_learn_bc by simp + with s'(2) show False by simp +qed + +theorem R1_not_in_BC: "\ \ BC" +proof - + have "U\<^sub>0 \ V\<^sub>0 \ \" + using V0_def U0_in_NUM by auto + then show ?thesis + using U0_V0_not_in_BC BC_closed_subseteq by auto +qed + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/ROOT b/thys/Inductive_Inference/ROOT new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/ROOT @@ -0,0 +1,21 @@ +chapter AFP + +session Inductive_Inference (AFP) = HOL + + options [timeout = 600] + sessions + "HOL-Library" + theories + Partial_Recursive + Universal + Standard_Results + Inductive_Inference_Basics + CP_FIN_NUM + CONS_LIM + Lemma_R + LIM_BC + TOTAL_CONS + R1_BC + Union + document_files + "root.tex" + "root.bib" diff --git a/thys/Inductive_Inference/Standard_Results.thy b/thys/Inductive_Inference/Standard_Results.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/Standard_Results.thy @@ -0,0 +1,1594 @@ +theory Standard_Results + imports Universal +begin + +section \Kleene normal form and the number of $\mu$-operations\ + +text \Kleene's original normal form theorem~\cite{Kleene43} states that +every partial recursive $f$ can be expressed as $f(x) = u(\mu y[t(i, x, y) = +0]$ for some $i$, where $u$ and $t$ are specially crafted primitive recursive +functions tied to Kleene's definition of partial recursive functions. +Rogers~\cite[p.~29f.]{Rogers87} relaxes the theorem by allowing $u$ and $t$ +to be any primitive recursive functions of arity one and three, respectively. +Both versions require a separate $t$-predicate for every arity. We will show +a unified version for all arities by treating $x$ as an encoded list of +arguments. + +Our universal function @{thm[display,names_short] "r_univ_def"} can represent +all partial recursive functions (see theorem @{thm[source] r_univ}). Moreover +@{term "r_result"}, @{term "r_dec"}, and @{term "r_not"} are primitive +recursive. As such @{term r_univ} could almost serve as the right-hand side +$u(\mu y[t(i, x, y) = 0]$. Its only flaw is that the outer function, the +composition of @{term r_dec} and @{term r_result}, is ternary rather than +unary.\ + +lemma r_univ_almost_kleene_nf: + "r_univ \ + (let u = Cn 3 r_dec [r_result]; + t = Cn 3 r_not [r_result] + in Cn 2 u [Mn 2 t, Id 2 0, Id 2 1])" + unfolding r_univ_def by (rule exteqI) simp_all + +text \We can remedy the wrong arity with some encoding and +projecting.\ + +definition r_nf_t :: recf where + "r_nf_t \ Cn 3 r_and + [Cn 3 r_eq [Cn 3 r_pdec2 [Id 3 0], Cn 3 r_prod_encode [Id 3 1, Id 3 2]], + Cn 3 r_not + [Cn 3 r_result + [Cn 3 r_pdec1 [Id 3 0], + Cn 3 r_pdec12 [Id 3 0], + Cn 3 r_pdec22 [Id 3 0]]]]" + +lemma r_nf_t_prim: "prim_recfn 3 r_nf_t" + unfolding r_nf_t_def by simp + +definition r_nf_u :: recf where + "r_nf_u \ Cn 1 r_dec [Cn 1 r_result [r_pdec1, r_pdec12, r_pdec22]]" + +lemma r_nf_u_prim: "prim_recfn 1 r_nf_u" + unfolding r_nf_u_def by simp + +lemma r_nf_t_0: + assumes "eval r_result [pdec1 y, pdec12 y, pdec22 y] \\ 0" + and "pdec2 y = prod_encode (i, x)" + shows "eval r_nf_t [y, i, x] \= 0" + unfolding r_nf_t_def using assms by auto + +lemma r_nf_t_1: + assumes "eval r_result [pdec1 y, pdec12 y, pdec22 y] \= 0 \ pdec2 y \ prod_encode (i, x)" + shows "eval r_nf_t [y, i, x] \= 1" + unfolding r_nf_t_def using assms r_result_total by auto + +text \The next function is just as universal as @{term r_univ}, but +satisfies the conditions of the Kleene normal form theorem because the +outer funtion @{term r_nf_u} is unary.\ + +definition "r_normal_form \ Cn 2 r_nf_u [Mn 2 r_nf_t]" + +lemma r_normal_form_recfn: "recfn 2 r_normal_form" + unfolding r_normal_form_def using r_nf_u_prim r_nf_t_prim by simp + +lemma r_univ_exteq_r_normal_form: "r_univ \ r_normal_form" +proof (rule exteqI) + show arity: "arity r_univ = arity r_normal_form" + using r_normal_form_recfn by simp + show "eval r_univ xs = eval r_normal_form xs" if "length xs = arity r_univ" for xs + proof - + have "length xs = 2" + using that by simp + then obtain i x where ix: "[i, x] = xs" + by (metis length_0_conv length_Suc_conv numeral_2_eq_2) + have "eval r_univ [i, x] = eval r_normal_form [i, x]" + proof (cases "\t. eval r_result [t, i, x] \= 0") + case True + then have "eval r_univ [i, x] \" + unfolding r_univ_def by simp + moreover have "eval r_normal_form [i, x] \" + proof - + have "eval r_nf_t [y, i, x] \= 1" for y + using True r_nf_t_1[of y i x] by fastforce + then show ?thesis + unfolding r_normal_form_def using r_nf_u_prim r_nf_t_prim by simp + qed + ultimately show ?thesis by simp + next + case False + then have "\t. eval r_result [t, i, x] \\ 0" + by (simp add: r_result_total) + then obtain t where "eval r_result [t, i, x] \\ 0" + by auto + then have "eval r_nf_t [triple_encode t i x, i, x] \= 0" + using r_nf_t_0 by simp + then obtain y where y: "eval (Mn 2 r_nf_t) [i, x] \= y" + using r_nf_t_prim Mn_free_imp_total by fastforce + then have "eval r_nf_t [y, i, x] \= 0" + using r_nf_t_prim Mn_free_imp_total eval_Mn_convergE(2)[of 2 r_nf_t "[i, x]" y] + by simp + then have r_result: "eval r_result [pdec1 y, pdec12 y, pdec22 y] \\ 0" + and pdec2: "pdec2 y = prod_encode (i, x)" + using r_nf_t_0[of y i x] r_nf_t_1[of y i x] r_result_total by auto + then have "eval r_result [pdec1 y, i, x] \\ 0" + by simp + then obtain v where v: + "eval r_univ [pdec12 y, pdec22 y] \= v" + "eval r_result [pdec1 y, pdec12 y, pdec22 y] \= Suc v" + using r_result r_result_bivalent'[of "pdec12 y" "pdec22 y" _ "pdec1 y"] + r_result_diverg'[of "pdec12 y" "pdec22 y" "pdec1 y"] + by auto + + have "eval r_normal_form [i, x] = eval r_nf_u [y]" + unfolding r_normal_form_def using y r_nf_t_prim r_nf_u_prim by simp + also have "... = eval r_dec [the (eval (Cn 1 r_result [r_pdec1, r_pdec12, r_pdec22]) [y])]" + unfolding r_nf_u_def using r_result by simp + also have "... = eval r_dec [Suc v]" + using v by simp + also have "... \= v" + by simp + finally have "eval r_normal_form [i, x] \= v" . + moreover have "eval r_univ [i, x] \= v" + using v(1) pdec2 by simp + ultimately show ?thesis by simp + qed + with ix show ?thesis by simp + qed +qed + +theorem normal_form: + assumes "recfn n f" + obtains i where "\x. e_length x = n \ eval r_normal_form [i, x] = eval f (list_decode x)" +proof - + have "eval r_normal_form [encode f, x] = eval f (list_decode x)" if "e_length x = n" for x + using r_univ_exteq_r_normal_form assms that exteq_def r_univ' by auto + then show ?thesis using that by auto +qed + +text \As a consequence of the normal form theorem every partial +recursive function can be represented with exactly one application of the +$\mu$-operator.\ + +fun count_Mn :: "recf \ nat" where + "count_Mn Z = 0" +| "count_Mn S = 0" +| "count_Mn (Id m n) = 0" +| "count_Mn (Cn n f gs) = count_Mn f + sum_list (map count_Mn gs)" +| "count_Mn (Pr n f g) = count_Mn f + count_Mn g" +| "count_Mn (Mn n f) = Suc (count_Mn f)" + +lemma count_Mn_zero_iff_prim: "count_Mn f = 0 \ Mn_free f" + by (induction f) auto + +text \The normal form has only one $\mu$-recursion.\ + +lemma count_Mn_normal_form: "count_Mn r_normal_form = 1" + unfolding r_normal_form_def r_nf_u_def r_nf_t_def using count_Mn_zero_iff_prim by simp + +lemma one_Mn_suffices: + assumes "recfn n f" + shows "\g. count_Mn g = 1 \ g \ f" +proof - + have "n > 0" + using assms wellf_arity_nonzero by auto + obtain i where i: + "\x. e_length x = n \ eval r_normal_form [i, x] = eval f (list_decode x)" + using normal_form[OF assms(1)] by auto + define g where "g \ Cn n r_normal_form [r_constn (n - 1) i, r_list_encode (n - 1)]" + then have "recfn n g" + using r_normal_form_recfn `n > 0` by simp + then have "g \ f" + using g_def r_list_encode i assms by (intro exteqI) simp_all + moreover have "count_Mn g = 1" + unfolding g_def using count_Mn_normal_form count_Mn_zero_iff_prim by simp + ultimately show ?thesis by auto +qed + +text \The previous lemma could have been obtained without @{term +"r_normal_form"} directly from @{term "r_univ"}.\ + + +section \The $s$-$m$-$n$ theorem\ + +text \For all $m, n > 0$ there is an $(m + 1)$-ary primitive recursive +function $s^m_n$ with +\[ + \varphi_p^{(m + n)}(c_1, \dots,c_m, x_1, \dots, x_n) = + \varphi_{s^m_n(p, c_1, \dots,c_m)}^{(n)}(x_1, \dots, x_n) +\] +for all $p, c_1, \ldots, c_m, x_1, \ldots, x_n$. Here, $\varphi^{(n)}$ is a +function universal for $n$-ary partial recursive functions, which we will +represent by @{term "r_universal n"}\ + +text \The $s^m_n$ functions compute codes of functions. We start simple: +computing codes of the unary constant functions.\ + +fun code_const1 :: "nat \ nat" where + "code_const1 0 = 0" +| "code_const1 (Suc c) = quad_encode 3 1 1 (singleton_encode (code_const1 c))" + +lemma code_const1: "code_const1 c = encode (r_const c)" + by (induction c) simp_all + +definition "r_code_const1_aux \ + Cn 3 r_prod_encode + [r_constn 2 3, + Cn 3 r_prod_encode + [r_constn 2 1, + Cn 3 r_prod_encode + [r_constn 2 1, Cn 3 r_singleton_encode [Id 3 1]]]]" + +lemma r_code_const1_aux_prim: "prim_recfn 3 r_code_const1_aux" + by (simp_all add: r_code_const1_aux_def) + +lemma r_code_const1_aux: + "eval r_code_const1_aux [i, r, c] \= quad_encode 3 1 1 (singleton_encode r)" + by (simp add: r_code_const1_aux_def) + +definition "r_code_const1 \ r_shrink (Pr 1 Z r_code_const1_aux)" + +lemma r_code_const1_prim: "prim_recfn 1 r_code_const1" + by (simp_all add: r_code_const1_def r_code_const1_aux_prim) + +lemma r_code_const1: "eval r_code_const1 [c] \= code_const1 c" +proof - + let ?h = "Pr 1 Z r_code_const1_aux" + have "eval ?h [c, x] \= code_const1 c" for x + using r_code_const1_aux r_code_const1_def + by (induction c) (simp_all add: r_code_const1_aux_prim) + then show ?thesis by (simp add: r_code_const1_def r_code_const1_aux_prim) +qed + +text \Functions that compute codes of higher-arity constant functions:\ + +definition code_constn :: "nat \ nat \ nat" where + "code_constn n c \ + if n = 1 then code_const1 c + else quad_encode 3 n (code_const1 c) (singleton_encode (triple_encode 2 n 0))" + +lemma code_constn: "code_constn (Suc n) c = encode (r_constn n c)" + unfolding code_constn_def using code_const1 r_constn_def + by (cases "n = 0") simp_all + +definition r_code_constn :: "nat \ recf" where + "r_code_constn n \ + if n = 1 then r_code_const1 + else + Cn 1 r_prod_encode + [r_const 3, + Cn 1 r_prod_encode + [r_const n, + Cn 1 r_prod_encode + [r_code_const1, + Cn 1 r_singleton_encode + [Cn 1 r_prod_encode + [r_const 2, Cn 1 r_prod_encode [r_const n, Z]]]]]]" + +lemma r_code_constn_prim: "prim_recfn 1 (r_code_constn n)" + by (simp_all add: r_code_constn_def r_code_const1_prim) + +lemma r_code_constn: "eval (r_code_constn n) [c] \= code_constn n c" + by (auto simp add: r_code_constn_def r_code_const1 code_constn_def r_code_const1_prim) + +text \Computing codes of $m$-ary projections:\ + +definition code_id :: "nat \ nat \ nat" where + "code_id m n \ triple_encode 2 m n" + +lemma code_id: "encode (Id m n) = code_id m n" + unfolding code_id_def by simp + +text \The functions $s^m_n$ are represented by the following function. +The value $m$ corresponds to the length of @{term "cs"}.\ + +definition smn :: "nat \ nat \ nat list \ nat" where + "smn n p cs \ quad_encode + 3 + n + (encode (r_universal (n + length cs))) + (list_encode (code_constn n p # map (code_constn n) cs @ map (code_id n) [0.. 0" + shows "smn n p cs = encode + (Cn n + (r_universal (n + length cs)) + (r_constn (n - 1) p # map (r_constn (n - 1)) cs @ (map (Id n) [0..The next function is to help us define @{typ recf}s corresponding +to the $s^m_n$ functions. It maps $m + 1$ arguments $p, c_1, \ldots, c_m$ to +an encoded list of length $m + n + 1$. The list comprises the $m + 1$ codes +of the $n$-ary constants $p, c_1, \ldots, c_m$ and the $n$ codes for all +$n$-ary projections.\ + +definition r_smn_aux :: "nat \ nat \ recf" where + "r_smn_aux n m \ + Cn (Suc m) + (r_list_encode (m + n)) + (map (\i. Cn (Suc m) (r_code_constn n) [Id (Suc m) i]) [0..i. r_constn m (code_id n i)) [0.. 0 \ prim_recfn (Suc m) (r_smn_aux n m)" + by (auto simp add: r_smn_aux_def r_code_constn_prim) + +lemma r_smn_aux: + assumes "n > 0" and "length cs = m" + shows "eval (r_smn_aux n m) (p # cs) \= + list_encode (map (code_constn n) (p # cs) @ map (code_id n) [0..g. eval g (p # cs)) ?xs = map Some (map (code_constn n) (p # cs))" + proof (intro nth_equalityI) + show len: "length (map (\g. eval g (p # cs)) ?xs) = + length (map Some (map (code_constn n) (p # cs)))" + by (simp add: assms(2)) + + have "map (\g. eval g (p # cs)) ?xs ! i = map Some (map (code_constn n) (p # cs)) ! i" + if "i < Suc m" for i + proof - + have "map (\g. eval g (p # cs)) ?xs ! i = (\g. eval g (p # cs)) (?xs ! i)" + using len_xs that by (metis nth_map) + also have "... = eval (Cn (Suc m) (r_code_constn n) [Id (Suc m) i]) (p # cs)" + using that len_xs + by (metis (no_types, lifting) add.left_neutral length_map nth_map nth_upt) + also have "... = eval (r_code_constn n) [the (eval (Id (Suc m) i) (p # cs))]" + using r_code_constn_prim assms(2) that by simp + also have "... = eval (r_code_constn n) [(p # cs) ! i]" + using len that by simp + finally have "map (\g. eval g (p # cs)) ?xs ! i \= code_constn n ((p # cs) ! i)" + using r_code_constn by simp + then show ?thesis + using len_xs len that by (metis length_map nth_map) + qed + moreover have "length (map (\g. eval g (p # cs)) ?xs) = Suc m" by simp + ultimately show "\i. i < length (map (\g. eval g (p # cs)) ?xs) \ + map (\g. eval g (p # cs)) ?xs ! i = + map Some (map (code_constn n) (p # cs)) ! i" + by simp + qed + moreover have "map (\g. eval g (p # cs)) ?ys = map Some (map (code_id n) [0..g. eval g (p # cs)) (?xs @ ?ys) = + map Some (map (code_constn n) (p # cs) @ map (code_id n) [0..x. the (eval x (p # cs))) (?xs @ ?ys) = + map the (map (\x. eval x (p # cs)) (?xs @ ?ys))" + by simp + ultimately have *: "map (\g. the (eval g (p # cs))) (?xs @ ?ys) = + (map (code_constn n) (p # cs) @ map (code_id n) [0..ig. eval g (p # cs)) ?xs ! i" + by (metis nth_map) + then have + "\ii" + using assms map_xs by (metis length_map nth_map option.simps(3)) + then have xs_converg: "\z\set ?xs. eval z (p # cs) \" + by (metis in_set_conv_nth) + + have "\ix. eval x (p # cs)) ?ys ! i" + by simp + then have + "\ii" + by simp + then have "\z\set (?xs @ ?ys). eval z (p # cs) \" + using xs_converg by auto + moreover have "recfn (length (p # cs)) (Cn (Suc m) (r_list_encode (m + n)) (?xs @ ?ys))" + using assms r_code_constn_prim by auto + ultimately have "eval (r_smn_aux n m) (p # cs) = + eval (r_list_encode (m + n)) (map (\g. the (eval g (p # cs))) (?xs @ ?ys))" + unfolding r_smn_aux_def using assms by simp + then have "eval (r_smn_aux n m) (p # cs) = + eval (r_list_encode (m + n)) (map (code_constn n) (p # cs) @ map (code_id n) [0..For all $m, n > 0$, the @{typ recf} corresponding to $s^m_n$ is +given by the next function.\ + +definition r_smn :: "nat \ nat \ recf" where + "r_smn n m \ + Cn (Suc m) r_prod_encode + [r_constn m 3, + Cn (Suc m) r_prod_encode + [r_constn m n, + Cn (Suc m) r_prod_encode + [r_constn m (encode (r_universal (n + m))), r_smn_aux n m]]]" + +lemma r_smn_prim [simp]: "n > 0 \ prim_recfn (Suc m) (r_smn n m)" + by (simp_all add: r_smn_def r_smn_aux_prim) + +lemma r_smn: + assumes "n > 0" and "length cs = m" + shows "eval (r_smn n m) (p # cs) \= smn n p cs" + using assms r_smn_def r_smn_aux smn_def r_smn_aux_prim by simp + +lemma map_eval_Some_the: + assumes "map (\g. eval g xs) gs = map Some ys" + shows "map (\g. the (eval g xs)) gs = ys" + using assms + by (metis (no_types, lifting) length_map nth_equalityI nth_map option.sel) + +text \The essential part of the $s$-$m$-$n$ theorem: For all $m, n > 0$ +the function $s^m_n$ satisfies +\[ + \varphi_p^{(m + n)}(c_1, \dots,c_m, x_1, \dots, x_n) = + \varphi_{s^m_n(p, c_1, \dots,c_m)}^{(n)}(x_1, \dots, x_n) +\] for all $p, c_i, x_j$.\ + +lemma smn_lemma: + assumes "n > 0" and len_cs: "length cs = m" and len_xs: "length xs = n" + shows "eval (r_universal (m + n)) (p # cs @ xs) = + eval (r_universal n) ((the (eval (r_smn n m) (p # cs))) # xs)" +proof - + let ?s = "r_smn n m" + let ?f = "Cn n + (r_universal (n + length cs)) + (r_constn (n - 1) p # map (r_constn (n - 1)) cs @ (map (Id n) [0..= smn n p cs" + using assms r_smn by simp + then have eval_s: "eval ?s (p # cs) \= encode ?f" + by (simp add: assms(1) smn) + + have "recfn n ?f" + using len_cs assms by auto + then have *: "eval (r_universal n) ((encode ?f) # xs) = eval ?f xs" + using r_universal[of ?f n, OF _ len_xs] by simp + + let ?gs = "r_constn (n - 1) p # map (r_constn (n - 1)) cs @ map (Id n) [0..g\set ?gs. eval g xs \" + using len_cs len_xs assms by auto + then have "eval ?f xs = + eval (r_universal (n + length cs)) (map (\g. the (eval g xs)) ?gs)" + using len_cs len_xs assms `recfn n ?f` by simp + then have "eval ?f xs = eval (r_universal (m + n)) (map (\g. the (eval g xs)) ?gs)" + by (simp add: len_cs add.commute) + then have "eval (r_universal n) ((the (eval ?s (p # cs))) # xs) = + eval (r_universal (m + n)) (map (\g. the (eval g xs)) ?gs)" + using eval_s * by simp + moreover have "map (\g. the (eval g xs)) ?gs = p # cs @ xs" + proof (intro nth_equalityI) + show "length (map (\g. the (eval g xs)) ?gs) = length (p # cs @ xs)" + by (simp add: len_xs) + have len: "length (map (\g. the (eval g xs)) ?gs) = Suc (m + n)" + by (simp add: len_cs) + moreover have "map (\g. the (eval g xs)) ?gs ! i = (p # cs @ xs) ! i" + if "i < Suc (m + n)" for i + proof - + from that consider "i = 0" | "i > 0 \ i < Suc m" | "Suc m \ i \ i < Suc (m + n)" + using not_le_imp_less by auto + then show ?thesis + proof (cases) + case 1 + then show ?thesis using assms(1) len_xs by simp + next + case 2 + then have "?gs ! i = (map (r_constn (n - 1)) cs) ! (i - 1)" + using len_cs + by (metis One_nat_def Suc_less_eq Suc_pred length_map + less_numeral_extra(3) nth_Cons' nth_append) + then have "map (\g. the (eval g xs)) ?gs ! i = + (\g. the (eval g xs)) ((map (r_constn (n - 1)) cs) ! (i - 1))" + using len by (metis length_map nth_map that) + also have "... = the (eval ((r_constn (n - 1) (cs ! (i - 1)))) xs)" + using 2 len_cs by auto + also have "... = cs ! (i - 1)" + using r_constn len_xs assms(1) by simp + also have "... = (p # cs @ xs) ! i" + using 2 len_cs + by (metis diff_Suc_1 less_Suc_eq_0_disj less_numeral_extra(3) nth_Cons' nth_append) + finally show ?thesis . + next + case 3 + then have "?gs ! i = (map (Id n) [0..g. the (eval g xs)) ?gs ! i = + (\g. the (eval g xs)) ((map (Id n) [0..g. the (eval g xs)) ?gs ! i = (p # cs @ xs) ! i" + if "i < length (map (\g. the (eval g xs)) ?gs)" for i + using that by simp + qed + ultimately show ?thesis by simp +qed + +theorem smn_theorem: + assumes "n > 0" + shows "\s. prim_recfn (Suc m) s \ + (\p cs xs. length cs = m \ length xs = n \ + eval (r_universal (m + n)) (p # cs @ xs) = + eval (r_universal n) ((the (eval s (p # cs))) # xs))" + using smn_lemma exI[of _ "r_smn n m"] assms by simp + +text \For every numbering, that is, binary partial recursive function, +$\psi$ there is a total recursive function $c$ that translates $\psi$-indices +into $\varphi$-indices.\ + +lemma numbering_translation: + assumes "recfn 2 psi" + obtains c where + "recfn 1 c" + "total c" + "\i x. eval psi [i, x] = eval r_phi [the (eval c [i]), x]" +proof - + let ?p = "encode psi" + define c where "c = Cn 1 (r_smn 1 1) [r_const ?p, Id 1 0]" + then have "prim_recfn 1 c" by simp + moreover from this have "total c" + by auto + moreover have "eval r_phi [the (eval c [i]), x] = eval psi [i, x]" for i x + proof - + have "eval c [i] = eval (r_smn 1 1) [?p, i]" + using c_def by simp + then have "eval (r_universal 1) [the (eval c [i]), x] = + eval (r_universal 1) [the (eval (r_smn 1 1) [?p, i]), x]" + by simp + also have "... = eval (r_universal (1 + 1)) (?p # [i] @ [x])" + using smn_lemma[of 1 "[i]" 1 "[x]" ?p] by simp + also have "... = eval (r_universal 2) [?p, i, x]" + by (metis append_eq_Cons_conv nat_1_add_1) + also have "... = eval psi [i, x]" + using r_universal[OF assms, of "[i, x]"] by simp + finally have "eval (r_universal 1) [the (eval c [i]), x] = eval psi [i, x]" . + then show ?thesis using r_phi_def by simp + qed + ultimately show ?thesis using that by auto +qed + + +section \Fixed-point theorems\ + +text \Fixed-point theorems (also known as recursion theorems) come in +many shapes. We prove the minimum we need for Chapter~\ref{c:iirf}.\ + + +subsection \Rogers's fixed-point theorem\ + +text \In this section we prove a theorem that Rogers~\cite{Rogers87} +credits to Kleene, but admits that it is a special case and not the original +formulation. We follow Wikipedia~\cite{wiki-krt} and call it the Rogers's +fixed-point theorem.\ + +lemma s11_inj: "inj (\x. smn 1 p [x])" +proof + fix x\<^sub>1 x\<^sub>2 :: nat + assume "smn 1 p [x\<^sub>1] = smn 1 p [x\<^sub>2]" + then have "list_encode [code_constn 1 p, code_constn 1 x\<^sub>1, code_id 1 0] = + list_encode [code_constn 1 p, code_constn 1 x\<^sub>2, code_id 1 0]" + using smn_def by (simp add: prod_encode_eq) + then have "[code_constn 1 p, code_constn 1 x\<^sub>1, code_id 1 0] = + [code_constn 1 p, code_constn 1 x\<^sub>2, code_id 1 0]" + using list_decode_encode by metis + then have "code_constn 1 x\<^sub>1 = code_constn 1 x\<^sub>2" by simp + then show "x\<^sub>1 = x\<^sub>2" + using code_const1 code_constn code_constn_def encode_injective r_constn + by (metis One_nat_def length_Cons list.size(3) option.simps(1)) +qed + +definition "r_univuniv \ Cn 2 r_phi [Cn 2 r_phi [Id 2 0, Id 2 0], Id 2 1]" + +lemma r_univuniv_recfn: "recfn 2 r_univuniv" + by (simp add: r_univuniv_def) + +lemma r_univuniv_converg: + assumes "eval r_phi [x, x] \" + shows "eval r_univuniv [x, y] = eval r_phi [the (eval r_phi [x, x]), y]" + unfolding r_univuniv_def using assms r_univuniv_recfn r_phi_recfn by simp + +text \Strictly speaking this is a generalization of Rogers's theorem in +that it shows the existence of infinitely many fixed-points. In conventional +terms it says that for every total recursive $f$ and $k \in \mathbb{N}$ there is +an $n \geq k$ with $\varphi_n = \varphi_{f(n)}$.\ + +theorem rogers_fixed_point_theorem: + fixes k :: nat + assumes "recfn 1 f" and "total f" + shows "\n\k. \x. eval r_phi [n, x] = eval r_phi [the (eval f [n]), x]" +proof - + let ?p = "encode r_univuniv" + define h where "h = Cn 1 (r_smn 1 1) [r_const ?p, Id 1 0]" + then have "prim_recfn 1 h" + by simp + then have "total h" + by blast + have "eval h [x] = eval (Cn 1 (r_smn 1 1) [r_const ?p, Id 1 0]) [x]" for x + unfolding h_def by simp + then have h: "the (eval h [x]) = smn 1 ?p [x]" for x + by (simp add: r_smn) + + have "eval r_phi [the (eval h [x]), y] = eval r_univuniv [x, y]" for x y + proof - + have "eval r_phi [the (eval h [x]), y] = eval r_phi [smn 1 ?p [x], y]" + using h by simp + also have "... = eval r_phi [the (eval (r_smn 1 1) [?p, x]), y]" + by (simp add: r_smn) + also have "... = eval (r_universal 2) [?p, x, y]" + using r_phi_def smn_lemma[of 1 "[x]" 1 "[y]" ?p] + by (metis Cons_eq_append_conv One_nat_def Suc_1 length_Cons + less_numeral_extra(1) list.size(3) plus_1_eq_Suc) + finally show "eval r_phi [the (eval h [x]), y] = eval r_univuniv [x, y]" + using r_universal r_univuniv_recfn by simp + qed + then have *: "eval r_phi [the (eval h [x]), y] = eval r_phi [the (eval r_phi [x, x]), y]" + if "eval r_phi [x, x] \" for x y + using r_univuniv_converg that by simp + + let ?fh = "Cn 1 f [h]" + have "recfn 1 ?fh" + using `prim_recfn 1 h` assms by simp + then have "infinite {r. recfn 1 r \ r \ ?fh}" + using exteq_infinite[of ?fh 1] by simp + then have "infinite (encode ` {r. recfn 1 r \ r \ ?fh})" (is "infinite ?E") + using encode_injective by (meson finite_imageD inj_onI) + then have "infinite ((\x. smn 1 ?p [x]) ` ?E)" + using s11_inj[of ?p] by (simp add: finite_image_iff inj_on_subset) + moreover have "(\x. smn 1 ?p [x]) ` ?E = {smn 1 ?p [encode r] |r. recfn 1 r \ r \ ?fh}" + by auto + ultimately have "infinite {smn 1 ?p [encode r] |r. recfn 1 r \ r \ ?fh}" + by simp + then obtain n where "n \ k" "n \ {smn 1 ?p [encode r] |r. recfn 1 r \ r \ ?fh}" + by (meson finite_nat_set_iff_bounded_le le_cases) + then obtain r where r: "recfn 1 r" "n = smn 1 ?p [encode r]" "recfn 1 r \ r \ ?fh" + by auto + then have eval_r: "eval r [encode r] = eval ?fh [encode r]" + by (simp add: exteq_def) + then have eval_r': "eval r [encode r] = eval f [the (eval h [encode r])]" + using assms `total h` `prim_recfn 1 h` by simp + then have "eval r [encode r] \" + using `prim_recfn 1 h` assms(1,2) by simp + then have "eval r_phi [encode r, encode r] \" + by (simp add: \recfn 1 r\ r_phi) + then have "eval r_phi [the (eval h [encode r]), y] = + eval r_phi [(the (eval r_phi [encode r, encode r])), y]" + for y + using * by simp + then have "eval r_phi [the (eval h [encode r]), y] = + eval r_phi [(the (eval r [encode r])), y]" + for y + by (simp add: \recfn 1 r\ r_phi) + moreover have "n = the (eval h [encode r])" by (simp add: h r(2)) + ultimately have "eval r_phi [n, y] = eval r_phi [the (eval r [encode r]), y]" for y + by simp + then have "eval r_phi [n, y] = eval r_phi [the (eval ?fh [encode r]), y]" for y + using r by (simp add: eval_r) + moreover have "eval ?fh [encode r] = eval f [n]" + using eval_r eval_r' \n = the (eval h [encode r])\ by auto + ultimately have "eval r_phi [n, y] = eval r_phi [the (eval f [n]), y]" for y + by simp + with `n \ k` show ?thesis by auto +qed + + +subsection \Kleene's fixed-point theorem\ + +text \The next theorem is what Rogers~\cite[p.~214]{Rogers87} calls +Kleene's version of what we call Rogers's fixed-point theorem. More precisely +this would be Kleene's \emph{second} fixed-point theorem, but since we do not +cover the first one, we leave out the number.\ + +theorem kleene_fixed_point_theorem: + fixes k :: nat + assumes "recfn 2 psi" + shows "\n\k. \x. eval r_phi [n, x] = eval psi [n, x]" +proof - + from numbering_translation[OF assms] obtain c where c: + "recfn 1 c" + "total c" + "\i x. eval psi [i, x] = eval r_phi [the (eval c [i]), x]" + by auto + then obtain n where "n \ k" and "\x. eval r_phi [n, x] = eval r_phi [the (eval c [n]), x]" + using rogers_fixed_point_theorem by blast + with c(3) have "\x. eval r_phi [n, x] = eval psi [n, x]" + by simp + with `n \ k` show ?thesis by auto +qed + +text \Kleene's fixed-point theorem can be generalized to arbitrary +arities. But we need to generalize it only to binary functions in order to +show Smullyan's double fixed-point theorem in +Section~\ref{s:smullyan}.\ + +definition "r_univuniv2 \ + Cn 3 r_phi [Cn 3 (r_universal 2) [Id 3 0, Id 3 0, Id 3 1], Id 3 2]" + +lemma r_univuniv2_recfn: "recfn 3 r_univuniv2" + by (simp add: r_univuniv2_def) + +lemma r_univuniv2_converg: + assumes "eval (r_universal 2) [u, u, x] \" + shows "eval r_univuniv2 [u, x, y] = eval r_phi [the (eval (r_universal 2) [u, u, x]), y]" + unfolding r_univuniv2_def using assms r_univuniv2_recfn by simp + +theorem kleene_fixed_point_theorem_2: + assumes "recfn 2 f" and "total f" + shows "\n. + recfn 1 n \ + total n \ + (\x y. eval r_phi [(the (eval n [x])), y] = eval r_phi [(the (eval f [the (eval n [x]), x])), y])" +proof - + let ?p = "encode r_univuniv2" + let ?s = "r_smn 1 2" + define h where "h = Cn 2 ?s [r_dummy 1 (r_const ?p), Id 2 0, Id 2 1]" + then have [simp]: "prim_recfn 2 h" by simp + { + fix u x y + have "eval h [u, x] = eval (Cn 2 ?s [r_dummy 1 (r_const ?p), Id 2 0, Id 2 1]) [u, x]" + using h_def by simp + then have "the (eval h [u, x]) = smn 1 ?p [u, x]" + by (simp add: r_smn) + then have "eval r_phi [the (eval h [u, x]), y] = eval r_phi [smn 1 ?p [u, x], y]" + by simp + also have "... = + eval r_phi + [encode (Cn 1 (r_universal 3) (r_constn 0 ?p # r_constn 0 u # r_constn 0 x # [Id 1 0])), + y]" + using smn[of 1 ?p "[u, x]"] by (simp add: numeral_3_eq_3) + also have "... = + eval r_phi + [encode (Cn 1 (r_universal 3) (r_const ?p # r_const u # r_const x # [Id 1 0])), y]" + (is "_ = eval r_phi [encode ?f, y]") + by (simp add: r_constn_def) + also have "... = eval ?f [y]" + using r_phi'[of ?f] by auto + also have "... = eval (r_universal 3) [?p, u, x, y]" + using r_univuniv2_recfn r_universal r_phi by auto + also have "... = eval r_univuniv2 [u, x, y]" + using r_universal by (simp add: r_univuniv2_recfn) + finally have "eval r_phi [the (eval h [u, x]), y] = eval r_univuniv2 [u, x, y]" . + } + then have *: "eval r_phi [the (eval h [u, x]), y] = + eval r_phi [the (eval (r_universal 2) [u, u, x]), y]" + if "eval (r_universal 2) [u, u, x] \" for u x y + using r_univuniv2_converg that by simp + + let ?fh = "Cn 2 f [h, Id 2 1]" + let ?e = "encode ?fh" + have "recfn 2 ?fh" + using assms by simp + have "total h" + by auto + then have "total ?fh" + using assms Cn_total totalI2[of ?fh] by fastforce + + let ?n = "Cn 1 h [r_const ?e, Id 1 0]" + have "recfn 1 ?n" + using assms by simp + moreover have "total ?n" + using `total h` totalI1[of ?n] by simp + moreover { + fix x y + have "eval r_phi [(the (eval ?n [x])), y] = eval r_phi [(the (eval h [?e, x])), y]" + by simp + also have "... = eval r_phi [the (eval (r_universal 2) [?e, ?e, x]), y]" + using * r_universal[of _ 2] totalE[of ?fh 2] \total ?fh\ \recfn 2 ?fh\ + by (metis length_Cons list.size(3) numeral_2_eq_2) + also have "... = eval r_phi [the (eval f [the (eval h [?e, x]), x]), y]" + proof - + have "eval (r_universal 2) [?e, ?e, x] \" + using totalE[OF `total ?fh`] `recfn 2 ?fh` r_universal + by (metis length_Cons list.size(3) numeral_2_eq_2) + moreover have "eval (r_universal 2) [?e, ?e, x] = eval ?fh [?e, x]" + by (metis \recfn 2 ?fh\ length_Cons list.size(3) numeral_2_eq_2 r_universal) + then show ?thesis using assms `total h` by simp + qed + also have "... = eval r_phi [(the (eval f [the (eval ?n [x]), x])), y]" + by simp + finally have "eval r_phi [(the (eval ?n [x])), y] = + eval r_phi [(the (eval f [the (eval ?n [x]), x])), y]" . + } + ultimately show ?thesis by blast +qed + + +subsection \Smullyan's double fixed-point theorem\label{s:smullyan}\ + +theorem smullyan_double_fixed_point_theorem: + assumes "recfn 2 g" and "total g" and "recfn 2 h" and "total h" + shows "\m n. + (\x. eval r_phi [m, x] = eval r_phi [the (eval g [m, n]), x]) \ + (\x. eval r_phi [n, x] = eval r_phi [the (eval h [m, n]), x])" +proof - + obtain m where + "recfn 1 m" and + "total m" and + m: "\x y. eval r_phi [the (eval m [x]), y] = + eval r_phi [the (eval g [the (eval m [x]), x]), y]" + using kleene_fixed_point_theorem_2[of g] assms(1,2) by auto + define k where "k = Cn 1 h [m, Id 1 0]" + then have "recfn 1 k" + using `recfn 1 m` assms(3) by simp + have "total (Id 1 0)" + by (simp add: Mn_free_imp_total) + then have "total k" + using `total m` assms(4) Cn_total k_def `recfn 1 k` by simp + obtain n where n: "\x. eval r_phi [n, x] = eval r_phi [the (eval k [n]), x]" + using rogers_fixed_point_theorem[of k] `recfn 1 k` `total k` by blast + obtain mm where mm: "eval m [n] \= mm" + using `total m` `recfn 1 m` by fastforce + then have "\x. eval r_phi [mm, x] = eval r_phi [the (eval g [mm, n]), x]" + by (metis m option.sel) + moreover have "\x. eval r_phi [n, x] = eval r_phi [the (eval h [mm, n]), x]" + using k_def assms(3) `total m` `recfn 1 m` mm n by simp + ultimately show ?thesis by blast +qed + + +section \Decidable and recursively enumerable sets\label{s:decidable}\ + +text \We defined @{term decidable} already back in +Section~\ref{s:halting}: @{thm[display] decidable_def}\ + +text \The next theorem is adapted from @{thm[source] +halting_problem_undecidable}.\ + +theorem halting_problem_phi_undecidable: "\ decidable {x. eval r_phi [x, x] \}" + (is "\ decidable ?K") +proof + assume "decidable ?K" + then obtain f where "recfn 1 f" and f: "\x. eval f [x] \= (if x \ ?K then 1 else 0)" + using decidable_def by auto + define g where "g \ Cn 1 r_ifeq_else_diverg [f, Z, Z]" + then have "recfn 1 g" + using `recfn 1 f` r_ifeq_else_diverg_recfn by simp + then obtain i where i: "eval r_phi [i, x] = eval g [x]" for x + using r_phi' by auto + from g_def have "eval g [x] = (if x \ ?K then Some 0 else None)" for x + using r_ifeq_else_diverg_recfn `recfn 1 f` f by simp + then have "eval g [i] \ \ i \ ?K" by simp + also have "... \ eval r_phi [i, i] \" by simp + also have "... \ eval g [i] \" + using i by simp + finally have "eval g [i] \ \ eval g [i] \" . + then show False by auto +qed + +lemma decidable_complement: "decidable X \ decidable (- X)" +proof - + assume "decidable X" + then obtain f where f: "recfn 1 f" "\x. eval f [x] \= (if x \ X then 1 else 0)" + using decidable_def by auto + define g where "g = Cn 1 r_not [f]" + then have "recfn 1 g" + by (simp add: f(1)) + moreover have "eval g [x] \= (if x \ X then 0 else 1)" for x + by (simp add: g_def f) + ultimately show ?thesis using decidable_def by auto +qed + +text \Finite sets are decidable.\ + +fun r_contains :: "nat list \ recf" where + "r_contains [] = Z" +| "r_contains (x # xs) = Cn 1 r_ifeq [Id 1 0, r_const x, r_const 1, r_contains xs]" + +lemma r_contains_prim: "prim_recfn 1 (r_contains xs)" + by (induction xs) auto + +lemma r_contains: "eval (r_contains xs) [x] \= (if x \ set xs then 1 else 0)" +proof (induction xs arbitrary: x) + case Nil + then show ?case by simp +next + case (Cons a xs) + have "eval (r_contains (a # xs)) [x] = eval r_ifeq [x, a, 1, the (eval (r_contains xs) [x])]" + using r_contains_prim prim_recfn_total by simp + also have "... \= (if x = a then 1 else if x \ set xs then 1 else 0)" + using Cons.IH by simp + also have "... \= (if x = a \ x \ set xs then 1 else 0)" + by simp + finally show ?case by simp +qed + +lemma finite_set_decidable: "finite X \ decidable X" +proof - + fix X :: "nat set" + assume "finite X" + then obtain xs where "X = set xs" + using finite_list by auto + then have "\x. eval (r_contains xs) [x] \= (if x \ X then 1 else 0)" + using r_contains by simp + then show "decidable X" + using decidable_def r_contains_prim by blast +qed + +definition semidecidable :: "nat set \ bool" where + "semidecidable X \ (\f. recfn 1 f \ (\x. eval f [x] = (if x \ X then Some 1 else None)))" + +text \The semidecidable sets are the domains of partial recursive functions.\ + +lemma semidecidable_iff_domain: + "semidecidable X \ (\f. recfn 1 f \ (\x. eval f [x] \ \ x \ X))" +proof + show "semidecidable X \ \f. recfn 1 f \ (\x. (eval f [x] \) = (x \ X))" + using semidecidable_def by (metis option.distinct(1)) + show "semidecidable X" if "\f. recfn 1 f \ (\x. (eval f [x] \) = (x \ X))" for X + proof - + from that obtain f where f: "recfn 1 f" "\x. (eval f [x] \) = (x \ X)" + by auto + let ?g = "Cn 1 (r_const 1) [f]" + have "recfn 1 ?g" + using f(1) by simp + moreover have "\x. eval ?g [x] = (if x \ X then Some 1 else None)" + using f by simp + ultimately show "semidecidable X" + using semidecidable_def by blast + qed +qed + +lemma decidable_imp_semidecidable: "decidable X \ semidecidable X" +proof - + assume "decidable X" + then obtain f where f: "recfn 1 f" "\x. eval f [x] \= (if x \ X then 1 else 0)" + using decidable_def by auto + define g where "g = Cn 1 r_ifeq_else_diverg [f, r_const 1, r_const 1]" + then have "recfn 1 g" + by (simp add: f(1)) + have "eval g [x] = eval r_ifeq_else_diverg [if x \ X then 1 else 0, 1, 1]" for x + by (simp add: g_def f) + then have "\x. x \ X \ eval g [x] \= 1" and "\x. x \ X \ eval g [x] \" + by simp_all + then show ?thesis + using `recfn 1 g` semidecidable_def by auto +qed + +text \A set is recursively enumerable if it is empty or the image of a +total recursive function.\ + +definition recursively_enumerable :: "nat set \ bool" where + "recursively_enumerable X \ + X = {} \ (\f. recfn 1 f \ total f \ X = {the (eval f [x]) |x. x \ UNIV})" + +theorem recursively_enumerable_iff_semidecidable: + "recursively_enumerable X \ semidecidable X" +proof + show "semidecidable X" if "recursively_enumerable X" for X + proof (cases) + assume "X = {}" + then show ?thesis + using finite_set_decidable decidable_imp_semidecidable + recursively_enumerable_def semidecidable_def + by blast + next + assume "X \ {}" + with that obtain f where f: "recfn 1 f" "total f" "X = {the (eval f [x]) |x. x \ UNIV}" + using recursively_enumerable_def by blast + define h where "h = Cn 2 r_eq [Cn 2 f [Id 2 0], Id 2 1]" + then have "recfn 2 h" + using f(1) by simp + from h_def have h: "eval h [x, y] \= 0 \ the (eval f [x]) = y" for x y + using f(1,2) by simp + from h_def `recfn 2 h` totalI2 f(2) have "total h" by simp + define g where "g = Mn 1 h" + then have "recfn 1 g" + using h_def f(1) by simp + then have "eval g [y] = + (if (\x. eval h [x, y] \= 0 \ (\x')) + then Some (LEAST x. eval h [x, y] \= 0) + else None)" for y + using g_def `total h` f(2) by simp + then have "eval g [y] = + (if \x. eval h [x, y] \= 0 + then Some (LEAST x. eval h [x, y] \= 0) + else None)" for y + using `total h` `recfn 2 h` by simp + then have "eval g [y] \ \ (\x. eval h [x, y] \= 0)" for y + by simp + with h have "eval g [y] \ \ (\x. the (eval f [x]) = y)" for y + by simp + with f(3) have "eval g [y] \ \ y \ X" for y + by auto + with `recfn 1 g` semidecidable_iff_domain show ?thesis by auto + qed + + show "recursively_enumerable X" if "semidecidable X" for X + proof (cases) + assume "X = {}" + then show ?thesis using recursively_enumerable_def by simp + next + assume "X \ {}" + then obtain x\<^sub>0 where "x\<^sub>0 \ X" by auto + from that semidecidable_iff_domain obtain f where f: "recfn 1 f" "\x. eval f [x] \ \ x \ X" + by auto + let ?i = "encode f" + have i: "\x. eval f [x] = eval r_phi [?i, x]" + using r_phi' f(1) by simp + with `x\<^sub>0 \ X` f(2) have "eval r_phi [?i, x\<^sub>0] \" by simp + then obtain g where g: "recfn 1 g" "total g" "\x. eval r_phi [?i, x] \ = (\y. eval g [y] \= x)" + using f(1) nonempty_domain_enumerable by blast + with f(2) i have "\x. x \ X = (\y. eval g [y] \= x)" + by simp + then have "\x. x \ X = (\y. the (eval g [y]) = x)" + using totalE[OF g(2) g(1)] + by (metis One_nat_def length_Cons list.size(3) option.collapse option.sel) + then have "X = {the (eval g [y]) |y. y \ UNIV}" + by auto + with g(1,2) show ?thesis using recursively_enumerable_def by auto + qed +qed + +text \The next goal is to show that a set is decidable iff. it and its +complement are semidecidable. For this we use the concurrent evaluation +function.\ + +lemma semidecidable_decidable: + assumes "semidecidable X" and "semidecidable (- X)" + shows "decidable X" +proof - + obtain f where f: "recfn 1 f \ (\x. eval f [x] \ \ x \ X)" + using assms(1) semidecidable_iff_domain by auto + let ?i = "encode f" + obtain g where g: "recfn 1 g \ (\x. eval g [x] \ \ x \ (- X))" + using assms(2) semidecidable_iff_domain by auto + let ?j = "encode g" + define d where "d = Cn 1 r_pdec1 [Cn 1 r_parallel [r_const ?j, r_const ?i, Id 1 0]]" + then have "recfn 1 d" + by (simp add: d_def) + have *: "\x. eval r_phi [?i, x] = eval f [x]" "\x. eval r_phi [?j, x] = eval g [x]" + using f g r_phi' by simp_all + have "eval d [x] \= 1" if "x \ X" for x + proof - + have "eval f [x] \" + using f that by simp + moreover have "eval g [x] \" + using g that by blast + ultimately have "eval r_parallel [?j, ?i, x] \= prod_encode (1, the (eval f [x]))" + using * r_parallel(3) by simp + with d_def show ?thesis by simp + qed + moreover have "eval d [x] \= 0" if "x \ X" for x + proof - + have "eval g [x] \" + using g that by simp + moreover have "eval f [x] \" + using f that by blast + ultimately have "eval r_parallel [?j, ?i, x] \= prod_encode (0, the (eval g [x]))" + using * r_parallel(2) by blast + with d_def show ?thesis by simp + qed + ultimately show ?thesis + using decidable_def `recfn 1 d` by auto +qed + +theorem decidable_iff_semidecidable_complement: + "decidable X \ semidecidable X \ semidecidable (- X)" + using semidecidable_decidable decidable_imp_semidecidable decidable_complement + by blast + + +section \Rice's theorem\ + +definition index_set :: "nat set \ bool" where + "index_set I \ \i j. i \ I \ (\x. eval r_phi [i, x] = eval r_phi [j, x]) \ j \ I" + +lemma index_set_closed_in: + assumes "index_set I" and "i \ I" and "\x. eval r_phi [i, x] = eval r_phi [j, x]" + shows "j \ I" + using index_set_def assms by simp + +lemma index_set_closed_not_in: + assumes "index_set I" and "i \ I" and "\x. eval r_phi [i, x] = eval r_phi [j, x]" + shows "j \ I" + using index_set_def assms by metis + +theorem rice_theorem: + assumes "index_set I" and "I \ UNIV" and "I \ {}" + shows "\ decidable I" +proof + assume "decidable I" + then obtain d where d: "recfn 1 d" "\i. eval d [i] \= (if i \ I then 1 else 0)" + using decidable_def by auto + obtain j\<^sub>1 j\<^sub>2 where "j\<^sub>1 \ I" and "j\<^sub>2 \ I" + using assms(2,3) by auto + let ?if = "Cn 2 r_ifz [Cn 2 d [Id 2 0], r_dummy 1 (r_const j\<^sub>2), r_dummy 1 (r_const j\<^sub>1)]" + define psi where "psi = Cn 2 r_phi [?if, Id 2 1] " + then have "recfn 2 psi" + by (simp add: d) + have "eval ?if [x, y] = Some (if x \ I then j\<^sub>1 else j\<^sub>2)" for x y + by (simp add: d) + moreover have "eval psi [x, y] = eval (Cn 2 r_phi [?if, Id 2 1]) [x, y]" for x y + using psi_def by simp + ultimately have psi: "eval psi [x, y] = eval r_phi [if x \ I then j\<^sub>1 else j\<^sub>2, y]" for x y + by (simp add: d) + then have in_I: "eval psi [x, y] = eval r_phi [j\<^sub>1, y]" if "x \ I" for x y + by (simp add: that) + have not_in_I: "eval psi [x, y] = eval r_phi [j\<^sub>2, y]" if "x \ I" for x y + by (simp add: psi that) + obtain n where n: "\x. eval r_phi [n, x] = eval psi [n, x]" + using kleene_fixed_point_theorem[OF `recfn 2 psi`] by auto + show False + proof cases + assume "n \ I" + then have "\x. eval r_phi [n, x] = eval r_phi [j\<^sub>1, x]" + using n in_I by simp + then have "n \ I" + using `j\<^sub>1 \ I` index_set_closed_not_in[OF assms(1)] by simp + with `n \ I` show False by simp + next + assume "n \ I" + then have "\x. eval r_phi [n, x] = eval r_phi [j\<^sub>2, x]" + using n not_in_I by simp + then have "n \ I" + using `j\<^sub>2 \ I` index_set_closed_in[OF assms(1)] by simp + with `n \ I` show False by simp + qed +qed + + +section \Partial recursive functions as actual functions\label{s:alternative}\ + +text \A well-formed @{typ recf} describes an algorithm. Usually, +however, partial recursive functions are considered to be partial functions, +that is, right-unique binary relations. This distinction did not matter much +until now, because we were mostly concerned with the \emph{existence} of +partial recursive functions, which is equivalent to the existence of +algorithms. Whenever it did matter, we could use the extensional equivalence +@{term "exteq"}. In Chapter~\ref{c:iirf}, however, we will deal with sets of +functions and sets of sets of functions. + +For illustration consider the singleton set containing only the unary zero +function. It could be expressed by @{term "{Z}"}, but this would not contain +@{term[names_short] "Cn 1 (Id 1 0) [Z]"}, which computes the same function. +The alternative representation as @{term "{f. f \ Z}"} is not a +singleton set. Another alternative would be to identify partial recursive +functions with the equivalence classes of @{term "exteq"}. This would work +for all arities. But since we will only need unary and binary functions, we +can go for the less general but simpler alternative of regarding partial +recursive functions as certain functions of types @{typ "nat \ +nat option"} and @{typ "nat \ nat \ nat option"}. +With this notation we can represent the aforementioned set by @{term +"{\_. Some (0::nat)}"} and express that the function @{term "\_. +Some (0::nat)"} is total recursive. + +In addition terms get shorter, for instance, @{term "eval r_func [i, x]"} +becomes @{term "func i x"}.\ + + +subsection \The definitions\ + +type_synonym partial1 = "nat \ nat option" + +type_synonym partial2 = "nat \ nat \ nat option" + +definition total1 :: "partial1 \ bool" where + "total1 f \ \x. f x \" + +definition total2 :: "partial2 \ bool" where + "total2 f \ \x y. f x y \" + +lemma total1I [intro]: "(\x. f x \) \ total1 f" + using total1_def by simp + +lemma total2I [intro]: "(\x y. f x y \) \ total2 f" + using total2_def by simp + +lemma total1E [dest, simp]: "total1 f \ f x \" + using total1_def by simp + +lemma total2E [dest, simp]: "total2 f \ f x y \" + using total2_def by simp + +definition P1 :: "partial1 set" ("\

") where + "\

\ {\x. eval r [x] |r. recfn 1 r}" + +definition P2 :: "partial2 set" ("\

\<^sup>2") where + "\

\<^sup>2 \ {\x y. eval r [x, y] |r. recfn 2 r}" + +definition R1 :: "partial1 set" ("\") where + "\ \ {\x. eval r [x] |r. recfn 1 r \ total r}" + +definition R2 :: "partial2 set" ("\\<^sup>2") where + "\\<^sup>2 \ {\x y. eval r [x, y] |r. recfn 2 r \ total r}" + +definition Prim1 :: "partial1 set" where + "Prim1 \ {\x. eval r [x] |r. prim_recfn 1 r}" + +definition Prim2 :: "partial2 set" where + "Prim2 \ {\x y. eval r [x, y] |r. prim_recfn 2 r}" + +lemma R1_imp_P1 [simp, elim]: "f \ \ \ f \ \

" + using R1_def P1_def by auto + +lemma R2_imp_P2 [simp, elim]: "f \ \\<^sup>2 \ f \ \

\<^sup>2" + using R2_def P2_def by auto + +lemma Prim1_imp_R1 [simp, elim]: "f \ Prim1 \ f \ \" + unfolding Prim1_def R1_def by auto + +lemma Prim2_imp_R2 [simp, elim]: "f \ Prim2 \ f \ \\<^sup>2" + unfolding Prim2_def R2_def by auto + +lemma P1E [elim]: + assumes "f \ \

" + obtains r where "recfn 1 r" and "\x. eval r [x] = f x" + using assms P1_def by force + +lemma P2E [elim]: + assumes "f \ \

\<^sup>2" + obtains r where "recfn 2 r" and "\x y. eval r [x, y] = f x y" + using assms P2_def by force + +lemma P1I [intro]: + assumes "recfn 1 r" and "(\x. eval r [x]) = f" + shows "f \ \

" + using assms P1_def by auto + +lemma P2I [intro]: + assumes "recfn 2 r" and "\x y. eval r [x, y] = f x y" + shows "f \ \

\<^sup>2" +proof - + have "(\x y. eval r [x, y]) = f" + using assms(2) by simp + then show ?thesis + using assms(1) P2_def by auto +qed + +lemma R1I [intro]: + assumes "recfn 1 r" and "total r" and "\x. eval r [x] = f x" + shows "f \ \" + unfolding R1_def + using CollectI[of "\f. \r. f = (\x. eval r [x]) \ recfn 1 r \ total r" f] assms + by metis + +lemma R1E [elim]: + assumes "f \ \" + obtains r where "recfn 1 r" and "total r" and "f = (\x. eval r [x])" + using assms R1_def by auto + +lemma R2I [intro]: + assumes "recfn 2 r" and "total r" and "\x y. eval r [x, y] = f x y" + shows "f \ \\<^sup>2" + unfolding R2_def + using CollectI[of "\f. \r. f = (\x y. eval r [x, y]) \ recfn 2 r \ total r" f] assms + by metis + +lemma R1_SOME: + assumes "f \ \" + and "r = (SOME r'. recfn 1 r' \ total r' \ f = (\x. eval r' [x]))" + (is "r = (SOME r'. ?P r')") + shows "recfn 1 r" + and "\x. eval r [x] \" + and "\x. f x = eval r [x]" + and "f = (\x. eval r [x])" +proof - + obtain r' where "?P r'" + using R1E[OF assms(1)] by auto + then show "recfn 1 r" "\b. eval r [b] \" "\x. f x = eval r [x]" + using someI[of ?P r'] assms(2) totalE[of r] by (auto, metis) + then show "f = (\x. eval r [x])" by auto +qed + +lemma R2E [elim]: + assumes "f \ \\<^sup>2" + obtains r where "recfn 2 r" and "total r" and "f = (\x\<^sub>1 x\<^sub>2. eval r [x\<^sub>1, x\<^sub>2])" + using assms R2_def by auto + +lemma R1_imp_total1 [simp]: "f \ \ \ total1 f" + using total1I by fastforce + +lemma R2_imp_total2 [simp]: "f \ \\<^sup>2 \ total2 f" + using totalE by fastforce + +lemma Prim1I [intro]: + assumes "prim_recfn 1 r" and "\x. f x = eval r [x]" + shows "f \ Prim1" + using assms Prim1_def by blast + +lemma Prim2I [intro]: + assumes "prim_recfn 2 r" and "\x y. f x y = eval r [x, y]" + shows "f \ Prim2" + using assms Prim2_def by blast + +lemma P1_total_imp_R1 [intro]: + assumes "f \ \

" and "total1 f" + shows "f \ \" + using assms totalI1 by force + +lemma P2_total_imp_R2 [intro]: + assumes "f \ \

\<^sup>2 " and "total2 f" + shows "f \ \\<^sup>2" + using assms totalI2 by force + + +subsection \Some simple properties\ + +text \In order to show that a @{typ partial1} or @{typ partial2} +function is in @{term "\

"}, @{term "\

\<^sup>2"}, @{term "\"}, @{term +"\\<^sup>2"}, @{term "Prim1"}, or @{term "Prim2"} we will usually have to +find a suitable @{typ recf}. But for some simple or frequent cases this +section provides shortcuts.\ + +lemma identity_in_R1: "Some \ \" +proof - + have "\x. eval (Id 1 0) [x] \= x" by simp + moreover have "recfn 1 (Id 1 0)" by simp + moreover have "total (Id 1 0)" + by (simp add: totalI1) + ultimately show ?thesis by blast +qed + +lemma P2_proj_P1 [simp, elim]: + assumes "\ \ \

\<^sup>2" + shows "\ i \ \

" +proof - + from assms obtain u where u: "recfn 2 u" "(\x\<^sub>1 x\<^sub>2. eval u [x\<^sub>1, x\<^sub>2]) = \" + by auto + define v where "v \ Cn 1 u [r_const i, Id 1 0]" + then have "recfn 1 v" "(\x. eval v [x]) = \ i" + using u by auto + then show ?thesis by auto +qed + +lemma R2_proj_R1 [simp, elim]: + assumes "\ \ \\<^sup>2" + shows "\ i \ \" +proof - + from assms have "\ \ \

\<^sup>2" by simp + then have "\ i \ \

" by auto + moreover have "total1 (\ i)" + using assms by (simp add: total1I) + ultimately show ?thesis by auto +qed + +lemma const_in_Prim1: "(\_. Some c) \ Prim1" +proof - + define r where "r = r_const c" + then have "\x. eval r [x] = Some c" by simp + moreover have "recfn 1 r" "Mn_free r" + using r_def by simp_all + ultimately show ?thesis by auto +qed + +lemma concat_P1_P1: + assumes "f \ \

" and "g \ \

" + shows "(\x. if g x \ \ f (the (g x)) \ then Some (the (f (the (g x)))) else None) \ \

" + (is "?h \ \

") +proof - + obtain rf where rf: "recfn 1 rf" "\x. eval rf [x] = f x" + using assms(1) by auto + obtain rg where rg: "recfn 1 rg" "\x. eval rg [x] = g x" + using assms(2) by auto + let ?rh = "Cn 1 rf [rg]" + have "recfn 1 ?rh" + using rf(1) rg(1) by simp + moreover have "eval ?rh [x] = ?h x" for x + using rf rg by simp + ultimately show ?thesis by blast +qed + +lemma P1_update_P1: + assumes "f \ \

" + shows "f(x:=z) \ \

" +proof (cases z) + case None + define re where "re \ Mn 1 (r_constn 1 1)" + from assms obtain r where r: "recfn 1 r" "(\u. eval r [u]) = f" + by auto + define r' where "r' = Cn 1 (r_lifz re r) [Cn 1 r_eq [Id 1 0, r_const x], Id 1 0]" + have "recfn 1 r'" + using r(1) r'_def re_def by simp + then have "eval r' [u] = eval (r_lifz re r) [if u = x then 0 else 1, u]" for u + using r'_def by simp + with r(1) have "eval r' [u] = (if u = x then None else eval r [u])" for u + using re_def re_def by simp + with r(2) have "eval r' [u] = (f(x:=None)) u" for u + by auto + then have "(\u. eval r' [u]) = f(x:=None)" + by auto + with None `recfn 1 r'` show ?thesis by auto +next + case (Some y) + from assms obtain r where r: "recfn 1 r" "(\u. eval r [u]) = f" + by auto + define r' where + "r' \ Cn 1 (r_lifz (r_const y) r) [Cn 1 r_eq [Id 1 0, r_const x], Id 1 0]" + have "recfn 1 r'" + using r(1) r'_def by simp + then have "eval r' [u] = eval (r_lifz (r_const y) r) [if u = x then 0 else 1, u]" for u + using r'_def by simp + with r(1) have "eval r' [u] = (if u = x then Some y else eval r [u])" for u + by simp + with r(2) have "eval r' [u] = (f(x:=Some y)) u" for u + by auto + then have "(\u. eval r' [u]) = f(x:=Some y)" + by auto + with Some `recfn 1 r'` show ?thesis by auto +qed + +lemma swap_P2: + assumes "f \ \

\<^sup>2" + shows "(\x y. f y x) \ \

\<^sup>2" +proof - + obtain r where r: "recfn 2 r" "\x y. eval r [x, y] = f x y" + using assms by auto + then have "eval (r_swap r) [x, y] = f y x" for x y + by simp + moreover have "recfn 2 (r_swap r)" + using r_swap_recfn r(1) by simp + ultimately show ?thesis by auto +qed + +lemma swap_R2: + assumes "f \ \\<^sup>2" + shows "(\x y. f y x) \ \\<^sup>2" + using swap_P2[of f] assms + by (meson P2_total_imp_R2 R2_imp_P2 R2_imp_total2 total2E total2I) + +lemma skip_P1: + assumes "f \ \

" + shows "(\x. f (x + n)) \ \

" +proof - + obtain r where r: "recfn 1 r" "\x. eval r [x] = f x" + using assms by auto + let ?s = "Cn 1 r [Cn 1 r_add [Id 1 0, r_const n]]" + have "recfn 1 ?s" + using r by simp + have "eval ?s [x] = eval r [x + n]" for x + using r by simp + with r have "eval ?s [x] = f (x + n)" for x + by simp + with `recfn 1 ?s` show ?thesis by blast +qed + +lemma skip_R1: + assumes "f \ \" + shows "(\x. f (x + n)) \ \" + using assms skip_P1 R1_imp_total1 total1_def by auto + + +subsection \The Gödel numbering @{term \}\label{s:goedel_numbering}\ + +text \While the term \emph{Gödel numbering} is often used generically for +mappings between natural numbers and mathematical concepts, the inductive +inference literature uses it in a more specific sense. There it is equivalent +to the notion of acceptable numbering~\cite{Rogers87}: For every numbering +there is a recursive function mapping the numbering's indices to equivalent +ones of a Gödel numbering.\ + +definition goedel_numbering :: "partial2 \ bool" where + "goedel_numbering \ \ \ \ \

\<^sup>2 \ (\\\\

\<^sup>2. \c\\. \i. \ i = \ (the (c i)))" + +lemma goedel_numbering_P2: + assumes "goedel_numbering \" + shows "\ \ \

\<^sup>2" + using goedel_numbering_def assms by simp + +lemma goedel_numberingE: + assumes "goedel_numbering \" and "\ \ \

\<^sup>2" + obtains c where "c \ \" and "\i. \ i = \ (the (c i))" + using assms goedel_numbering_def by blast + +lemma goedel_numbering_universal: + assumes "goedel_numbering \" and "f \ \

" + shows "\i. \ i = f" +proof - + define \ :: partial2 where "\ = (\i. f)" + have "\ \ \

\<^sup>2" + proof - + obtain rf where rf: "recfn 1 rf" "\x. eval rf [x] = f x" + using assms(2) by auto + define r where "r = Cn 2 rf [Id 2 1]" + then have r: "recfn 2 r" "\i x. eval r [i, x] = eval rf [x]" + using rf(1) by simp_all + with rf(2) have "\i x. eval r [i, x] = f x" by simp + with r(1) show ?thesis using \_def by auto + qed + then obtain c where "c \ \" and "\i. \ i = \ (the (c i))" + using goedel_numbering_def assms(1) by auto + with \_def show ?thesis by auto +qed + +text \Our standard Gödel numbering is based on @{term r_phi}:\ + +definition phi :: partial2 ("\") where + "\ i x \ eval r_phi [i, x]" + +lemma phi_in_P2: "\ \ \

\<^sup>2" + unfolding phi_def using r_phi_recfn by blast + +text \Indices of any numbering can be translated into equivalent indices +of @{term phi}, which thus is a Gödel numbering.\ + +lemma numbering_translation_for_phi: + assumes "\ \ \

\<^sup>2" + shows "\c\\. \i. \ i = \ (the (c i))" +proof - + obtain psi where psi: "recfn 2 psi" "\i x. eval psi [i, x] = \ i x" + using assms by auto + with numbering_translation obtain b where + "recfn 1 b" "total b" "\i x. eval psi [i, x] = eval r_phi [the (eval b [i]), x]" + by blast + moreover from this obtain c where c: "c \ \" "\i. c i = eval b [i]" + by fast + ultimately have "\ i x = \ (the (c i)) x" for i x + using phi_def psi(2) by presburger + then have "\ i = \ (the (c i))" for i + by auto + then show ?thesis using c(1) by blast +qed + +corollary goedel_numbering_phi: "goedel_numbering \" + unfolding goedel_numbering_def using numbering_translation_for_phi phi_in_P2 by simp + +corollary phi_universal: + assumes "f \ \

" + obtains i where "\ i = f" + using goedel_numbering_universal[OF goedel_numbering_phi assms] by auto + + +subsection \Fixed-point theorems\ + +text \The fixed-point theorems look somewhat cleaner in the new +notation. We will only need the following ones in the next chapter.\ + +theorem kleene_fixed_point: + fixes k :: nat + assumes "\ \ \

\<^sup>2" + obtains i where "i \ k" and "\ i = \ i" +proof - + obtain r_psi where r_psi: "recfn 2 r_psi" "\i x. eval r_psi [i, x] = \ i x" + using assms by auto + then obtain i where i: "i \ k" "\x. eval r_phi [i, x] = eval r_psi [i, x]" + using kleene_fixed_point_theorem by blast + then have "\x. \ i x = \ i x" + using phi_def r_psi by simp + then show ?thesis using i that by blast +qed + +theorem smullyan_double_fixed_point: + assumes "g \ \\<^sup>2" and "h \ \\<^sup>2" + obtains m n where "\ m = \ (the (g m n))" and "\ n = \ (the (h m n))" +proof - + obtain rg where rg: "recfn 2 rg" "total rg" "g = (\x y. eval rg [x, y])" + using R2E[OF assms(1)] by auto + moreover obtain rh where rh: "recfn 2 rh" "total rh" "h = (\x y. eval rh [x, y])" + using R2E[OF assms(2)] by auto + ultimately obtain m n where + "\x. eval r_phi [m, x] = eval r_phi [the (eval rg [m, n]), x]" + "\x. eval r_phi [n, x] = eval r_phi [the (eval rh [m, n]), x]" + using smullyan_double_fixed_point_theorem[of rg rh] by blast + then have "\ m = \ (the (g m n))" and "\ n = \ (the (h m n))" + using phi_def rg rh by auto + then show ?thesis using that by simp +qed + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/TOTAL_CONS.thy b/thys/Inductive_Inference/TOTAL_CONS.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/TOTAL_CONS.thy @@ -0,0 +1,1468 @@ +section \TOTAL is a proper subset of CONS\label{s:total_cons}\ + +theory TOTAL_CONS + imports Lemma_R (* for r_auxhyp *) + CP_FIN_NUM (* for r_consistent *) + CONS_LIM (* for rmge2, goedel_at *) +begin + +text \We first show that TOTAL is a subset of CONS. Then we present a +separating class.\ + + +subsection \TOTAL is a subset of CONS\ + +text \A TOTAL strategy hypothesizes only total functions, for which the +consistency with the input prefix is decidable. A CONS strategy can thus run +a TOTAL strategy and check if its hypothesis is consistent. If so, it +outputs this hypothesis, otherwise some arbitrary consistent one. Since the +TOTAL strategy converges to a correct hypothesis, which is consistent, the +CONS strategy will converge to the same hypothesis.\ + +text \Without loss of generality we can assume that learning takes place +with respect to our Gödel numbering $\varphi$. So we need to decide +consistency only for this numbering.\ + +abbreviation r_consist_phi where + "r_consist_phi \ r_consistent r_phi" + +lemma r_consist_phi_recfn [simp]: "recfn 2 r_consist_phi" + by simp + +lemma r_consist_phi: + assumes "\k i k \" + shows "eval r_consist_phi [i, e] \= + (if \k i k \= e_nth e k then 0 else 1)" +proof - + have "\k" + using assms phi_def by simp + moreover have "recfn 2 r_phi" by simp + ultimately have "eval (r_consistent r_phi) [i, e] \= + (if \k= e_nth e k then 0 else 1)" + using r_consistent_converg assms by simp + then show ?thesis using phi_def by simp +qed + +lemma r_consist_phi_init: + assumes "f \ \" and "\ i \ \" + shows "eval r_consist_phi [i, f \ n] \= (if \k\n. \ i k = f k then 0 else 1)" + using assms r_consist_phi R1_imp_total1 total1E by (simp add: r_consist_phi) + +lemma TOTAL_subseteq_CONS: "TOTAL \ CONS" +proof + fix U assume "U \ TOTAL" + then have "U \ TOTAL_wrt \" + using TOTAL_wrt_phi_eq_TOTAL by blast + then obtain t' where t': "learn_total \ U t'" + using TOTAL_wrt_def by auto + then obtain t where t: "recfn 1 t" "\x. eval t [x] = t' x" + using learn_totalE(1) P1E by blast + then have t_converg: "eval t [f \ n] \" if "f \ U" for f n + using t' learn_totalE(1) that by auto + + define s where "s \ Cn 1 r_ifz [Cn 1 r_consist_phi [t, Id 1 0], t, r_auxhyp]" + then have "recfn 1 s" + using r_consist_phi_recfn r_auxhyp_prim t(1) by simp + + have consist: "eval r_consist_phi [the (eval t [f \ n]), f \ n] \= + (if \k\n. \ (the (eval t [f \ n])) k = f k then 0 else 1)" + if "f \ U" for f n + proof - + have "eval r_consist_phi [the (eval t [f \ n]), f \ n] = + eval (Cn 1 r_consist_phi [t, Id 1 0]) [f \ n]" + using that t_converg t(1) by simp + also have "... \= (if \k\n. \ (the (eval t [f \ n])) k = f k then 0 else 1)" + proof - + from that have "f \ \" + using learn_totalE(1) t' by blast + moreover have "\ (the (eval t [f \ n])) \ \" + using t' t learn_totalE t_converg that by simp + ultimately show ?thesis + using r_consist_phi_init t_converg t(1) that by simp + qed + finally show ?thesis . + qed + + have s_eq_t: "eval s [f \ n] = eval t [f \ n]" + if "\k\n. \ (the (eval t [f \ n])) k = f k" and "f \ U" for f n + using that consist s_def t r_auxhyp_prim prim_recfn_total + by simp + + have s_eq_aux: "eval s [f \ n] = eval r_auxhyp [f \ n]" + if "\ (\k\n. \ (the (eval t [f \ n])) k = f k)" and "f \ U" for f n + proof - + from that have "eval r_consist_phi [the (eval t [f \ n]), f \ n] \= 1" + using consist by simp + moreover have "t' (f \ n) \" using t' learn_totalE(1) that(2) by blast + ultimately show ?thesis + using s_def t r_auxhyp_prim t' learn_totalE by simp + qed + + have "learn_cons \ U (\e. eval s [e])" + proof (rule learn_consI) + have "eval s [f \ n] \" if "f \ U" for f n + using that t_converg[OF that, of n] s_eq_t[of n f] prim_recfn_total[of r_auxhyp 1] + r_auxhyp_prim s_eq_aux[OF _ that, of n] totalE + by fastforce + then show "environment \ U (\e. eval s [e])" + using t' `recfn 1 s` learn_totalE(1) by blast + show "\i. \ i = f \ (\\<^sup>\n. eval s [f \ n] \= i)" if "f \ U" for f + proof - + from that t' t learn_totalE obtain i n\<^sub>0 where + i_n0: "\ i = f \ (\n\n\<^sub>0. eval t [f \ n] \= i)" + by metis + then have "\n. n \ n\<^sub>0 \ \k\n. \ (the (eval t [f \ n])) k = f k" + by simp + with s_eq_t have "\n. n \ n\<^sub>0 \ eval s [f \ n] = eval t [f \ n]" + using that by simp + with i_n0 have "\n. n \ n\<^sub>0 \ eval s [f \ n] \= i" + by auto + with i_n0 show ?thesis by auto + qed + show "\k\n. \ (the (eval s [f \ n])) k = f k" if "f \ U" for f n + proof (cases "\k\n. \ (the (eval t [f \ n])) k = f k") + case True + with that s_eq_t show ?thesis by simp + next + case False + then have "eval s [f \ n] = eval r_auxhyp [f \ n]" + using that s_eq_aux by simp + moreover have "f \ \" + using learn_totalE(1)[OF t'] that by auto + ultimately show ?thesis using r_auxhyp by simp + qed + qed + then show "U \ CONS" using CONS_def by auto +qed + + +subsection \The separating class\ + + +subsubsection \Definition of the class\ + +text \The class that will be shown to be in @{term "CONS - TOTAL"} is +the union of the following two classes.\ + +definition V_constotal_1 :: "partial1 set" where + "V_constotal_1 \ {f. \j p. f = [j] \ p \ j \ 2 \ p \ \\<^sub>0\<^sub>1 \ \ j = f}" + +definition V_constotal_2 :: "partial1 set" where + "V_constotal_2 \ + {f. \j a k. + f = j # a @ [k] \ 0\<^sup>\ \ + j \ 2 \ + (\i 1) \ + k \ 2 \ + \ j = j # a \ \\<^sup>\ \ + \ k = f}" + +definition V_constotal :: "partial1 set" where + "V_constotal \ V_constotal_1 \ V_constotal_2" + +lemma V_constotal_2I: + assumes "f = j # a @ [k] \ 0\<^sup>\" + and "j \ 2" + and "\i 1" + and "k \ 2" + and "\ j = j # a \ \\<^sup>\" + and "\ k = f" + shows "f \ V_constotal_2" + using assms V_constotal_2_def by blast + +lemma V_subseteq_R1: "V_constotal \ \" +proof + fix f assume "f \ V_constotal" + then have "f \ V_constotal_1 \ f \ V_constotal_2" + using V_constotal_def by auto + then show "f \ \" + proof + assume "f \ V_constotal_1" + then obtain j p where "f = [j] \ p" "p \ \\<^sub>0\<^sub>1" + using V_constotal_1_def by blast + then show ?thesis using prepend_in_R1 RPred1_subseteq_R1 by auto + next + assume "f \ V_constotal_2" + then obtain j a k where "f = j # a @ [k] \ 0\<^sup>\" + using V_constotal_2_def by blast + then show ?thesis using almost0_in_R1 by auto + qed +qed + + +subsubsection \The class is in CONS\ + +text \The class can be learned by the strategy @{term rmge2}, which +outputs the rightmost value greater or equal two in the input $f^n$. If $f$ +is from $V_1$ then the strategy is correct right from the start. If $f$ is +from $V_2$ the strategy outputs the consistent hypothesis $j$ until it +encounters the correct hypothesis $k$, to which it converges.\ + +lemma V_in_CONS: "learn_cons \ V_constotal rmge2" +proof (rule learn_consI) + show "environment \ V_constotal rmge2" + using V_subseteq_R1 rmge2_in_R1 R1_imp_total1 phi_in_P2 by simp + have "(\i. \ i = f \ (\\<^sup>\n. rmge2 (f \ n) \= i)) \ + (\n. \k\n. \ (the (rmge2 (f \ n))) k = f k)" + if "f \ V_constotal" for f + proof (cases "f \ V_constotal_1") + case True + then obtain j p where + f: "f = [j] \ p" and + j: "j \ 2" and + p: "p \ \\<^sub>0\<^sub>1" and + phi_j: "\ j = f" + using V_constotal_1_def by blast + then have "f 0 \= j" by (simp add: prepend_at_less) + then have f_at_0: "the (f 0) \ 2" by (simp add: j) + have f_at_gr0: "the (f x) \ 1" if "x > 0" for x + using that f p by (simp add: RPred1_altdef Suc_leI prepend_at_ge) + have "total1 f" + using V_subseteq_R1 that R1_imp_total1 total1_def by auto + have "rmge2 (f \ n) \= j" for n + proof - + let ?P = "\i. i < Suc n \ the (f i) \ 2" + have "Greatest ?P = 0" + proof (rule Greatest_equality) + show "0 < Suc n \ 2 \ the (f 0)" + using f_at_0 by simp + show "\y. y < Suc n \ 2 \ the (f y) \ y \ 0" + using f_at_gr0 by fastforce + qed + then have "rmge2 (f \ n) = f 0" + using f_at_0 rmge2_init_total[of f n, OF `total1 f`] by auto + then show "rmge2 (f \ n) \= j" + by (simp add: \f 0 \= j\) + qed + then show ?thesis using phi_j by auto + next + case False + then have "f \ V_constotal_2" + using V_constotal_def that by auto + then obtain j a k where jak: + "f = j # a @ [k] \ 0\<^sup>\" + "j \ 2" + "\i 1" + "k \ 2" + "\ j = j # a \ \\<^sup>\ " + "\ k = f" + using V_constotal_2_def by blast + then have f_at_0: "f 0 \= j" by simp + have f_eq_a: "f x \= a ! (x - 1)" if "0 < x \ x < Suc (length a)" for x + proof - + have "x - 1 < length a" + using that by auto + then show ?thesis + by (simp add: jak(1) less_SucI nth_append that) + qed + then have f_at_a: "the (f x) \ 1" if "0 < x \ x < Suc (length a)" for x + using jak(3) that by auto + from jak have f_k: "f (Suc (length a)) \= k" by auto + from jak have f_at_big: "f x \= 0" if "x > Suc (length a)" for x + using that by simp + let ?P = "\n i. i < Suc n \ the (f i) \ 2" + have rmge2: "rmge2 (f \ n) = f (Greatest (?P n))" for n + proof - + have "\ (\i 2 \ the (f 0)" + using that by (simp add: jak(2) f_at_0) + show "\y. y < Suc n \ 2 \ the (f y) \ y \ 0" + using that f_at_a + by (metis Suc_1 dual_order.strict_trans leI less_Suc_eq not_less_eq_eq) + qed + with rmge2 f_at_0 have rmge2_small: + "rmge2 (f \ n) \= j" if "n < Suc (length a)" for n + using that by simp + have "Greatest (?P n) = Suc (length a)" if "n \ Suc (length a)" for n + proof (rule Greatest_equality) + show "Suc (length a) < Suc n \ 2 \ the (f (Suc (length a)))" + using that f_k by (simp add: jak(4) less_Suc_eq_le) + show "\y. y < Suc n \ 2 \ the (f y) \ y \ Suc (length a)" + using that f_at_big by (metis leI le_SucI not_less_eq_eq numeral_2_eq_2 option.sel) + qed + with rmge2 f_at_big f_k have rmge2_big: + "rmge2 (f \ n) \= k" if "n \ Suc (length a)" for n + using that by simp + then have "\i n\<^sub>0. \ i = f \ (\n\n\<^sub>0. rmge2 (f \ n) \= i)" + using jak(6) by auto + moreover have "\k\n. \ (the (rmge2 (f \ n))) k = f k" for n + proof (cases "n < Suc (length a)") + case True + then have "rmge2 (f \ n) \= j" + using rmge2_small by simp + then have "\ (the (rmge2 (f \ n))) = \ j" by simp + with True show ?thesis + using rmge2_small f_at_0 f_eq_a jak(5) prepend_at_less + by (metis le_less_trans le_zero_eq length_Cons not_le_imp_less nth_Cons_0 nth_Cons_pos) + next + case False + then show ?thesis using rmge2_big jak by simp + qed + ultimately show ?thesis by simp + qed + then show "\f. f \ V_constotal \ \i. \ i = f \ (\\<^sup>\n. rmge2 (f \ n) \= i)" + and "\f n. f \ V_constotal \ \k\n. \ (the (rmge2 (f \ n))) k = f k" + by simp_all +qed + + +subsubsection \The class is not in TOTAL\ + +text \Recall that $V$ is the union of $V_1 = \{jp \mid j\geq2 \land p \in +\mathcal{R}_{01} \land \varphi_j = jp\}$ and $V_2 = \{jak0^\infty \mid j\geq 2 \land a +\in \{0, 1\}^* \land k\geq 2 \land \varphi_j = ja\uparrow^\infty \land\ +\varphi_k = jak0^\infty\}$.\ + +text \The proof is adapted from a proof of a stronger result by +Freivalds, Kinber, and Wiehagen~\cite[Theorem~27]{fkw-iisde-95} concerning an +inference type not defined here. + +The proof is by contradiction. If $V$ was in TOTAL, there would be +a strategy $S$ learning $V$ in our standard Gödel numbering $\varphi$. +By Lemma R for TOTAL we can assume $S$ to be total. + +In order to construct a function $f\in V$ for which $S$ fails we employ a +computable process iteratively building function prefixes. For every $j$ the +process builds a function $\psi_j$. The initial prefix is the singleton +$[j]$. Given a prefix $b$, the next prefix is determined as follows: +\begin{enumerate} +\item Search for a $y \geq |b|$ with $\varphi_{S(b)}(y) \downarrow= v$ for +some $v$. +\item Set the new prefix $b0^{y - |b|}\bar{v}$, where $\bar v = 1 - v$. +\end{enumerate} + +Step~1 can diverge, for example, if $\varphi_{S(b)}$ is the empty function. +In this case $\psi_j$ will only be defined for a finite prefix. If, however, +Step~2 is reached, the prefix $b$ is extended to a $b'$ such that +$\varphi_{S(b)}(y) \neq b'_y$, which implies $S(b)$ is a wrong hypothesis for +every function starting with $b'$, in particular for $\psi_j$. Since $\bar v +\in \{0, 1\}$, Step~2 only appends zeros and ones, which is important for +showing membership in $V$. + +This process defines a numbering $\psi \in \mathcal{P}^2$, and by Kleene's +fixed-point theorem there is a $j \geq 2$ with $\varphi_j = \psi_j$. For this +$j$ there are two cases: +\begin{enumerate} +\item[Case 1.] Step~1 always succeeds. Then $\psi_j$ is total and + $\psi_j \in V_1$. But $S$ outputs wrong hypotheses on infinitely many + prefixes of $\psi_j$ (namely every prefix constructed by the process). + +\item[Case 2.] Step~1 diverges at some iteration, say when the state is $b = ja$ + for some $a \in \{0, 1\}^*$. + Then $\psi_j$ has the form $ja\uparrow^\infty$. The numbering $\chi$ with $\chi_k = + jak0^\infty$ is in $\mathcal{P}^2$, and by Kleene's fixed-point theorem there is a + $k\geq 2$ with $\varphi_k = \chi_k = jak0^\infty$. This $jak0^\infty$ is in + $V_2$ and has the prefix $ja$. But Step~1 diverged on this prefix, which + means there is no $y \geq |ja|$ with $\varphi_{S(ja)}(y)\downarrow$. In + other words $S$ hypothesizes a non-total function. +\end{enumerate} + +Thus, in both cases there is a function in $V$ where $S$ does not behave like +a TOTAL strategy. This is the desired contradiction. + +The following locale formalizes this proof sketch.\ + +locale total_cons = + fixes s :: partial1 + assumes s_in_R1: "s \ \" +begin + +definition r_s :: recf where + "r_s \ SOME r_s. recfn 1 r_s \ total r_s \ s = (\x. eval r_s [x])" + +lemma rs_recfn [simp]: "recfn 1 r_s" + and rs_total [simp]: "\x. eval r_s [x] \" + and eval_rs: "\x. s x = eval r_s [x]" + using r_s_def R1_SOME[OF s_in_R1, of r_s] by simp_all + +text \Performing Step~1 means enumerating the domain of +$\varphi_{S(b)}$ until a $y \geq |b|$ is found. The next function enumerates +all domain values and checks the condition for them.\ + +definition "r_search_enum \ + Cn 2 r_le [Cn 2 r_length [Id 2 1], Cn 2 r_enumdom [Cn 2 r_s [Id 2 1], Id 2 0]]" + +lemma r_search_enum_recfn [simp]: "recfn 2 r_search_enum" + by (simp add: r_search_enum_def Let_def) + +abbreviation search_enum :: partial2 where + "search_enum x b \ eval r_search_enum [x, b]" + +abbreviation enumdom :: partial2 where + "enumdom i y \ eval r_enumdom [i, y]" + +lemma enumdom_empty_domain: + assumes "\x. \ i x \" + shows "\y. enumdom i y \" + using assms r_enumdom_empty_domain by (simp add: phi_def) + +lemma enumdom_nonempty_domain: + assumes "\ i x\<^sub>0 \" + shows "\y. enumdom i y \" + and "\x. \ i x \ \ (\y. enumdom i y \= x)" + using assms r_enumdom_nonempty_domain phi_def by metis+ + +text \Enumerating the empty domain yields the empty function.\ + +lemma search_enum_empty: + fixes b :: nat + assumes "s b \= i" and "\x. \ i x \" + shows "\x. search_enum x b \" + using assms r_search_enum_def enumdom_empty_domain eval_rs by simp + +text \Enumerating a non-empty domain yields a total function.\ + +lemma search_enum_nonempty: + fixes b y0 :: nat + assumes "s b \= i" and "\ i y\<^sub>0 \" and "e = the (enumdom i x)" + shows "search_enum x b \= (if e_length b \ e then 0 else 1)" +proof - + let ?e = "\x. the (enumdom i x)" + let ?y = "Cn 2 r_enumdom [Cn 2 r_s [Id 2 1], Id 2 0]" + have "recfn 2 ?y" using assms(1) by simp + moreover have "\x. eval ?y [x, b] = enumdom i x" + using assms(1,2) eval_rs by auto + moreover from this have "\x. eval ?y [x, b] \" + using enumdom_nonempty_domain(1)[OF assms(2)] by simp + ultimately have "eval (Cn 2 r_le [Cn 2 r_length [Id 2 1], ?y]) [x, b] \= + (if e_length b \ ?e x then 0 else 1)" + by simp + then show ?thesis using assms by (simp add: r_search_enum_def) +qed + +text \If there is a $y$ as desired, the enumeration will eventually return +zero (representing ``true'').\ + +lemma search_enum_nonempty_eq0: + fixes b y :: nat + assumes "s b \= i" and "\ i y \" and "y \ e_length b" + shows "\x. search_enum x b \= 0" +proof - + obtain x where x: "enumdom i x \= y" + using enumdom_nonempty_domain(2)[OF assms(2)] assms(2) by auto + from assms(2) have "\ i y \" by simp + with x have "search_enum x b \= 0" + using search_enum_nonempty[where ?e=y] assms by auto + then show ?thesis by auto +qed + +text \If there is no $y$ as desired, the enumeration will never return +zero.\ + +lemma search_enum_nonempty_neq0: + fixes b y0 :: nat + assumes "s b \= i" + and "\ i y\<^sub>0 \" + and "\ (\y. \ i y \ \ y \ e_length b)" + shows "\ (\x. search_enum x b \= 0)" +proof + assume "\x. search_enum x b \= 0" + then obtain x where x: "search_enum x b \= 0" + by auto + obtain y where y: "enumdom i x \= y" + using enumdom_nonempty_domain[OF assms(2)] by blast + then have "search_enum x b \= (if e_length b \ y then 0 else 1)" + using assms(1-2) search_enum_nonempty by simp + with x have "e_length b \ y" + using option.inject by fastforce + moreover have "\ i y \" + using assms(2) enumdom_nonempty_domain(2) y by blast + ultimately show False using assms(3) by force +qed + +text \The next function corresponds to Step~1. Given a prefix $b$ it +computes a $y \geq |b|$ with $\varphi_{S(b)}(y)\downarrow$ if such a $y$ +exists; otherwise it diverges.\ + +definition "r_search \ Cn 1 r_enumdom [r_s, Mn 1 r_search_enum]" + +lemma r_search_recfn [simp]: "recfn 1 r_search" + using r_search_def by simp + +abbreviation search :: partial1 where + "search b \ eval r_search [b]" + +text \If $\varphi_{S(b)}$ is the empty function, the search process +diverges because already the enumeration of the domain diverges.\ + +lemma search_empty: + assumes "s b \= i" and "\x. \ i x \" + shows "search b \" +proof - + have "\x. search_enum x b \" + using search_enum_empty[OF assms] by simp + then have "eval (Mn 1 r_search_enum) [b] \" by simp + then show "search b \" unfolding r_search_def by simp +qed + +text \If $\varphi_{S(b)}$ is non-empty, but there is no $y$ with the +desired properties, the search process diverges.\ + +lemma search_nonempty_neq0: + fixes b y0 :: nat + assumes "s b \= i" + and "\ i y\<^sub>0 \" + and "\ (\y. \ i y \ \ y \ e_length b)" + shows "search b \" +proof - + have "\ (\x. search_enum x b \= 0)" + using assms search_enum_nonempty_neq0 by simp + moreover have "recfn 1 (Mn 1 r_search_enum)" + by (simp add: assms(1)) + ultimately have "eval (Mn 1 r_search_enum) [b] \" by simp + then show ?thesis using r_search_def by auto +qed + +text \If there is a $y$ as desired, the search process will return +one such $y$.\ + +lemma search_nonempty_eq0: + fixes b y :: nat + assumes "s b \= i" and "\ i y \" and "y \ e_length b" + shows "search b \" + and "\ i (the (search b)) \" + and "the (search b) \ e_length b" +proof - + have "\x. search_enum x b \= 0" + using assms search_enum_nonempty_eq0 by simp + moreover have "\x. search_enum x b \" + using assms search_enum_nonempty by simp + moreover have "recfn 1 (Mn 1 r_search_enum)" + by simp + ultimately have + 1: "search_enum (the (eval (Mn 1 r_search_enum) [b])) b \= 0" and + 2: "eval (Mn 1 r_search_enum) [b] \" + using eval_Mn_diverg eval_Mn_convergE[of 1 "r_search_enum" "[b]"] + by (metis (no_types, lifting) One_nat_def length_Cons list.size(3) option.collapse, + metis (no_types, lifting) One_nat_def length_Cons list.size(3)) + let ?x = "the (eval (Mn 1 r_search_enum) [b])" + have "search b = eval (Cn 1 r_enumdom [r_s, Mn 1 r_search_enum]) [b]" + unfolding r_search_def by simp + then have 3: "search b = enumdom i ?x" + using assms 2 eval_rs by simp + then have "the (search b) = the (enumdom i ?x)" (is "?y = _") + by simp + then have 4: "search_enum ?x b \= (if e_length b \ ?y then 0 else 1)" + using search_enum_nonempty assms by simp + from 3 have "\ i ?y \" + using enumdom_nonempty_domain assms(2) by (metis option.collapse) + then show "\ i ?y \" + using phi_def by simp + then show "?y \ e_length b" + using assms 4 1 option.inject by fastforce + show "search b \" + using 3 assms(2) enumdom_nonempty_domain(1) by auto +qed + +text \The converse of the previous lemma states that whenever +the search process returns a value it will be one with the +desired properties.\ + +lemma search_converg: + assumes "s b \= i" and "search b \" (is "?y \") + shows "\ i (the ?y) \" + and "the ?y \ e_length b" +proof - + have "\y. \ i y \" + using assms search_empty by meson + then have "\y. y \ e_length b \ \ i y \" + using search_nonempty_neq0 assms by meson + then obtain y where y: "y \ e_length b \ \ i y \" by auto + then have "\ i y \" + using phi_def by simp + then show "\ i (the (search b)) \" + and "(the (search b)) \ e_length b" + using y assms search_nonempty_eq0[OF assms(1) `\ i y \`] by simp_all +qed + +text \Likewise, if the search diverges, there is no appropriate $y$.\ + +lemma search_diverg: + assumes "s b \= i" and "search b \" + shows "\ (\y. \ i y \ \ y \ e_length b)" +proof + assume "\y. \ i y \ \ y \ e_length b" + then obtain y where y: "\ i y \" "y \ e_length b" + by auto + from y(1) have "\ i y \" + by (simp add: phi_def) + with y(2) search_nonempty_eq0 have "search b \" + using assms by blast + with assms(2) show False by simp +qed + +text \Step~2 extends the prefix by a block of the shape $0^n\bar v$. +The next function constructs such a block for given $n$ and $v$.\ + +definition "r_badblock \ + let f = Cn 1 r_singleton_encode [r_not]; + g = Cn 3 r_cons [r_constn 2 0, Id 3 1] + in Pr 1 f g" + +lemma r_badblock_prim [simp]: "recfn 2 r_badblock" + unfolding r_badblock_def by simp + +lemma r_badblock: "eval r_badblock [n, v] \= list_encode (replicate n 0 @ [1 - v])" +proof (induction n) + case 0 + let ?f = "Cn 1 r_singleton_encode [r_not]" + have "eval r_badblock [0, v] = eval ?f [v]" + unfolding r_badblock_def by simp + also have "... = eval r_singleton_encode [the (eval r_not [v])]" + by simp + also have "... \= list_encode [1 - v]" + by simp + finally show ?case by simp +next + case (Suc n) + let ?g = "Cn 3 r_cons [r_constn 2 0, Id 3 1]" + have "recfn 3 ?g" by simp + have "eval r_badblock [(Suc n), v] = eval ?g [n, the (eval r_badblock [n , v]), v]" + using `recfn 3 ?g` Suc by (simp add: r_badblock_def) + also have "... = eval ?g [n, list_encode (replicate n 0 @ [1 - v]), v]" + using Suc by simp + also have "... = eval r_cons [0, list_encode (replicate n 0 @ [1 - v])]" + by simp + also have "... \= e_cons 0 (list_encode (replicate n 0 @ [1 - v]))" + by simp + also have "... \= list_encode (0 # (replicate n 0 @ [1 - v]))" + by simp + also have "... \= list_encode (replicate (Suc n) 0 @ [1 - v])" + by simp + finally show ?case by simp +qed + +lemma r_badblock_only_01: "e_nth (the (eval r_badblock [n, v])) i \ 1" + using r_badblock by (simp add: nth_append) + +lemma r_badblock_last: "e_nth (the (eval r_badblock [n, v])) n = 1 - v" + using r_badblock by (simp add: nth_append) + +text \The following function computes the next prefix from the current +one. In other words, it performs Steps~1 and~2.\ + +definition "r_next \ + Cn 1 r_append + [Id 1 0, + Cn 1 r_badblock + [Cn 1 r_sub [r_search, r_length], + Cn 1 r_phi [r_s, r_search]]]" + +lemma r_next_recfn [simp]: "recfn 1 r_next" + unfolding r_next_def by simp + +text \The name @{text next} is unavailable, so we go for @{term nxt}.\ + +abbreviation nxt :: partial1 where + "nxt b \ eval r_next [b]" + +lemma nxt_diverg: + assumes "search b \" + shows "nxt b \" + unfolding r_next_def using assms by (simp add: Let_def) + +lemma nxt_converg: + assumes "search b \= y" + shows "nxt b \= + e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\ (the (s b)) y)]))" + unfolding r_next_def using assms r_badblock search_converg phi_def eval_rs + by fastforce + +lemma nxt_search_diverg: + assumes "nxt b \" + shows "search b \" +proof (rule ccontr) + assume "search b \" + then obtain y where "search b \= y" by auto + then show False + using nxt_converg assms by simp +qed + +text \If Step~1 finds a $y$, the hypothesis $S(b)$ is incorrect for +the new prefix.\ + +lemma nxt_wrong_hyp: + assumes "nxt b \= b'" and "s b \= i" + shows "\y i y \\ e_nth b' y" +proof - + obtain y where y: "search b \= y" + using assms nxt_diverg by fastforce + then have y_len: "y \ e_length b" + using assms search_converg(2) by fastforce + then have b': "b' = + (e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\ i y)])))" + using y assms nxt_converg by simp + then have "e_nth b' y = 1 - the (\ i y)" + using y_len e_nth_append_big r_badblock r_badblock_last by auto + moreover have "\ i y \" + using search_converg y y_len assms(2) by fastforce + ultimately have "\ i y \\ e_nth b' y" + by (metis gr_zeroI less_numeral_extra(4) less_one option.sel zero_less_diff) + moreover have "e_length b' = Suc y" + using y_len e_length_append b' by auto + ultimately show ?thesis by auto +qed + +text \If Step~1 diverges, the hypothesis $S(b)$ refers to a non-total +function.\ + +lemma nxt_nontotal_hyp: + assumes "nxt b \" and "s b \= i" + shows "\x. \ i x \" + using nxt_search_diverg[OF assms(1)] search_diverg[OF assms(2)] by auto + +text \The process only ever extends the given prefix.\ + +lemma nxt_stable: + assumes "nxt b \= b'" + shows "\x= y" + using assms nxt_diverg by fastforce + then have "y \ e_length b" + using search_converg(2) eval_rs rs_total by fastforce + show ?thesis + proof (rule allI, rule impI) + fix x assume "x < e_length b" + let ?i = "the (s b)" + have b': "b' = + (e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\ ?i y)])))" + using assms nxt_converg[OF y] by auto + then show "e_nth b x = e_nth b' x" + using e_nth_append_small \x < e_length b\ by auto + qed +qed + +text \The following properties of @{term r_next} will be +used to show that some of the constructed functions are in the class +$V$.\ + +lemma nxt_append_01: + assumes "nxt b \= b'" + shows "\x. x \ e_length b \ x < e_length b' \ e_nth b' x = 0 \ e_nth b' x = 1" +proof - + obtain y where y: "search b \= y" + using assms nxt_diverg by fastforce + let ?i = "the (s b)" + have b': "b' = (e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\ ?i y)])))" + (is "b' = (e_append b ?z)") + using assms y nxt_converg prod_encode_eq by auto + show ?thesis + proof (rule allI, rule impI) + fix x assume x: "e_length b \ x \ x < e_length b'" + then have "e_nth b' x = e_nth ?z (x - e_length b)" + using b' e_nth_append_big by blast + then show "e_nth b' x = 0 \ e_nth b' x = 1" + by (metis less_one nat_less_le option.sel r_badblock r_badblock_only_01) + qed +qed + +lemma nxt_monotone: + assumes "nxt b \= b'" + shows "e_length b < e_length b'" +proof - + obtain y where y: "search b \= y" + using assms nxt_diverg by fastforce + let ?i = "the (s b)" + have b': "b' = + (e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\ ?i y)])))" + using assms y nxt_converg prod_encode_eq by auto + then show ?thesis using e_length_append by auto +qed + +text \The next function computes the prefixes after each iteration of +the process @{term r_next} when started with the list $[j]$.\ + +definition r_prefixes :: recf where + "r_prefixes \ Pr 1 r_singleton_encode (Cn 3 r_next [Id 3 1])" + +lemma r_prefixes_recfn [simp]: "recfn 2 r_prefixes" + unfolding r_prefixes_def by (simp add: Let_def) + +abbreviation prefixes :: partial2 where + "prefixes t j \ eval r_prefixes [t, j]" + +lemma prefixes_at_0: "prefixes 0 j \= list_encode [j]" + unfolding r_prefixes_def by simp + +lemma prefixes_at_Suc: + assumes "prefixes t j \" (is "?b \") + shows "prefixes (Suc t) j = nxt (the ?b)" + using r_prefixes_def assms by auto + +lemma prefixes_at_Suc': + assumes "prefixes t j \= b" + shows "prefixes (Suc t) j = nxt b" + using r_prefixes_def assms by auto + +lemma prefixes_prod_encode: + assumes "prefixes t j \" + obtains b where "prefixes t j \= b" + using assms surj_prod_encode by force + +lemma prefixes_converg_le: + assumes "prefixes t j \" and "t' \ t" + shows "prefixes t' j \" + using r_prefixes_def assms eval_Pr_converg_le[of 1 _ _ "[j]"] + by simp + +lemma prefixes_diverg_add: + assumes "prefixes t j \" + shows "prefixes (t + d) j \" + using r_prefixes_def assms eval_Pr_diverg_add[of 1 _ _ "[j]"] + by simp + +text \Many properties of @{term r_prefixes} can be derived from similar +properties of @{term r_next}.\ + +lemma prefixes_length: + assumes "prefixes t j \= b" + shows "e_length b > t" +proof (insert assms, induction t arbitrary: b) + case 0 + then show ?case using prefixes_at_0 prod_encode_eq by auto +next + case (Suc t) + then have "prefixes t j \" + using prefixes_converg_le Suc_n_not_le_n nat_le_linear by blast + then obtain b' where b': "prefixes t j \= b'" + using prefixes_prod_encode by blast + with Suc have "e_length b' > t" by simp + have "prefixes (Suc t) j = nxt b'" + using b' prefixes_at_Suc' by simp + with Suc have "nxt b' \= b" by simp + then have "e_length b' < e_length b" + using nxt_monotone by simp + then show ?case using `e_length b' > t` by simp +qed + +lemma prefixes_monotone: + assumes "prefixes t j \= b" and "prefixes (t + d) j \= b'" + shows "e_length b \ e_length b'" +proof (insert assms, induction d arbitrary: b') + case 0 + then show ?case using prod_encode_eq by simp +next + case (Suc d) + moreover have "t + d \ t + Suc d" by simp + ultimately have "prefixes (t + d) j \" + using prefixes_converg_le by blast + then obtain b'' where b'': "prefixes (t + d) j \= b''" + using prefixes_prod_encode by blast + with Suc have "prefixes (t + Suc d) j = nxt b''" + by (simp add: prefixes_at_Suc') + with Suc have "nxt b'' \= b'" by simp + then show ?case using nxt_monotone Suc b'' by fastforce +qed + +lemma prefixes_stable: + assumes "prefixes t j \= b" and "prefixes (t + d) j \= b'" + shows "\x t + Suc d" by simp + ultimately have "prefixes (t + d) j \" + using prefixes_converg_le by blast + then obtain b'' where b'': "prefixes (t + d) j \= b''" + using prefixes_prod_encode by blast + with Suc have "prefixes (t + Suc d) j = nxt b''" + by (simp add: prefixes_at_Suc') + with Suc have b': "nxt b'' \= b'" by simp + show "\x e_length b''" + using x prefixes_monotone b'' Suc by fastforce + ultimately show "e_nth b x = e_nth b' x" + using b'' nxt_stable Suc b' prefixes_monotone x + by (metis leD le_neq_implies_less) + qed +qed + +lemma prefixes_tl_only_01: + assumes "prefixes t j \= b" + shows "\x>0. e_nth b x = 0 \ e_nth b x = 1" +proof (insert assms, induction t arbitrary: b) + case 0 + then show ?case using prefixes_at_0 prod_encode_eq by auto +next + case (Suc t) + then have "prefixes t j \" + using prefixes_converg_le Suc_n_not_le_n nat_le_linear by blast + then obtain b' where b': "prefixes t j \= b'" + using prefixes_prod_encode by blast + show "\x>0. e_nth b x = 0 \ e_nth b x = 1" + proof (rule allI, rule impI) + fix x :: nat + assume x: "x > 0" + show "e_nth b x = 0 \ e_nth b x = 1" + proof (cases "x < e_length b'") + case True + then show ?thesis + using Suc b' prefixes_at_Suc' nxt_stable x by metis + next + case False + then show ?thesis + using Suc.prems b' prefixes_at_Suc' nxt_append_01 by auto + qed + qed +qed + +lemma prefixes_hd: + assumes "prefixes t j \= b" + shows "e_nth b 0 = j" +proof - + obtain b' where b': "prefixes 0 j \= b'" + by (simp add: prefixes_at_0) + then have "b' = list_encode [j]" + by (simp add: prod_encode_eq prefixes_at_0) + then have "e_nth b' 0 = j" by simp + then show "e_nth b 0 = j" + using assms prefixes_stable[OF b', of t b] prefixes_length[OF b'] by simp +qed + +lemma prefixes_nontotal_hyp: + assumes "prefixes t j \= b" + and "prefixes (Suc t) j \" + and "s b \= i" + shows "\x. \ i x \" + using nxt_nontotal_hyp[OF _ assms(3)] assms(2) prefixes_at_Suc'[OF assms(1)] by simp + +text \We now consider the two cases from the proof sketch.\ + +abbreviation "case_two j \ \t. prefixes t j \" + +abbreviation "case_one j \ \ case_two j" + +text \In Case~2 there is a maximum convergent iteration because +iteration 0 converges.\ + +lemma case_two: + assumes "case_two j" + shows "\t. (\t'\t. prefixes t' j \) \ (\t'>t. prefixes t' j \)" +proof - + let ?P = "\t. prefixes t j \" + define t\<^sub>0 where "t\<^sub>0 = Least ?P" + then have "?P t\<^sub>0" + using assms LeastI_ex[of ?P] by simp + then have diverg: "?P t" if "t \ t\<^sub>0" for t + using prefixes_converg_le that by blast + from t\<^sub>0_def have converg: "\ ?P t" if "t < t\<^sub>0" for t + using Least_le[of ?P] that not_less by blast + have "t\<^sub>0 > 0" + proof (rule ccontr) + assume "\ 0 < t\<^sub>0" + then have "t\<^sub>0 = 0" by simp + with `?P t\<^sub>0` prefixes_at_0 show False by simp + qed + let ?t = "t\<^sub>0 - 1" + have "\t'\?t. prefixes t' j \" + using converg \0 < t\<^sub>0\ by auto + moreover have "\t'>?t. prefixes t' j \" + using diverg by simp + ultimately show ?thesis by auto +qed + +text \Having completed the modelling of the process, we can now define +the functions $\psi_j$ it computes. The value $\psi_j(x)$ is computed by +running @{term r_prefixes} until the prefix is longer than $x$ and then +taking the $x$-th element of the prefix.\ + +definition "r_psi \ + let f = Cn 3 r_less [Id 3 2, Cn 3 r_length [Cn 3 r_prefixes [Id 3 0, Id 3 1]]] + in Cn 2 r_nth [Cn 2 r_prefixes [Mn 2 f, Id 2 0], Id 2 1]" + +lemma r_psi_recfn: "recfn 2 r_psi" + unfolding r_psi_def by simp + +abbreviation psi :: partial2 ("\") where + "\ j x \ eval r_psi [j, x]" + +lemma psi_in_P2: "\ \ \

\<^sup>2" + using r_psi_recfn by auto + +text \The values of @{term "\"} can be read off the prefixes.\ + +lemma psi_eq_nth_prefix: + assumes "prefixes t j \= b" and "e_length b > x" + shows "\ j x \= e_nth b x" +proof - + let ?f = "Cn 3 r_less [Id 3 2, Cn 3 r_length [Cn 3 r_prefixes [Id 3 0, Id 3 1]]]" + let ?P = "\t. prefixes t j \ \ e_length (the (prefixes t j)) > x" + from assms have ex_t: "\t. ?P t" by auto + define t\<^sub>0 where "t\<^sub>0 = Least ?P" + then have "?P t\<^sub>0" + using LeastI_ex[OF ex_t] by simp + from ex_t have not_P: "\ ?P t" if "t < t\<^sub>0" for t + using ex_t that Least_le[of ?P] not_le t\<^sub>0_def by auto + + have "?P t" using assms by simp + with not_P have "t\<^sub>0 \ t" using leI by blast + then obtain b\<^sub>0 where b0: "prefixes t\<^sub>0 j \= b\<^sub>0" + using assms(1) prefixes_converg_le by blast + + have "eval ?f [t\<^sub>0, j, x] \= 0" + proof - + have "eval (Cn 3 r_prefixes [Id 3 0, Id 3 1]) [t\<^sub>0, j, x] \= b\<^sub>0" + using b0 by simp + then show ?thesis using `?P t\<^sub>0` by simp + qed + moreover have "eval ?f [t, j, x] \\ 0" if "t < t\<^sub>0" for t + proof - + obtain bt where bt: "prefixes t j \= bt" + using prefixes_converg_le[of t\<^sub>0 j t] b0 `t < t\<^sub>0` by auto + moreover have "\ ?P t" + using that not_P by simp + ultimately have "e_length bt \ x" by simp + moreover have "eval (Cn 3 r_prefixes [Id 3 0, Id 3 1]) [t, j, x] \= bt" + using bt by simp + ultimately show ?thesis by simp + qed + ultimately have "eval (Mn 2 ?f) [j, x] \= t\<^sub>0" + using eval_Mn_convergI[of 2 ?f "[j, x]" t\<^sub>0] by simp + then have "\ j x \= e_nth b\<^sub>0 x" + unfolding r_psi_def using b0 by simp + then show ?thesis + using `t\<^sub>0 \ t` assms(1) prefixes_stable[of t\<^sub>0 j b\<^sub>0 "t - t\<^sub>0" b] b0 `?P t\<^sub>0` + by simp +qed + +lemma psi_converg_imp_prefix: + assumes "\ j x \" + shows "\t b. prefixes t j \= b \ e_length b > x" +proof - + let ?f = "Cn 3 r_less [Id 3 2, Cn 3 r_length [Cn 3 r_prefixes [Id 3 0, Id 3 1]]]" + have "eval (Mn 2 ?f) [j, x] \" + proof (rule ccontr) + assume "\ eval (Mn 2 ?f) [j, x] \" + then have "eval (Mn 2 ?f) [j, x] \" by simp + then have "\ j x \" + unfolding r_psi_def by simp + then show False + using assms by simp + qed + then obtain t where t: "eval (Mn 2 ?f) [j, x] \= t" + by blast + have "recfn 2 (Mn 2 ?f)" by simp + then have f_zero: "eval ?f [t, j, x] \= 0" + using eval_Mn_convergE[OF _ t] + by (metis (no_types, lifting) One_nat_def Suc_1 length_Cons list.size(3)) + have "prefixes t j \" + proof (rule ccontr) + assume "\ prefixes t j \" + then have "prefixes t j \" by simp + then have "eval ?f [t, j, x] \" by simp + with f_zero show False by simp + qed + then obtain b' where b': "prefixes t j \= b'" by auto + moreover have "e_length b' > x" + proof (rule ccontr) + assume "\ e_length b' > x" + then have "eval ?f [t, j, x] \= 1" + using b' by simp + with f_zero show False by simp + qed + ultimately show ?thesis by auto +qed + +lemma psi_converg_imp_prefix': + assumes "\ j x \" + shows "\t b. prefixes t j \= b \ e_length b > x \ \ j x \= e_nth b x" + using psi_converg_imp_prefix[OF assms] psi_eq_nth_prefix by blast + +text \In both Case~1 and~2, $\psi_j$ starts with $j$.\ + +lemma psi_at_0: "\ j 0 \= j" + using prefixes_hd prefixes_length psi_eq_nth_prefix prefixes_at_0 by fastforce + +text \In Case~1, $\psi_j$ is total and made up of $j$ followed by zeros +and ones, just as required by the definition of $V_1$.\ + +lemma case_one_psi_total: + assumes "case_one j" and "x > 0" + shows "\ j x \= 0 \ \ j x \= 1" +proof - + obtain b where b: "prefixes x j \= b" + using assms(1) by auto + then have "e_length b > x" + using prefixes_length by simp + then have "\ j x \= e_nth b x" + using b psi_eq_nth_prefix by simp + moreover have "e_nth b x = 0 \ e_nth b x = 1" + using prefixes_tl_only_01[OF b] assms(2) by simp + ultimately show "\ j x \= 0 \ \ j x \= 1" + by simp +qed + +text \In Case~2, $\psi_j$ is defined only for a prefix starting with +$j$ and continuing with zeros and ones. This prefix corresponds to $ja$ from +the definition of $V_2$.\ + +lemma case_two_psi_only_prefix: + assumes "case_two j" + shows "\y. (\x. 0 < x \ x < y \ \ j x \= 0 \ \ j x \= 1) \ + (\x \ y. \ j x \)" +proof - + obtain t where + t_le: "\t'\t. prefixes t' j \" and + t_gr: "\t'>t. prefixes t' j \" + using assms case_two by blast + then obtain b where b: "prefixes t j \= b" + by auto + let ?y = "e_length b" + have "\ j x \= 0 \ \ j x \= 1" if "x > 0 \ x < ?y" for x + using t_le b that by (metis prefixes_tl_only_01 psi_eq_nth_prefix) + moreover have "\ j x \" if "x \ ?y" for x + proof (rule ccontr) + assume "\ j x \" + then obtain t' b' where t': "prefixes t' j \= b'" and "e_length b' > x" + using psi_converg_imp_prefix by blast + then have "e_length b' > ?y" + using that by simp + with t' have "t' > t" + using prefixes_monotone b by (metis add_diff_inverse_nat leD) + with t' t_gr show False by simp + qed + ultimately show ?thesis by auto +qed + +definition longest_prefix :: "nat \ nat" where + "longest_prefix j \ THE y. (\x j x \) \ (\x\y. \ j x \)" + +lemma longest_prefix: + assumes "case_two j" and "z = longest_prefix j" + shows "(\x j x \) \ (\x\z. \ j x \)" +proof - + let ?P = "\z. (\x j x \) \ (\x\z. \ j x \)" + obtain y where y: + "\x. 0 < x \ x < y \ \ j x \= 0 \ \ j x \= 1" + "\x\y. \ j x \" + using case_two_psi_only_prefix[OF assms(1)] by auto + have "?P (THE z. ?P z)" + proof (rule theI[of ?P y]) + show "?P y" + proof + show "\x j x \" + proof (rule allI, rule impI) + fix x assume "x < y" + show "\ j x \" + proof (cases "x = 0") + case True + then show ?thesis using psi_at_0 by simp + next + case False + then show ?thesis using y(1) `x < y` by auto + qed + qed + show "\x\y. \ j x \" using y(2) by simp + qed + show "z = y" if "?P z" for z + proof (rule ccontr, cases "z < y") + case True + moreover assume "z \ y" + ultimately show False + using that `?P y` by auto + next + case False + moreover assume "z \ y" + then show False + using that `?P y` y(2) by (meson linorder_cases order_refl) + qed + qed + then have "(\x<(THE z. ?P z). \ j x \) \ (\x\(THE z. ?P z). \ j x \)" + by blast + moreover have "longest_prefix j = (THE z. ?P z)" + unfolding longest_prefix_def by simp + ultimately show ?thesis using assms(2) by metis +qed + +lemma case_two_psi_longest_prefix: + assumes "case_two j" and "y = longest_prefix j" + shows "(\x. 0 < x \ x < y \ \ j x \= 0 \ \ j x \= 1) \ + (\x \ y. \ j x \)" + using assms longest_prefix case_two_psi_only_prefix + by (metis prefixes_tl_only_01 psi_converg_imp_prefix') + +text \The prefix cannot be empty because the process starts with prefix $[j]$.\ + +lemma longest_prefix_gr_0: + assumes "case_two j" + shows "longest_prefix j > 0" + using assms case_two_psi_longest_prefix psi_at_0 by force + +lemma psi_not_divergent_init: + assumes "prefixes t j \= b" + shows "(\ j) \ (e_length b - 1) = b" +proof (intro initI) + show "0 < e_length b" + using assms prefixes_length by fastforce + show "\ j x \= e_nth b x" if "x < e_length b" for x + using that assms psi_eq_nth_prefix by simp +qed + +text \In Case~2, the strategy $S$ outputs a non-total hypothesis on +some prefix of $\psi_j$.\ + +lemma case_two_nontotal_hyp: + assumes "case_two j" + shows "\n total1 (\ (the (s ((\ j) \ n))))" +proof - + obtain t where "\t'\t. prefixes t' j \" and t_gr: "\t'>t. prefixes t' j \" + using assms case_two by blast + then obtain b where b: "prefixes t j \= b" + by auto + moreover obtain i where i: "s b \= i" + using eval_rs by fastforce + moreover have div: "prefixes (Suc t) j \" + using t_gr by simp + ultimately have "\x. \ i x \" + using prefixes_nontotal_hyp by simp + then obtain x where "\ i x \" by auto + moreover have init: "\ j \ (e_length b - 1) = b" (is "_ \ ?n = b") + using psi_not_divergent_init[OF b] by simp + ultimately have "\ (the (s (\ j \ ?n))) x \" + using i by simp + then have "\ total1 (\ (the (s (\ j \ ?n))))" + by auto + moreover have "?n < longest_prefix j" + using case_two_psi_longest_prefix init b div psi_eq_nth_prefix + by (metis length_init lessI not_le_imp_less option.simps(3)) + ultimately show ?thesis by auto +qed + +text \Consequently, in Case~2 the strategy does not TOTAL-learn +any function starting with the longest prefix of $\psi_j$.\ + +lemma case_two_not_learn: + assumes "case_two j" + and "f \ \" + and "\x. x < longest_prefix j \ f x = \ j x" + shows "\ learn_total \ {f} s" +proof - + obtain n where n: + "n < longest_prefix j" + "\ total1 (\ (the (s (\ j \ n))))" + using case_two_nontotal_hyp[OF assms(1)] by auto + have "f \ n = \ j \ n" + using assms(3) n(1) by (intro init_eqI) auto + with n(2) show ?thesis by (metis R1_imp_total1 learn_totalE(3) singletonI) +qed + +text \In Case~1 the strategy outputs a wrong hypothesis +on infinitely many prefixes of $\psi_j$ and thus does not +learn $\psi_j$ in the limit, much less in the sense of TOTAL.\ + +lemma case_one_wrong_hyp: + assumes "case_one j" + shows "\n>k. \ (the (s ((\ j) \ n))) \ \ j" +proof - + have all_t: "\t. prefixes t j \" + using assms by simp + then obtain b where b: "prefixes (Suc k) j \= b" + by auto + then have length: "e_length b > Suc k" + using prefixes_length by simp + then have init: "\ j \ (e_length b - 1) = b" + using psi_not_divergent_init b by simp + obtain i where i: "s b \= i" + using eval_rs by fastforce + from all_t obtain b' where b': "prefixes (Suc (Suc k)) j \= b'" + by auto + then have "\ j \ (e_length b' - 1) = b'" + using psi_not_divergent_init by simp + moreover have "\y i y \\ e_nth b' y" + using nxt_wrong_hyp b b' i prefixes_at_Suc by auto + ultimately have "\y i y \ \ j y" + using b' psi_eq_nth_prefix by auto + then have "\ i \ \ j" by auto + then show ?thesis + using init length i by (metis Suc_less_eq length_init option.sel) +qed + +lemma case_one_not_learn: + assumes "case_one j" + shows "\ learn_lim \ {\ j} s" +proof (rule infinite_hyp_wrong_not_Lim[of "\ j"]) + show "\ j \ {\ j}" by simp + show "\n. \m>n. \ (the (s (\ j \ m))) \ \ j" + using case_one_wrong_hyp[OF assms] by simp +qed + +lemma case_one_not_learn_V: + assumes "case_one j" and "j \ 2" and "\ j = \ j" + shows "\ learn_lim \ V_constotal s" +proof - + have "\ j \ V_constotal_1" + proof - + define p where "p = (\x. (\ j) (x + 1))" + have "p \ \\<^sub>0\<^sub>1" + proof - + from p_def have "p \ \

" + using skip_P1[of "\ j" 1] psi_in_P2 P2_proj_P1 by blast + moreover have "p x \= 0 \ p x \= 1" for x + using p_def assms(1) case_one_psi_total by auto + moreover from this have "total1 p" by fast + ultimately show ?thesis using RPred1_def by auto + qed + moreover have "\ j = [j] \ p" + by (intro prepend_eqI, simp add: psi_at_0, simp add: p_def) + ultimately show ?thesis using assms(2,3) V_constotal_1_def by blast + qed + then have "\ j \ V_constotal" using V_constotal_def by auto + moreover have "\ learn_lim \ {\ j} s" + using case_one_not_learn assms(1) by simp + ultimately show ?thesis using learn_lim_closed_subseteq by auto +qed + +text \The next lemma embodies the construction of $\chi$ followed by +the application of Kleene's fixed-point theorem as described in the +proof sketch.\ + +lemma goedel_after_prefixes: + fixes vs :: "nat list" and m :: nat + shows "\n\m. \ n = vs @ [n] \ 0\<^sup>\" +proof - + define f :: partial1 where "f \ vs \ 0\<^sup>\" + then have "f \ \" + using almost0_in_R1 by auto + then obtain n where n: + "n \ m" + "\ n = (\x. if x = length vs then Some n else f x)" + using goedel_at[of f m "length vs"] by auto + moreover have "\ n x = (vs @ [n] \ 0\<^sup>\) x" for x + proof - + consider "x < length vs" | "x = length vs" | "x > length vs" + by linarith + then show ?thesis + using n f_def by (cases) (auto simp add: prepend_associative) + qed + ultimately show ?thesis by blast +qed + +text \If Case~2 holds for a $j\geq 2$ with $\varphi_j = \psi_j$, that +is, if $\psi_j\in V_1$, then there is a function in $V$, namely $\psi_j$, on +which $S$ fails. Therefore $S$ does not learn $V$.\ + +lemma case_two_not_learn_V: + assumes "case_two j" and "j \ 2" and "\ j = \ j" + shows "\ learn_total \ V_constotal s" +proof - + define z where "z = longest_prefix j" + then have "z > 0" + using longest_prefix_gr_0[OF assms(1)] by simp + define vs where "vs = prefix (\ j) (z - 1)" + then have "vs ! 0 = j" + using psi_at_0 `z > 0` by simp + define a where "a = tl vs" + then have vs: "vs = j # a" + using vs_def `vs ! 0 = j` + by (metis length_Suc_conv length_prefix list.sel(3) nth_Cons_0) + obtain k where k: "k \ 2" and phi_k: "\ k = j # a @ [k] \ 0\<^sup>\" + using goedel_after_prefixes[of 2 "j # a"] by auto + have phi_j: "\ j = j # a \ \\<^sup>\ " + proof (rule prepend_eqI) + show "\x. x < length (j # a) \ \ j x \= (j # a) ! x" + using assms(1,3) vs vs_def \0 < z\ + length_prefix[of "\ j" "z - 1"] + prefix_nth[of _ _ "\ j"] + psi_at_0[of j] + case_two_psi_longest_prefix[OF _ z_def] + longest_prefix[OF _ z_def] + by (metis One_nat_def Suc_pred option.collapse) + show "\x. \ j (length (j # a) + x) \" + using assms(3) vs_def + by (simp add: vs assms(1) case_two_psi_longest_prefix z_def) + qed + moreover have "\ k \ V_constotal_2" + proof (intro V_constotal_2I[of _ j a k]) + show "\ k = j # a @ [k] \ 0\<^sup>\" + using phi_k . + show "2 \ j" + using `2 \ j` . + show "2 \ k" + using `2 \ k` . + show "\i 1" + proof (rule allI, rule impI) + fix i assume i: "i < length a" + then have "Suc i < z" + using z_def vs_def length_prefix \0 < z\ vs + by (metis One_nat_def Suc_mono Suc_pred length_Cons) + have "a ! i = vs ! (Suc i)" + using vs by simp + also have "... = the (\ j (Suc i))" + using vs_def vs i length_Cons length_prefix prefix_nth + by (metis Suc_mono) + finally show "a ! i \ 1" + using case_two_psi_longest_prefix `Suc i < z` z_def + by (metis assms(1) less_or_eq_imp_le not_le_imp_less not_one_less_zero + option.sel zero_less_Suc) + qed + qed (auto simp add: phi_j) + then have "\ k \ V_constotal" + using V_constotal_def by auto + moreover have "\ learn_total \ {\ k} s" + proof - + have "\ k \ \" + by (simp add: phi_k almost0_in_R1) + moreover have "\x. x < longest_prefix j \ \ k x = \ j x" + using phi_k vs_def z_def length_prefix phi_j prepend_associative prepend_at_less + by (metis One_nat_def Suc_pred \0 < z\ \vs = j # a\ append_Cons assms(3)) + ultimately show ?thesis + using case_two_not_learn[OF assms(1)] by simp + qed + ultimately show "\ learn_total \ V_constotal s" + using learn_total_closed_subseteq by auto +qed + +text \The strategy $S$ does not learn $V$ in either case.\ + +lemma not_learn_total_V: "\ learn_total \ V_constotal s" +proof - + obtain j where "j \ 2" "\ j = \ j" + using kleene_fixed_point psi_in_P2 by auto + then show ?thesis + using case_one_not_learn_V learn_total_def case_two_not_learn_V + by (cases "case_two j") auto +qed + +end + + +lemma V_not_in_TOTAL: "V_constotal \ TOTAL" +proof (rule ccontr) + assume "\ V_constotal \ TOTAL" + then have "V_constotal \ TOTAL" by simp + then have "V_constotal \ TOTAL_wrt \" + by (simp add: TOTAL_wrt_phi_eq_TOTAL) + then obtain s where "learn_total \ V_constotal s" + using TOTAL_wrt_def by auto + then obtain s' where s': "s' \ \" "learn_total \ V_constotal s'" + using lemma_R_for_TOTAL_simple by blast + then interpret total_cons s' + by (simp add: total_cons_def) + have "\ learn_total \ V_constotal s'" + by (simp add: not_learn_total_V) + with s'(2) show False by simp +qed + +lemma TOTAL_neq_CONS: "TOTAL \ CONS" + using V_not_in_TOTAL V_in_CONS CONS_def by auto + +text \The main result of this section:\ + +theorem TOTAL_subset_CONS: "TOTAL \ CONS" + using TOTAL_subseteq_CONS TOTAL_neq_CONS by simp + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/Union.thy b/thys/Inductive_Inference/Union.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/Union.thy @@ -0,0 +1,283 @@ +section \The union of classes\label{s:union}\ + +theory Union + imports R1_BC TOTAL_CONS +begin + +text \None of the inference types introduced in this chapter are closed +under union of classes. For all inference types except FIN this follows from +@{thm[source] "U0_V0_not_in_BC"}.\ + +lemma not_closed_under_union: + "\\\{CP, TOTAL, CONS, LIM, BC}. U\<^sub>0 \ \ \ V\<^sub>0 \ \ \ U\<^sub>0 \ V\<^sub>0 \ \" + using U0_in_CP U0_in_NUM V0_in_FIN + FIN_subseteq_CP + NUM_subseteq_TOTAL + CP_subseteq_TOTAL + TOTAL_subseteq_CONS + CONS_subseteq_Lim + Lim_subseteq_BC + U0_V0_not_in_BC + by blast + +text \In order to show the analogous result for FIN consider the +classes $\{0^\infty\}$ and $\{0^n10^\infty \mid n \in \mathbb{N}\}$. The +former can be learned finitely by a strategy that hypothesizes $0^\infty$ for +every input. The latter can be learned finitely by a strategy that waits for +the 1 and hypothesizes the only function in the class with a 1 at that +position. However, the union of both classes is not in FIN. This is because +any FIN strategy has to hypothesize $0^\infty$ on some prefix of the form +$0^n$. But the strategy then fails for the function $0^n10^\infty$.\ + +lemma singleton_in_FIN: "f \ \ \ {f} \ FIN" +proof - + assume "f \ \" + then obtain i where i: "\ i = f" + using phi_universal by blast + define s :: partial1 where "s = (\_. Some (Suc i))" + then have "s \ \" + using const_in_Prim1[of "Suc i"] by simp + have "learn_fin \ {f} s" + proof (intro learn_finI) + show "environment \ {f} s" + using `s \ \` `f \ \` by (simp add: phi_in_P2) + show "\i n\<^sub>0. \ i = g \ (\n0. s (g \ n) \= 0) \ (\n\n\<^sub>0. s (g \ n) \= Suc i)" + if "g \ {f}" for g + proof - + from that have "g = f" by simp + then have "\ i = g" + using i by simp + moreover have "\n<0. s (g \ n) \= 0" by simp + moreover have "\n\0. s (g \ n) \= Suc i" + using s_def by simp + ultimately show ?thesis by auto + qed + qed + then show "{f} \ FIN" using FIN_def by auto +qed + +definition U_single :: "partial1 set" where + "U_single \ {(\x. if x = n then Some 1 else Some 0)| n. n \ UNIV}" + +lemma U_single_in_FIN: "U_single \ FIN" +proof - + define psi :: partial2 where "psi \ \n x. if x = n then Some 1 else Some 0" + have "psi \ \\<^sup>2" + using psi_def by (intro R2I[of "Cn 2 r_not [r_eq]"]) auto + define s :: partial1 where + "s \ \b. if findr b \= e_length b then Some 0 else Some (Suc (the (findr b)))" + have "s \ \" + proof (rule R1I) + let ?r = "Cn 1 r_ifeq [r_findr, r_length, Z, Cn 1 S [r_findr]]" + show "recfn 1 ?r" by simp + show "total ?r" by auto + show "eval ?r [b] = s b" for b + proof - + let ?b = "the (findr b)" + have "eval ?r [b] = (if ?b = e_length b then Some 0 else Some (Suc (?b)))" + using findr_total by simp + then show "eval ?r [b] = s b" + by (metis findr_total option.collapse option.inject s_def) + qed + qed + have "U_single \ \" + proof + fix f + assume "f \ U_single" + then obtain n where "f = (\x. if x = n then Some 1 else Some 0)" + using U_single_def by auto + then have "f = psi n" + using psi_def by simp + then show "f \ \" + using `psi \ \\<^sup>2` by simp + qed + have "learn_fin psi U_single s" + proof (rule learn_finI) + show "environment psi U_single s" + using `psi \ \\<^sup>2` `s \ \` `U_single \ \` by simp + show "\i n\<^sub>0. psi i = f \ (\n0. s (f \ n) \= 0) \ (\n\n\<^sub>0. s (f \ n) \= Suc i)" + if "f \ U_single" for f + proof - + from that obtain i where i: "f = (\x. if x = i then Some 1 else Some 0)" + using U_single_def by auto + then have "psi i = f" + using psi_def by simp + moreover have "\n n) \= 0" + using i s_def findr_def by simp + moreover have "\n\i. s (f \ n) \= Suc i" + proof (rule allI, rule impI) + fix n + assume "n \ i" + let ?e = "init f n" + have "\i 0" + using `n \ i` i by simp + then have less: "the (findr ?e) < e_length ?e" + and nth_e: "e_nth ?e (the (findr ?e)) \ 0" + using findr_ex by blast+ + then have "s ?e \= Suc (the (findr ?e))" + using s_def by auto + moreover have "the (findr ?e) = i" + using nth_e less i by (metis length_init nth_init option.sel) + ultimately show "s ?e \= Suc i" by simp + qed + ultimately show ?thesis by auto + qed + qed + then show "U_single \ FIN" using FIN_def by blast +qed + +lemma zero_U_single_not_in_FIN: "{0\<^sup>\} \ U_single \ FIN" +proof + assume "{0\<^sup>\} \ U_single \ FIN" + then obtain psi s where learn: "learn_fin psi ({0\<^sup>\} \ U_single) s" + using FIN_def by blast + then have "learn_fin psi {0\<^sup>\} s" + using learn_fin_closed_subseteq by auto + then obtain i n\<^sub>0 where i: + "psi i = 0\<^sup>\" + "\n0. s (0\<^sup>\ \ n) \= 0" + "\n\n\<^sub>0. s (0\<^sup>\ \ n) \= Suc i" + using learn_finE(2) by blast + let ?f = "\x. if x = Suc n\<^sub>0 then Some 1 else Some 0" + have "?f \ 0\<^sup>\" by (metis option.inject zero_neq_one) + have "?f \ U_single" + using U_single_def by auto + then have "learn_fin psi {?f} s" + using learn learn_fin_closed_subseteq by simp + then obtain j m\<^sub>0 where j: + "psi j = ?f" + "\n0. s (?f \ n) \= 0" + "\n\m\<^sub>0. s (?f \ n) \= Suc j" + using learn_finE(2) by blast + consider + (less) "m\<^sub>0 < n\<^sub>0" | (eq) "m\<^sub>0 = n\<^sub>0" | (gr) "m\<^sub>0 > n\<^sub>0" + by linarith + then show False + proof (cases) + case less + then have "s (0\<^sup>\\ m\<^sub>0) \= 0" + using i by simp + moreover have "0\<^sup>\ \ m\<^sub>0 = ?f \ m\<^sub>0" + using less init_eqI[of m\<^sub>0 ?f "0\<^sup>\"] by simp + ultimately have "s (?f \ m\<^sub>0) \= 0" by simp + then show False using j by simp + next + case eq + then have "0\<^sup>\ \ m\<^sub>0 = ?f \ m\<^sub>0" + using init_eqI[of m\<^sub>0 ?f "0\<^sup>\"] by simp + then have "s (0\<^sup>\ \ m\<^sub>0) = s (?f \ m\<^sub>0)" by simp + then have "i = j" + using i j eq by simp + then have "psi i = psi j" by simp + then show False using `?f \ 0\<^sup>\` i j by simp + next + case gr + have "0\<^sup>\ \ n\<^sub>0 = ?f \ n\<^sub>0" + using init_eqI[of n\<^sub>0 ?f "0\<^sup>\"] by simp + moreover have "s (0\<^sup>\ \ n\<^sub>0) \= Suc i" + using i by simp + moreover have "s (?f \ n\<^sub>0) \= 0" + using j gr by simp + ultimately show False by simp + qed +qed + +lemma FIN_not_closed_under_union: "\U V. U \ FIN \ V \ FIN \ U \ V \ FIN" +proof - + have "{0\<^sup>\} \ FIN" + using singleton_in_FIN const_in_Prim1 by simp + moreover have "U_single \ FIN" + using U_single_in_FIN by simp + ultimately show ?thesis + using zero_U_single_not_in_FIN by blast +qed + +text \In contrast to the inference types, NUM is closed under the union +of classes. The total numberings that exist for each NUM class can be +interleaved to produce a total numbering encompassing the union of the +classes. To define the interleaving, modulo and division by two will be +helpful.\ + +definition "r_div2 \ + r_shrink + (Pr 1 Z + (Cn 3 r_ifle + [Cn 3 r_mul [r_constn 2 2, Cn 3 S [Id 3 0]], Id 3 2, Cn 3 S [Id 3 1], Id 3 1]))" + +lemma r_div2_prim [simp]: "prim_recfn 1 r_div2" + unfolding r_div2_def by simp + +lemma r_div2 [simp]: "eval r_div2 [n] \= n div 2" +proof - + let ?p = "Pr 1 Z + (Cn 3 r_ifle + [Cn 3 r_mul [r_constn 2 2, Cn 3 S [Id 3 0]], Id 3 2, Cn 3 S [Id 3 1], Id 3 1])" + have "eval ?p [i, n] \= min (n div 2) i" for i + by (induction i) auto + then have "eval ?p [n, n] \= n div 2" by simp + then show ?thesis unfolding r_div2_def by simp +qed + +definition "r_mod2 \ Cn 1 r_sub [Id 1 0, Cn 1 r_mul [r_const 2, r_div2]]" + +lemma r_mod2_prim [simp]: "prim_recfn 1 r_mod2" + unfolding r_mod2_def by simp + +lemma r_mod2 [simp]: "eval r_mod2 [n] \= n mod 2" + unfolding r_mod2_def using Rings.semiring_modulo_class.minus_mult_div_eq_mod + by auto + +lemma NUM_closed_under_union: + assumes "U \ NUM" and "V \ NUM" + shows "U \ V \ NUM" +proof - + from assms obtain psi_u psi_v where + psi_u: "psi_u \ \\<^sup>2" "\f. f \ U \ \i. psi_u i = f" and + psi_v: "psi_v \ \\<^sup>2" "\f. f \ V \ \i. psi_v i = f" + by fastforce + define psi where "psi \ \i. if i mod 2 = 0 then psi_u (i div 2) else psi_v (i div 2)" + from psi_u(1) obtain u where u: "recfn 2 u" "total u" "\x y. eval u [x, y] = psi_u x y" + by auto + from psi_v(1) obtain v where v: "recfn 2 v" "total v" "\x y. eval v [x, y] = psi_v x y" + by auto + let ?r_psi = "Cn 2 r_ifz + [Cn 2 r_mod2 [Id 2 0], + Cn 2 u [Cn 2 r_div2 [Id 2 0], Id 2 1], + Cn 2 v [Cn 2 r_div2 [Id 2 0], Id 2 1]]" + show ?thesis + proof (rule NUM_I[of psi]) + show "psi \ \\<^sup>2" + proof (rule R2I) + show "recfn 2 ?r_psi" + using u(1) v(1) by simp + show "eval ?r_psi [x, y] = psi x y" for x y + using u v psi_def prim_recfn_total R2_imp_total2[OF psi_u(1)] + R2_imp_total2[OF psi_v(1)] + by simp + moreover have "psi x y \" for x y + using psi_def psi_u(1) psi_v(1) by simp + ultimately show "total ?r_psi" + using `recfn 2 ?r_psi` totalI2 by simp + qed + show "\i. psi i = f" if "f \ U \ V" for f + proof (cases "f \ U") + case True + then obtain j where "psi_u j = f" + using psi_u(2) by auto + then have "psi (2 * j) = f" + using psi_def by simp + then show ?thesis by auto + next + case False + then have "f \ V" + using that by simp + then obtain j where "psi_v j = f" + using psi_v(2) by auto + then have "psi (Suc (2 * j)) = f" + using psi_def by simp + then show ?thesis by auto + qed + qed +qed + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/Universal.thy b/thys/Inductive_Inference/Universal.thy new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/Universal.thy @@ -0,0 +1,2537 @@ +section \A universal partial recursive function\ + +theory Universal + imports Partial_Recursive +begin + +text \The main product of this section is a universal partial recursive +function, which given a code $i$ of an $n$-ary partial recursive function $f$ +and an encoded list @{term xs} of $n$ arguments, computes @{term "eval f +xs"}. From this we can derive fixed-arity universal functions satisfying the +usual results such as the $s$-$m$-$n$ theorem. To represent the code $i$, we +need a way to encode @{typ recf}s as natural numbers (Section~\ref{s:recf_enc}). To +construct the universal function, we devise a ternary function taking $i$, +$xs$, and a step bound $t$ and simulating the execution of $f$ on input $xs$ for +$t$ steps. This function is useful in its own right, enabling techniques like +dovetailing or ``concurrent'' evaluation of partial recursive functions. + +The notion of a ``step'' is not part of the definition of (the evaluation of) +partial recursive functions, but one can simulate the evaluation on an +abstract machine (Section~\ref{s:step}). This machine's configurations can be +encoded as natural numbers, and this leads us to a step function @{typ "nat +\ nat"} on encoded configurations (Section~\ref{s:step_enc}). +This function in turn can be computed by a primitive recursive function, from +which we develop the aforementioned ternary function of $i$, @{term xs}, and +$t$ (Section~\ref{s:step_recf}). From this we can finally derive +a universal function (Section~\ref{s:the_universal}).\ + +subsection \A step function\label{s:step}\ + +text \We simulate the stepwise execution of a partial recursive +function in a fairly straightforward way reminiscent of the execution of +function calls in an imperative programming language. A configuration of the +abstract machine is a pair consisting of: +\begin{enumerate} +\item A stack of frames. A frame represents the execution of a function and is + a triple @{term "(f, xs, locals)"} of + \begin{enumerate} + \item a @{typ recf} @{term f} being executed, + \item a @{typ "nat list"} of arguments of @{term f}, + \item a @{typ "nat list"} of local variables, which holds intermediate + values when @{term f} is of the form @{term Cn}, @{term Pr}, or @{term Mn}. + \end{enumerate} +\item A register of type @{typ "nat option"} representing the return value of + the last function call: @{term None} signals that in the previous step the + stack was not popped and hence no value was returned, whereas @{term "Some + v"} means that in the previous step a function returned @{term v}. +\end{enumerate} +For computing @{term h} on input @{term xs}, the initial configuration is +@{term "([(h, xs, [])], None)"}. When the computation for a frame ends, it is +popped off the stack, and its return value is put in the register. The entire +computation ends when the stack is empty. In such a final configuration the +register contains the value of @{term h} at @{term xs}. If no final +configuration is ever reached, @{term h} diverges at @{term xs}. + +The execution of one step depends on the topmost (that is, active) frame. In +the step when a frame @{term "(h, xs, locals)"} is pushed onto the stack, the +local variables are @{term "locals = []"}. The following happens until the +frame is popped off the stack again (if it ever is): +\begin{itemize} +\item For the base functions @{term "h = Z"}, @{term "h = S"}, + @{term[names_short] "h = Id m n"}, the frame is popped off the stack right away, + and the return value is placed in the register. +\item For @{term "h = Cn n f gs"}, for each function $g$ in @{term gs}: + \begin{enumerate} + \item A new frame of the form @{term "(g, xs, [])"} is pushed onto the stack. + \item When (and if) this frame + is eventually popped, the value in the register is @{term "eval g xs"}. This value + is appended to the list @{term locals} of local variables. + \end{enumerate} + When all $g$ in $gs$ have been evaluated in this manner, $f$ is evaluated on the local variables + by pushing @{term "(f, locals, [])"}. The resulting register value is kept + and the active frame for $h$ is popped off the stack. +\item For @{text "h = Pr n f g"}, let @{term "xs = y # ys"}. First @{term "(f, + ys, [])"} is pushed and the return value stored in the @{term + locals}. Then @{term "(g, x # v # ys, [])"} is pushed, + where $x$ is the length of @{term locals} and $v$ the most recently + appended value. The return value is appended to @{term locals}. This is + repeated until the length of @{term locals} reaches @{term y}. Then the most + recently appended local is placed in the register, and the stack is popped. +\item For @{text "h = Mn n f"}, frames @{term "(f, x # xs, [])"} are pushed + for $x = 0, 1, 2, \ldots$ until one of them returns $0$. Then this + $x$ is placed in the register and the stack is popped. Until then $x$ is + stored in @{term locals}. If none of these evaluations return $0$, the + stack never shrinks, and thus the machine never reaches a final state. +\end{itemize}\ + +type_synonym frame = "recf \ nat list \ nat list" + +type_synonym configuration = "frame list \ nat option" + + +subsubsection \Definition of the step function\ + +fun step :: "configuration \ configuration" where + "step ([], rv) = ([], rv)" +| "step (((Z, _, _) # fs), rv) = (fs, Some 0)" +| "step (((S, xs, _) # fs), rv) = (fs, Some (Suc (hd xs)))" +| "step (((Id m n, xs, _) # fs), rv) = (fs, Some (xs ! n))" +| "step (((Cn n f gs, xs, ls) # fs), rv) = + (if length ls = length gs + then if rv = None + then ((f, ls, []) # (Cn n f gs, xs, ls) # fs, None) + else (fs, rv) + else if rv = None + then if length ls < length gs + then ((gs ! (length ls), xs, []) # (Cn n f gs, xs, ls) # fs, None) + else (fs, rv) \\cannot occur, so don't-care term\ + else ((Cn n f gs, xs, ls @ [the rv]) # fs, None))" +| "step (((Pr n f g, xs, ls) # fs), rv) = + (if ls = [] + then if rv = None + then ((f, tl xs, []) # (Pr n f g, xs, ls) # fs, None) + else ((Pr n f g, xs, [the rv]) # fs, None) + else if length ls = Suc (hd xs) + then (fs, Some (hd ls)) + else if rv = None + then ((g, (length ls - 1) # hd ls # tl xs, []) # (Pr n f g, xs, ls) # fs, None) + else ((Pr n f g, xs, (the rv) # ls) # fs, None))" +| "step (((Mn n f, xs, ls) # fs), rv) = + (if ls = [] + then ((f, 0 # xs, []) # (Mn n f, xs, [0]) # fs, None) + else if rv = Some 0 + then (fs, Some (hd ls)) + else ((f, (Suc (hd ls)) # xs, []) # (Mn n f, xs, [Suc (hd ls)]) # fs, None))" + +definition reachable :: "configuration \ configuration \ bool" where + "reachable x y \ \t. iterate t step x = y" + +lemma step_reachable [intro]: + assumes "step x = y" + shows "reachable x y" + unfolding reachable_def using assms by (metis iterate.simps(1,2) comp_id) + +lemma reachable_transitive [trans]: + assumes "reachable x y" and "reachable y z" + shows "reachable x z" + using assms iterate_additive[where ?f=step] reachable_def by metis + +lemma reachable_refl: "reachable x x" + unfolding reachable_def by (metis iterate.simps(1) eq_id_iff) + +text \From a final configuration, that is, when the stack is empty, +only final configurations are reachable.\ + +lemma step_empty_stack: + assumes "fst x = []" + shows "fst (step x) = []" + using assms by (metis prod.collapse step.simps(1)) + +lemma reachable_empty_stack: + assumes "fst x = []" and "reachable x y" + shows "fst y = []" +proof - + have "fst (iterate t step x) = []" for t + using assms step_empty_stack by (induction t) simp_all + then show ?thesis + using reachable_def assms(2) by auto +qed + +abbreviation nonterminating :: "configuration \ bool" where + "nonterminating x \ \t. fst (iterate t step x) \ []" + +lemma reachable_nonterminating: + assumes "reachable x y" and "nonterminating y" + shows "nonterminating x" +proof - + from assms(1) obtain t\<^sub>1 where t1: "iterate t\<^sub>1 step x = y" + using reachable_def by auto + have "fst (iterate t step x) \ []" for t + proof (cases "t \ t\<^sub>1") + case True + then show ?thesis + using t1 assms(2) reachable_def reachable_empty_stack iterate_additive' + by (metis le_Suc_ex) + next + case False + then have "iterate t step x = iterate (t\<^sub>1 + (t - t\<^sub>1)) step x" + by simp + then have "iterate t step x = iterate (t - t\<^sub>1) step (iterate t\<^sub>1 step x)" + by (simp add: iterate_additive') + then have "iterate t step x = iterate (t - t\<^sub>1) step y" + using t1 by simp + then show "fst (iterate t step x) \ []" + using assms(2) by simp + qed + then show ?thesis .. +qed + +text \The function @{term step} is underdefined, for example, when the +top frame contains a non-well-formed @{typ recf} or too few arguments. All is +well, though, if every frame contains a well-formed @{typ recf} whose arity +matches the number of arguments. Such stacks will be called +\emph{valid}.\ + +definition valid :: "frame list \ bool" where + "valid stack \ \s\set stack. recfn (length (fst (snd s))) (fst s)" + +lemma valid_frame: "valid (s # ss) \ valid ss \ recfn (length (fst (snd s))) (fst s)" + using valid_def by simp + +lemma valid_ConsE: "valid ((f, xs, locs) # rest) \ valid rest \ recfn (length xs) f" + using valid_def by simp + +lemma valid_ConsI: "valid rest \ recfn (length xs) f \ valid ((f, xs, locs) # rest)" + using valid_def by simp + +text \Stacks in initial configurations are valid, and performing a step +maintains the validity of the stack.\ + +lemma step_valid: "valid stack \ valid (fst (step (stack, rv)))" +proof (cases stack) + case Nil + then show ?thesis using valid_def by simp +next + case (Cons s ss) + assume valid: "valid stack" + then have *: "valid ss \ recfn (length (fst (snd s))) (fst s)" + using valid_frame Cons by simp + show ?thesis + proof (cases "fst s") + case Z + then show ?thesis using Cons valid * by (metis fstI prod.collapse step.simps(2)) + next + case S + then show ?thesis using Cons valid * by (metis fst_conv prod.collapse step.simps(3)) + next + case Id + then show ?thesis using Cons valid * by (metis fstI prod.collapse step.simps(4)) + next + case (Cn n f gs) + then obtain xs ls where "s = (Cn n f gs, xs, ls)" + using Cons by (metis prod.collapse) + moreover consider + "length ls = length gs \ rv \" + | "length ls = length gs \ rv \" + | "length ls < length gs \ rv \" + | "length ls \ length gs \ rv \" + | "length ls > length gs \ rv \" + by linarith + ultimately show ?thesis using valid Cons valid_def by (cases) auto + next + case (Pr n f g) + then obtain xs ls where s: "s = (Pr n f g, xs, ls)" + using Cons by (metis prod.collapse) + consider + "length ls = 0 \ rv \" + | "length ls = 0 \ rv \" + | "length ls \ 0 \ length ls = Suc (hd xs)" + | "length ls \ 0 \ length ls \ Suc (hd xs) \ rv \" + | "length ls \ 0 \ length ls \ Suc (hd xs) \ rv \" + by linarith + then show ?thesis using Cons * valid_def s by (cases) auto + next + case (Mn n f) + then obtain xs ls where s: "s = (Mn n f, xs, ls)" + using Cons by (metis prod.collapse) + consider + "length ls = 0" + | "length ls \ 0 \ rv \" + | "length ls \ 0 \ rv \" + by linarith + then show ?thesis using Cons * valid_def s by (cases) auto + qed +qed + +corollary iterate_step_valid: + assumes "valid stack" + shows "valid (fst (iterate t step (stack, rv)))" + using assms +proof (induction t) + case 0 + then show ?case by simp +next + case (Suc t) + moreover have "iterate (Suc t) step (stack, rv) = step (iterate t step (stack, rv))" + by simp + ultimately show ?case using step_valid valid_def by (metis prod.collapse) +qed + + +subsubsection \Correctness of the step function\ + +text \The function @{term step} works correctly for a @{typ recf} $f$ +on arguments @{term xs} in some configuration if (1) in case $f$ converges, @{term +step} reaches a configuration with the topmost frame popped and @{term "eval +f xs"} in the register, and (2) in case $f$ diverges, @{term step} does not +reach a final configuration.\ + +fun correct :: "configuration \ bool" where + "correct ([], r) = True" +| "correct ((f, xs, ls) # rest, r) = + (if eval f xs \ then reachable ((f, xs, ls) # rest, r) (rest, eval f xs) + else nonterminating ((f, xs, ls) # rest, None))" + +lemma correct_convergI: + assumes "eval f xs \" and "reachable ((f, xs, ls) # rest, None) (rest, eval f xs)" + shows "correct ((f, xs, ls) # rest, None)" + using assms by auto + +lemma correct_convergE: + assumes "correct ((f, xs, ls) # rest, None)" and "eval f xs \" + shows "reachable ((f, xs, ls) # rest, None) (rest, eval f xs)" + using assms by simp + +text \The correctness proof for @{term step} is by structural induction +on the @{typ recf} in the top frame. The base cases @{term Z}, @{term S}, +and @{term[names_short] Id} are simple. For @{text "X = Cn, Pr, Mn"}, the +lemmas named @{text reachable_X} show which configurations are reachable for +@{typ recf}s of shape @{text X}. Building on those, the lemmas named @{text +step_X_correct} show @{term step}'s correctness for @{text X}.\ + +lemma reachable_Cn: + assumes "valid (((Cn n f gs), xs, []) # rest)" (is "valid ?stack") + and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" + and "\g xs rest. + g \ set gs \ valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" + and "\i" + and "k \ length gs" + shows "reachable + (?stack, None) + ((Cn n f gs, xs, take k (map (\g. the (eval g xs)) gs)) # rest, None)" + using assms(4,5) +proof (induction k) + case 0 + then show ?case using reachable_refl by simp +next + case (Suc k) + let ?ys = "map (\g. the (eval g xs)) gs" + from Suc have "k < length gs" by simp + have valid: "recfn (length xs) (Cn n f gs)" "valid rest" + using assms(1) valid_ConsE[of "(Cn n f gs)"] by simp_all + from Suc have "reachable (?stack, None) ((Cn n f gs, xs, take k ?ys) # rest, None)" + (is "_ (?stack1, None)") + by simp + also have "reachable ... ((gs ! k, xs, []) # ?stack1, None)" + using step_reachable `k < length gs` by (simp add: min_absorb2) + also have "reachable ... (?stack1, eval (gs ! k) xs)" + (is "_ (_, ?rv)") + using Suc.prems(1) \k < length gs\ assms(3) valid valid_ConsI by auto + also have "reachable ... ((Cn n f gs, xs, (take (Suc k) ?ys)) # rest, None)" + (is "_ (?stack2, None)") + proof - + have "step (?stack1, ?rv) = ((Cn n f gs, xs, (take k ?ys) @ [the ?rv]) # rest, None)" + using Suc by auto + also have "... = ((Cn n f gs, xs, (take (Suc k) ?ys)) # rest, None)" + by (simp add: \k < length gs\ take_Suc_conv_app_nth) + finally show ?thesis + using step_reachable by simp + qed + finally show "reachable (?stack, None) (?stack2, None)" . +qed + +lemma step_Cn_correct: + assumes "valid (((Cn n f gs), xs, []) # rest)" (is "valid ?stack") + and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" + and "\g xs rest. + g \ set gs \ valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" + shows "correct (?stack, None)" +proof - + have valid: "recfn (length xs) (Cn n f gs)" "valid rest" + using valid_ConsE[OF assms(1)] by auto + let ?ys = "map (\g. the (eval g xs)) gs" + consider + (diverg_f) "\g\set gs. eval g xs \" and "eval f ?ys \" + | (diverg_gs) "\g\set gs. eval g xs \" + | (converg) "eval (Cn n f gs) xs \" + using valid_ConsE[OF assms(1)] by fastforce + then show ?thesis + proof (cases) + case diverg_f + then have "\i" by simp + then have "reachable (?stack, None) ((Cn n f gs, xs, ?ys) # rest, None)" + (is "_ (?stack1, None)") + using reachable_Cn[OF assms, where ?k="length gs"] by simp + also have "reachable ... ((f, ?ys, []) # ?stack1, None)" (is "_ (?stack2, None)") + by (simp add: step_reachable) + finally have "reachable (?stack, None) (?stack2, None)" . + moreover have "nonterminating (?stack2, None)" + using diverg_f(2) assms(2)[of ?ys ?stack1] valid_ConsE[OF assms(1)] valid_ConsI + by auto + ultimately have "nonterminating (?stack, None)" + using reachable_nonterminating by simp + moreover have "eval (Cn n f gs) xs \" + using diverg_f(2) assms(1) eval_Cn valid_ConsE by presburger + ultimately show ?thesis by simp + next + case diverg_gs + then have ex_i: "\i" + using in_set_conv_nth[of _ gs] by auto + define k where "k = (LEAST i. i < length gs \ eval (gs ! i) xs \)" (is "_ = Least ?P") + then have gs_k: "eval (gs ! k) xs \" + using LeastI_ex[OF ex_i] by simp + have "\i" + using k_def not_less_Least[of _ ?P] LeastI_ex[OF ex_i] by simp + moreover from this have "k < length gs" + using ex_i less_le_trans not_le by blast + ultimately have "reachable (?stack, None) ((Cn n f gs, xs, take k ?ys) # rest, None)" + using reachable_Cn[OF assms] by simp + also have "reachable ... + ((gs ! (length (take k ?ys)), xs, []) # (Cn n f gs, xs, take k ?ys) # rest, None)" + (is "_ (?stack1, None)") + proof - + have "length (take k ?ys) < length gs" + by (simp add: \k < length gs\ less_imp_le_nat min_less_iff_disj) + then show ?thesis using step_reachable by simp + qed + finally have "reachable (?stack, None) (?stack1, None)" . + moreover have "nonterminating (?stack1, None)" + proof - + have "recfn (length xs) (gs ! k)" + using \k < length gs\ valid(1) by simp + then have "correct (?stack1, None)" + using \k < length gs\ nth_mem valid valid_ConsI + assms(3)[of "gs ! (length (take k ?ys))" xs] + by auto + moreover have "length (take k ?ys) = k" + by (simp add: \k < length gs\ less_imp_le_nat min_absorb2) + ultimately show ?thesis using gs_k by simp + qed + ultimately have "nonterminating (?stack, None)" + using reachable_nonterminating by simp + moreover have "eval (Cn n f gs) xs \" + using diverg_gs valid by fastforce + ultimately show ?thesis by simp + next + case converg + then have f: "eval f ?ys \" and g: "\g. g \ set gs \ eval g xs \" + using valid(1) by (metis eval_Cn)+ + then have "\i" + by simp + then have "reachable (?stack, None) ((Cn n f gs, xs, take (length gs) ?ys) # rest, None)" + using reachable_Cn assms by blast + also have "reachable ... ((Cn n f gs, xs, ?ys) # rest, None)" (is "_ (?stack1, None)") + by (simp add: reachable_refl) + also have "reachable ... ((f, ?ys, []) # ?stack1, None)" + using step_reachable by simp + also have "reachable ... (?stack1, eval f ?ys)" + using assms(2)[of "?ys"] correct_convergE valid f valid_ConsI by auto + also have "reachable (?stack1, eval f ?ys) (rest, eval f ?ys)" + using f by auto + finally have "reachable (?stack, None) (rest, eval f ?ys)" . + moreover have "eval (Cn n f gs) xs = eval f ?ys" + using g valid(1) by auto + ultimately show ?thesis + using converg correct_convergI by auto + qed +qed + +text \During the execution of a frame with a partial recursive function +of shape @{term "Pr n f g"} and arguments @{term "x # xs"}, the list of local +variables collects all the function values up to @{term x} in reversed +order. We call such a list a @{term trace} for short.\ + +definition trace :: "nat \ recf \ recf \ nat list \ nat \ nat list" where + "trace n f g xs x \ map (\y. the (eval (Pr n f g) (y # xs))) (rev [0..xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" + and "\xs rest. valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" + and "y \ x" + and "eval (Pr n f g) (y # xs) \" + shows "reachable (?stack, None) ((Pr n f g, x # xs, trace n f g xs y) # rest, None)" + using assms(4,5) +proof (induction y) + case 0 + have valid: "recfn (length (x # xs)) (Pr n f g)" "valid rest" + using valid_ConsE[OF assms(1)] by simp_all + then have f: "eval f xs \" using 0 by simp + let ?as = "x # xs" + have "reachable (?stack, None) ((f, xs, []) # ((Pr n f g), ?as, []) # rest, None)" + using step_reachable by simp + also have "reachable ... (?stack, eval f xs)" + using assms(2)[of xs "((Pr n f g), ?as, []) # rest"] + correct_convergE[OF _ f] f valid valid_ConsI + by simp + also have "reachable ... ((Pr n f g, ?as, [the (eval f xs)]) # rest, None)" + using step_reachable valid(1) f by simp + finally have "reachable (?stack, None) ((Pr n f g, ?as, [the (eval f xs)]) # rest, None)" . + then show ?case using trace_def valid(1) by simp +next + case (Suc y) + have valid: "recfn (length (x # xs)) (Pr n f g)" "valid rest" + using valid_ConsE[OF assms(1)] by simp_all + let ?ls = "trace n f g xs y" + have lenls: "length ?ls = Suc y" + using trace_length by auto + moreover have hdls: "hd ?ls = the (eval (Pr n f g) (y # xs))" + using Suc trace_hd by auto + ultimately have g: + "eval g (y # hd ?ls # xs) \" + "eval (Pr n f g) (Suc y # xs) = eval g (y # hd ?ls # xs)" + using eval_Pr_Suc_converg hdls valid(1) Suc by simp_all + then have "reachable (?stack, None) ((Pr n f g, x # xs, ?ls) # rest, None)" + (is "_ (?stack1, None)") + using Suc valid(1) by fastforce + also have "reachable ... ((g, y # hd ?ls # xs, []) # (Pr n f g, x # xs, ?ls) # rest, None)" + using Suc.prems lenls by fastforce + also have "reachable ... (?stack1, eval g (y # hd ?ls # xs))" + (is "_ (_, ?rv)") + using assms(3) g(1) valid valid_ConsI by auto + also have "reachable ... ((Pr n f g, x # xs, (the ?rv) # ?ls) # rest, None)" + using Suc.prems(1) g(1) lenls by auto + finally have "reachable (?stack, None) ((Pr n f g, x # xs, (the ?rv) # ?ls) # rest, None)" . + moreover have "trace n f g xs (Suc y) = (the ?rv) # ?ls" + using g(2) trace_Suc by simp + ultimately show ?case by simp +qed + +lemma step_Pr_correct: + assumes "valid (((Pr n f g), xs, []) # rest)" (is "valid ?stack") + and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" + and "\xs rest. valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" + shows "correct (?stack, None)" +proof - + have valid: "valid rest" "recfn (length xs) (Pr n f g)" + using valid_ConsE[OF assms(1)] by simp_all + then have "length xs > 0" + by auto + then obtain y ys where y_ys: "xs = y # ys" + using list.exhaust_sel by auto + let ?t = "trace n f g ys" + consider + (converg) "eval (Pr n f g) xs \" + | (diverg_f) "eval (Pr n f g) xs \" and "eval f ys \" + | (diverg) "eval (Pr n f g) xs \" and "eval f ys \" + by auto + then show ?thesis + proof (cases) + case converg + then have "\z. z \ y \ reachable (?stack, None) (((Pr n f g), xs, ?t z) # rest, None)" + using assms valid by (simp add: eval_Pr_converg_le reachable_Pr y_ys) + then have "reachable (?stack, None) (((Pr n f g), xs, ?t y) # rest, None)" + by simp + moreover have "reachable (((Pr n f g), xs, ?t y) # rest, None) (rest, Some (hd (?t y)))" + using trace_length step_reachable y_ys by fastforce + ultimately have "reachable (?stack, None) (rest, Some (hd (?t y)))" + using reachable_transitive by blast + then show ?thesis + using assms(1) trace_hd converg y_ys by simp + next + case diverg_f + have *: "step (?stack, None) = ((f, ys, []) # ((Pr n f g), xs, []) # tl ?stack, None)" + (is "_ = (?stack1, None)") + using assms(1,2) y_ys by simp + then have "reachable (?stack, None) (?stack1, None)" + using step_reachable by simp + moreover have "nonterminating (?stack1, None)" + using assms diverg_f valid valid_ConsI * by auto + ultimately have "nonterminating (?stack, None)" + using reachable_nonterminating by blast + then show ?thesis using diverg_f(1) assms(1) by simp + next + case diverg + let ?h = "\z. the (eval (Pr n f g) (z # ys))" + let ?Q = "\z. z < y \ eval (Pr n f g) (z # ys) \" + have "?Q 0" + using assms diverg neq0_conv y_ys valid by fastforce + define zmax where "zmax = Greatest ?Q" + then have "?Q zmax" + using `?Q 0` GreatestI_nat[of ?Q 0 y] by simp + have le_zmax: "\z. ?Q z \ z \ zmax" + using Greatest_le_nat[of ?Q _ y] zmax_def by simp + have len: "length (?t zmax) < Suc y" + by (simp add: \?Q zmax\ trace_length) + have "eval (Pr n f g) (y # ys) \" if "y \ zmax" for y + using that zmax_def `?Q zmax` assms eval_Pr_converg_le[of n f g ys zmax y] valid y_ys + by simp + then have "reachable (?stack, None) (((Pr n f g), xs, ?t y) # rest, None)" + if "y \ zmax" for y + using that `?Q zmax` diverg y_ys assms reachable_Pr by simp + then have "reachable (?stack, None) (((Pr n f g), xs, ?t zmax) # rest, None)" + (is "reachable _ (?stack1, None)") + by simp + also have "reachable ... + ((g, zmax # ?h zmax # tl xs, []) # (Pr n f g, xs, ?t zmax) # rest, None)" + (is "_ (?stack2, None)") + proof (rule step_reachable) + have "length (?t zmax) \ Suc (hd xs)" + using len y_ys by simp + moreover have "hd (?t zmax) = ?h zmax" + using trace_hd by auto + moreover have "length (?t zmax) = Suc zmax" + using trace_length by simp + ultimately show "step (?stack1, None) = (?stack2, None)" + by auto + qed + finally have "reachable (?stack, None) (?stack2, None)" . + moreover have "nonterminating (?stack2, None)" + proof - + have "correct (?stack2, None)" + using y_ys assms valid_ConsI valid by simp + moreover have "eval g (zmax # ?h zmax # ys) \" + using \?Q zmax\ diverg le_zmax len less_Suc_eq trace_length y_ys valid + by fastforce + ultimately show ?thesis using y_ys by simp + qed + ultimately have "nonterminating (?stack, None)" + using reachable_nonterminating by simp + then show ?thesis using diverg assms(1) by simp + qed +qed + +lemma reachable_Mn: + assumes "valid ((Mn n f, xs, []) # rest)" (is "valid ?stack") + and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" + and "\y {None, Some 0}" + shows "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" + using assms(3) +proof (induction z) + case 0 + then have "step (?stack, None) = ((f, 0 # xs, []) # (Mn n f, xs, [0]) # rest, None)" + using assms by simp + then show ?case + using step_reachable assms(1) by simp +next + case (Suc z) + have valid: "valid rest" "recfn (length xs) (Mn n f)" + using valid_ConsE[OF assms(1)] by auto + have f: "eval f (z # xs) \ {None, Some 0}" + using Suc by simp + have "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" + using Suc by simp + also have "reachable ... ((Mn n f, xs, [z]) # rest, eval f (z # xs))" + using f assms(2)[of "z # xs"] valid correct_convergE valid_ConsI by auto + also have "reachable ... ((f, (Suc z) # xs, []) # (Mn n f, xs, [Suc z]) # rest, None)" + (is "_ (?stack1, None)") + using step_reachable f by simp + finally have "reachable (?stack, None) (?stack1, None)" . + then show ?case by simp +qed + +lemma iterate_step_empty_stack: "iterate t step ([], rv) = ([], rv)" + using step_empty_stack by (induction t) simp_all + +lemma reachable_iterate_step_empty_stack: + assumes "reachable cfg ([], rv)" + shows "\t. iterate t step cfg = ([], rv) \ (\t' [])" +proof - + let ?P = "\t. iterate t step cfg = ([], rv)" + from assms have "\t. ?P t" + by (simp add: reachable_def) + moreover define tmin where "tmin = Least ?P" + ultimately have "?P tmin" + using LeastI_ex[of ?P] by simp + have "fst (iterate t' step cfg) \ []" if "t' < tmin" for t' + proof + assume "fst (iterate t' step cfg) = []" + then obtain v where v: "iterate t' step cfg = ([], v)" + by (metis prod.exhaust_sel) + then have "iterate t'' step ([], v) = ([], v)" for t'' + using iterate_step_empty_stack by simp + then have "iterate (t' + t'') step cfg = ([], v)" for t'' + using v iterate_additive by fast + moreover obtain t'' where "t' + t'' = tmin" + using \t' < tmin\ less_imp_add_positive by auto + ultimately have "iterate tmin step cfg = ([], v)" + by auto + then have "v = rv" + using `?P tmin` by simp + then have "iterate t' step cfg = ([], rv)" + using v by simp + moreover have "\t' ?P t'" + unfolding tmin_def using not_less_Least[of _ ?P] by simp + ultimately show False + using that by simp + qed + then show ?thesis using `?P tmin` by auto +qed + +lemma step_Mn_correct: + assumes "valid ((Mn n f, xs, []) # rest)" (is "valid ?stack") + and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" + shows "correct (?stack, None)" +proof - + have valid: "valid rest" "recfn (length xs) (Mn n f)" + using valid_ConsE[OF assms(1)] by auto + consider + (diverg) "eval (Mn n f) xs \" and "\z. eval f (z # xs) \" + | (diverg_f) "eval (Mn n f) xs \" and "\z. eval f (z # xs) \" + | (converg) "eval (Mn n f) xs \" + by fast + then show ?thesis + proof (cases) + case diverg + then have "\z. eval f (z # xs) \ Some 0" + using eval_Mn_diverg[OF valid(2)] by simp + then have "\y {None, Some 0}" for z + using diverg by simp + then have reach_z: + "\z. reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" + using reachable_Mn[OF assms] diverg by simp + + define h :: "nat \ configuration" where + "h z \ ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" for z + then have h_inj: "\x y. x \ y \ h x \ h y" and z_neq_Nil: "\z. fst (h z) \ []" + by simp_all + + have z: "\z\<^sub>0. \z>z\<^sub>0. \ (\t'\t. iterate t' step (?stack, None) = h z)" for t + proof (induction t) + case 0 + then show ?case by (metis h_inj le_zero_eq less_not_refl3) + next + case (Suc t) + then show ?case + using h_inj by (metis (no_types, hide_lams) le_Suc_eq less_not_refl3 less_trans) + qed + + have "nonterminating (?stack, None)" + proof (rule ccontr) + assume "\ nonterminating (?stack, None)" + then obtain t where t: "fst (iterate t step (?stack, None)) = []" + by auto + then obtain z\<^sub>0 where "\z>z\<^sub>0. \ (\t'\t. iterate t' step (?stack, None) = h z)" + using z by auto + then have not_h: "\t'\t. iterate t' step (?stack, None) \ h (Suc z\<^sub>0)" + by simp + have "\t'\t. fst (iterate t' step (?stack, None)) = []" + using t iterate_step_empty_stack iterate_additive'[of t] + by (metis le_Suc_ex prod.exhaust_sel) + then have "\t'\t. iterate t' step (?stack, None) \ h (Suc z\<^sub>0)" + using z_neq_Nil by auto + then have "\t'. iterate t' step (?stack, None) \ h (Suc z\<^sub>0)" + using not_h nat_le_linear by auto + then have "\ reachable (?stack, None) (h (Suc z\<^sub>0))" + using reachable_def by simp + then show False + using reach_z[of "Suc z\<^sub>0"] h_def by simp + qed + then show ?thesis using diverg by simp + next + case diverg_f + let ?P = "\z. eval f (z # xs) \" + define zmin where "zmin \ Least ?P" + then have "\y {None, Some 0}" + using diverg_f eval_Mn_diverg[OF valid(2)] less_trans not_less_Least[of _ ?P] + by blast + moreover have f_zmin: "eval f (zmin # xs) \" + using diverg_f LeastI_ex[of ?P] zmin_def by simp + ultimately have + "reachable (?stack, None) ((f, zmin # xs, []) # (Mn n f, xs, [zmin]) # rest, None)" + (is "reachable _ (?stack1, None)") + using reachable_Mn[OF assms] by simp + moreover have "nonterminating (?stack1, None)" + using f_zmin assms valid diverg_f valid_ConsI by auto + ultimately have "nonterminating (?stack, None)" + using reachable_nonterminating by simp + then show ?thesis using diverg_f by simp + next + case converg + then obtain z where z: "eval (Mn n f) xs \= z" by auto + have f_z: "eval f (z # xs) \= 0" + and f_less_z: "\y. y < z \ eval f (y # xs) \\ 0" + using eval_Mn_convergE(2,3)[OF valid(2) z] by simp_all + then have + "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" + using reachable_Mn[OF assms] by simp + also have "reachable ... ((Mn n f, xs, [z]) # rest, eval f (z # xs))" + using assms(2)[of "z # xs"] valid f_z valid_ConsI correct_convergE + by auto + also have "reachable ... (rest, Some z)" + using f_z f_less_z step_reachable by simp + finally have "reachable (?stack, None) (rest, Some z)" . + then show ?thesis using z by simp + qed +qed + +theorem step_correct: + assumes "valid ((f, xs, []) # rest)" + shows "correct ((f, xs, []) # rest, None)" + using assms +proof (induction f arbitrary: xs rest) + case Z + then show ?case using valid_ConsE[of Z] step_reachable by simp +next + case S + then show ?case using valid_ConsE[of S] step_reachable by simp +next + case (Id m n) + then show ?case using valid_ConsE[of "Id m n"] by auto +next + case Cn + then show ?case using step_Cn_correct by presburger +next + case Pr + then show ?case using step_Pr_correct by simp +next + case Mn + then show ?case using step_Mn_correct by presburger +qed + + +subsection \Encoding partial recursive functions\label{s:recf_enc}\ + +text \In this section we define an injective, but not surjective, +mapping from @{typ recf}s to natural numbers.\ + +abbreviation triple_encode :: "nat \ nat \ nat \ nat" where + "triple_encode x y z \ prod_encode (x, prod_encode (y, z))" + +abbreviation quad_encode :: "nat \ nat \ nat \ nat \ nat" where + "quad_encode w x y z \ prod_encode (w, prod_encode (x, prod_encode (y, z)))" + +fun encode :: "recf \ nat" where + "encode Z = 0" +| "encode S = 1" +| "encode (Id m n) = triple_encode 2 m n" +| "encode (Cn n f gs) = quad_encode 3 n (encode f) (list_encode (map encode gs))" +| "encode (Pr n f g) = quad_encode 4 n (encode f) (encode g)" +| "encode (Mn n f) = triple_encode 5 n (encode f)" + +lemma prod_encode_gr1: "a > 1 \ prod_encode (a, x) > 1" + using le_prod_encode_1 less_le_trans by blast + +lemma encode_not_Z_or_S: "encode f = prod_encode (a, b) \ a > 1 \ f \ Z \ f \ S" + by (metis encode.simps(1) encode.simps(2) less_numeral_extra(4) not_one_less_zero + prod_encode_gr1) + +lemma encode_injective: "encode f = encode g \ f = g" +proof (induction g arbitrary: f) + case Z + have "\a x. a > 1 \ prod_encode (a, x) > 0" + using prod_encode_gr1 by (meson less_one less_trans) + then have "f \ Z \ encode f > 0" + by (cases f) auto + then have "encode f = 0 \ f = Z" by fastforce + then show ?case using Z by simp +next + case S + have "\a x. a > 1 \ prod_encode (a, x) \ Suc 0" + using prod_encode_gr1 by (metis One_nat_def less_numeral_extra(4)) + then have "encode f = 1 \ f = S" + by (cases f) auto + then show ?case using S by simp +next + case Id + then obtain z where *: "encode f = prod_encode (2, z)" by simp + show ?case + using Id by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq) +next + case Cn + then obtain z where *: "encode f = prod_encode (3, z)" by simp + show ?case + proof (cases f) + case Z + then show ?thesis using * encode_not_Z_or_S by simp + next + case S + then show ?thesis using * encode_not_Z_or_S by simp + next + case Id + then show ?thesis using * by (simp add: prod_encode_eq) + next + case Cn + then show ?thesis + using * Cn.IH Cn.prems list_decode_encode + by (smt encode.simps(4) fst_conv list.inj_map_strong prod_encode_eq snd_conv) + next + case Pr + then show ?thesis using * by (simp add: prod_encode_eq) + next + case Mn + then show ?thesis using * by (simp add: prod_encode_eq) + qed +next + case Pr + then obtain z where *: "encode f = prod_encode (4, z)" by simp + show ?case + using Pr by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq) +next + case Mn + then obtain z where *: "encode f = prod_encode (5, z)" by simp + show ?case + using Mn by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq) +qed + +definition encode_kind :: "nat \ nat" where + "encode_kind e \ if e = 0 then 0 else if e = 1 then 1 else pdec1 e" + +lemma encode_kind_0: "encode_kind (encode Z) = 0" + unfolding encode_kind_def by simp + +lemma encode_kind_1: "encode_kind (encode S) = 1" + unfolding encode_kind_def by simp + +lemma encode_kind_2: "encode_kind (encode (Id m n)) = 2" + unfolding encode_kind_def + by (metis encode.simps(1-3) encode_injective fst_conv prod_encode_inverse + recf.simps(16) recf.simps(8)) + +lemma encode_kind_3: "encode_kind (encode (Cn n f gs)) = 3" + unfolding encode_kind_def + by (metis encode.simps(1,2,4) encode_injective fst_conv prod_encode_inverse + recf.simps(10) recf.simps(18)) + +lemma encode_kind_4: "encode_kind (encode (Pr n f g)) = 4" + unfolding encode_kind_def + by (metis encode.simps(1,2,5) encode_injective fst_conv prod_encode_inverse + recf.simps(12) recf.simps(20)) + +lemma encode_kind_5: "encode_kind (encode (Mn n f)) = 5" + unfolding encode_kind_def + by (metis encode.simps(1,2,6) encode_injective fst_conv prod_encode_inverse + recf.simps(14) recf.simps(22)) + +lemmas encode_kind_n = + encode_kind_0 encode_kind_1 encode_kind_2 encode_kind_3 encode_kind_4 encode_kind_5 + +lemma encode_kind_Cn: + assumes "encode_kind (encode f) = 3" + shows "\n f' gs. f = Cn n f' gs" + using assms encode_kind_n by (cases f) auto + +lemma encode_kind_Pr: + assumes "encode_kind (encode f) = 4" + shows "\n f' g. f = Pr n f' g" + using assms encode_kind_n by (cases f) auto + +lemma encode_kind_Mn: + assumes "encode_kind (encode f) = 5" + shows "\n g. f = Mn n g" + using assms encode_kind_n by (cases f) auto + +lemma pdec2_encode_Id: "pdec2 (encode (Id m n)) = prod_encode (m, n)" + by simp + +lemma pdec2_encode_Pr: "pdec2 (encode (Pr n f g)) = triple_encode n (encode f) (encode g)" + by simp + + +subsection \The step function on encoded configurations\label{s:step_enc}\ + +text \In this section we construct a function @{text "estep :: nat +\ nat"} that is equivalent to the function @{text "step :: +configuration \ configuration"} except that it applies to encoded +configurations. We start by defining an encoding for configurations.\ + +definition encode_frame :: "frame \ nat" where + "encode_frame s \ + triple_encode (encode (fst s)) (list_encode (fst (snd s))) (list_encode (snd (snd s)))" + +lemma encode_frame: + "encode_frame (f, xs, ls) = triple_encode (encode f) (list_encode xs) (list_encode ls)" + unfolding encode_frame_def by simp + +abbreviation encode_option :: "nat option \ nat" where + "encode_option x \ if x = None then 0 else Suc (the x)" + +definition encode_config :: "configuration \ nat" where + "encode_config cfg \ + prod_encode (list_encode (map encode_frame (fst cfg)), encode_option (snd cfg))" + +lemma encode_config: + "encode_config (ss, rv) = prod_encode (list_encode (map encode_frame ss), encode_option rv)" + unfolding encode_config_def by simp + +text \Various projections from encoded configurations:\ + +definition e2stack where "e2stack e \ pdec1 e" +definition e2rv where "e2rv e \ pdec2 e" +definition e2tail where "e2tail e \ e_tl (e2stack e)" +definition e2frame where "e2frame e \ e_hd (e2stack e)" +definition e2i where "e2i e \ pdec1 (e2frame e)" +definition e2xs where "e2xs e \ pdec12 (e2frame e)" +definition e2ls where "e2ls e \ pdec22 (e2frame e)" +definition e2lenas where "e2lenas e \ e_length (e2xs e)" +definition e2lenls where "e2lenls e \ e_length (e2ls e)" + +lemma e2rv_rv [simp]: + "e2rv (encode_config (ss, rv)) = (if rv \ then 0 else Suc (the rv))" + unfolding e2rv_def using encode_config by simp + +lemma e2stack_stack [simp]: + "e2stack (encode_config (ss, rv)) = list_encode (map encode_frame ss)" + unfolding e2stack_def using encode_config by simp + +lemma e2tail_tail [simp]: + "e2tail (encode_config (s # ss, rv)) = list_encode (map encode_frame ss)" + unfolding e2tail_def using encode_config by fastforce + +lemma e2frame_frame [simp]: + "e2frame (encode_config (s # ss, rv)) = encode_frame s" + unfolding e2frame_def using encode_config by fastforce + +lemma e2i_f [simp]: + "e2i (encode_config ((f, xs, ls) # ss, rv)) = encode f" + unfolding e2i_def using encode_config e2frame_frame encode_frame by force + +lemma e2xs_xs [simp]: + "e2xs (encode_config ((f, xs, ls) # ss, rv)) = list_encode xs" + using e2xs_def e2frame_frame encode_frame by force + +lemma e2ls_ls [simp]: + "e2ls (encode_config ((f, xs, ls) # ss, rv)) = list_encode ls" + using e2ls_def e2frame_frame encode_frame by force + +lemma e2lenas_lenas [simp]: + "e2lenas (encode_config ((f, xs, ls) # ss, rv)) = length xs" + using e2lenas_def e2frame_frame encode_frame by simp + +lemma e2lenls_lenls [simp]: + "e2lenls (encode_config ((f, xs, ls) # ss, rv)) = length ls" + using e2lenls_def e2frame_frame encode_frame by simp + +lemma e2stack_0_iff_Nil: + assumes "e = encode_config (ss, rv)" + shows "e2stack e = 0 \ ss = []" + using assms + by (metis list_encode.simps(1) e2stack_stack list_encode_0 map_is_Nil_conv) + +lemma e2ls_0_iff_Nil [simp]: "list_decode (e2ls e) = [] \ e2ls e = 0" + by (metis list_decode.simps(1) list_encode_decode) + +text \We now define @{text eterm} piecemeal by considering the more +complicated cases @{text Cn}, @{text Pr}, and @{text Mn} separately.\ + +definition "estep_Cn e \ + if e2lenls e = e_length (pdec222 (e2i e)) + then if e2rv e = 0 + then prod_encode (e_cons (triple_encode (pdec122 (e2i e)) (e2ls e) 0) (e2stack e), 0) + else prod_encode (e2tail e, e2rv e) + else if e2rv e = 0 + then if e2lenls e < e_length (pdec222 (e2i e)) + then prod_encode + (e_cons + (triple_encode (e_nth (pdec222 (e2i e)) (e2lenls e)) (e2xs e) 0) + (e2stack e), + 0) + else prod_encode (e2tail e, e2rv e) + else prod_encode + (e_cons + (triple_encode (e2i e) (e2xs e) (e_snoc (e2ls e) (e2rv e - 1))) + (e2tail e), + 0)" + +lemma estep_Cn: + assumes "c = (((Cn n f gs, xs, ls) # fs), rv)" + shows "estep_Cn (encode_config c) = encode_config (step c)" + using encode_frame by (simp add: assms estep_Cn_def, simp add: encode_config assms) + +definition "estep_Pr e \ + if e2ls e = 0 + then if e2rv e = 0 + then prod_encode + (e_cons (triple_encode (pdec122 (e2i e)) (e_tl (e2xs e)) 0) (e2stack e), + 0) + else prod_encode + (e_cons (triple_encode (e2i e) (e2xs e) (singleton_encode (e2rv e - 1))) (e2tail e), + 0) + else if e2lenls e = Suc (e_hd (e2xs e)) + then prod_encode (e2tail e, Suc (e_hd (e2ls e))) + else if e2rv e = 0 + then prod_encode + (e_cons + (triple_encode + (pdec222 (e2i e)) + (e_cons (e2lenls e - 1) (e_cons (e_hd (e2ls e)) (e_tl (e2xs e)))) + 0) + (e2stack e), + 0) + else prod_encode + (e_cons + (triple_encode (e2i e) (e2xs e) (e_cons (e2rv e - 1) (e2ls e))) (e2tail e), + 0)" + +lemma estep_Pr1: + assumes "c = (((Pr n f g, xs, ls) # fs), rv)" + and "ls \ []" + and "length ls \ Suc (hd xs)" + and "rv \ None" + and "recfn (length xs) (Pr n f g)" + shows "estep_Pr (encode_config c) = encode_config (step c)" +proof - + let ?e = "encode_config c" + from assms(5) have "length xs > 0" by auto + then have eq: "hd xs = e_hd (e2xs ?e)" + using assms e_hd_def by auto + have "step c = ((Pr n f g, xs, (the rv) # ls) # fs, None)" + (is "step c = (?t # ?ss, None)") + using assms by simp + then have "encode_config (step c) = + prod_encode (list_encode (map encode_frame (?t # ?ss)), 0)" + using encode_config by simp + also have "... = + prod_encode (e_cons (encode_frame ?t) (list_encode (map encode_frame (?ss))), 0)" + by simp + also have "... = prod_encode (e_cons (encode_frame ?t) (e2tail ?e), 0)" + using assms(1) by simp + also have "... = prod_encode + (e_cons + (triple_encode (e2i ?e) (e2xs ?e) (e_cons (e2rv ?e - 1) (e2ls ?e))) + (e2tail ?e), + 0)" + by (simp add: assms encode_frame) + finally show ?thesis + using assms eq estep_Pr_def by auto +qed + +lemma estep_Pr2: + assumes "c = (((Pr n f g, xs, ls) # fs), rv)" + and "ls \ []" + and "length ls \ Suc (hd xs)" + and "rv = None" + and "recfn (length xs) (Pr n f g)" + shows "estep_Pr (encode_config c) = encode_config (step c)" +proof - + let ?e = "encode_config c" + from assms(5) have "length xs > 0" by auto + then have eq: "hd xs = e_hd (e2xs ?e)" + using assms e_hd_def by auto + have "step c = ((g, (length ls - 1) # hd ls # tl xs, []) # (Pr n f g, xs, ls) # fs, None)" + (is "step c = (?t # ?ss, None)") + using assms by simp + then have "encode_config (step c) = + prod_encode (list_encode (map encode_frame (?t # ?ss)), 0)" + using encode_config by simp + also have "... = + prod_encode (e_cons (encode_frame ?t) (list_encode (map encode_frame (?ss))), 0)" + by simp + also have "... = prod_encode (e_cons (encode_frame ?t) (e2stack ?e), 0)" + using assms(1) by simp + also have "... = prod_encode + (e_cons + (triple_encode + (pdec222 (e2i ?e)) + (e_cons (e2lenls ?e - 1) (e_cons (e_hd (e2ls ?e)) (e_tl (e2xs ?e)))) + 0) + (e2stack ?e), + 0)" + using assms(1,2) encode_frame[of g "(length ls - 1) # hd ls # tl xs" "[]"] + pdec2_encode_Pr[of n f g] e2xs_xs e2i_f e2lenls_lenls e2ls_ls e_hd + by (metis list_encode.simps(1) list.collapse list_decode_encode + prod_encode_inverse snd_conv) + finally show ?thesis + using assms eq estep_Pr_def by auto +qed + +lemma estep_Pr3: + assumes "c = (((Pr n f g, xs, ls) # fs), rv)" + and "ls \ []" + and "length ls = Suc (hd xs)" + and "recfn (length xs) (Pr n f g)" + shows "estep_Pr (encode_config c) = encode_config (step c)" +proof - + let ?e = "encode_config c" + from assms(4) have "length xs > 0" by auto + then have "hd xs = e_hd (e2xs ?e)" + using assms e_hd_def by auto + then have "(length ls = Suc (hd xs)) = (e2lenls ?e = Suc (e_hd (e2xs ?e)))" + using assms by simp + then have *: "estep_Pr ?e = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" + using assms estep_Pr_def by auto + have "step c = (fs, Some (hd ls))" + using assms(1,2,3) by simp + then have "encode_config (step c) = + prod_encode (list_encode (map encode_frame fs), encode_option (Some (hd ls)))" + using encode_config by simp + also have "... = + prod_encode (list_encode (map encode_frame fs), encode_option (Some (e_hd (e2ls ?e))))" + using assms(1,2) e_hd_def by auto + also have "... = prod_encode (list_encode (map encode_frame fs), Suc (e_hd (e2ls ?e)))" + by simp + also have "... = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" + using assms(1) by simp + finally have "encode_config (step c) = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" . + then show ?thesis + using estep_Pr_def * by presburger +qed + +lemma estep_Pr4: + assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "ls = []" + shows "estep_Pr (encode_config c) = encode_config (step c)" + using encode_frame + by (simp add: assms estep_Pr_def, simp add: encode_config assms) + +lemma estep_Pr: + assumes "c = (((Pr n f g, xs, ls) # fs), rv)" + and "recfn (length xs) (Pr n f g)" + shows "estep_Pr (encode_config c) = encode_config (step c)" + using assms estep_Pr1 estep_Pr2 estep_Pr3 estep_Pr4 by simp + +definition "estep_Mn e \ + if e2ls e = 0 + then prod_encode + (e_cons + (triple_encode (pdec22 (e2i e)) (e_cons 0 (e2xs e)) 0) + (e_cons + (triple_encode (e2i e) (e2xs e) (singleton_encode 0)) + (e2tail e)), + 0) + else if e2rv e = 1 + then prod_encode (e2tail e, Suc (e_hd (e2ls e))) + else prod_encode + (e_cons + (triple_encode (pdec22 (e2i e)) (e_cons (Suc (e_hd (e2ls e))) (e2xs e)) 0) + (e_cons + (triple_encode (e2i e) (e2xs e) (singleton_encode (Suc (e_hd (e2ls e))))) + (e2tail e)), + 0)" + +lemma estep_Mn: + assumes "c = (((Mn n f, xs, ls) # fs), rv)" + shows "estep_Mn (encode_config c) = encode_config (step c)" +proof - + let ?e = "encode_config c" + consider "ls \ []" and "rv \ Some 0" | "ls \ []" and "rv = Some 0" | "ls = []" + by auto + then show ?thesis + proof (cases) + case 1 + then have step_c: "step c = + ((f, (Suc (hd ls)) # xs, []) # (Mn n f, xs, [Suc (hd ls)]) # fs, None)" + (is "step c = ?cfg") + using assms by simp + have "estep_Mn ?e = + prod_encode + (e_cons + (triple_encode (encode f) (e_cons (Suc (hd ls)) (list_encode xs)) 0) + (e_cons + (triple_encode (encode (Mn n f)) (list_encode xs) (singleton_encode (Suc (hd ls)))) + (list_encode (map encode_frame fs))), + 0)" + using 1 assms e_hd_def estep_Mn_def by auto + also have "... = encode_config ?cfg" + using encode_config by (simp add: encode_frame) + finally show ?thesis + using step_c by simp + next + case 2 + have "estep_Mn ?e = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" + using 2 assms estep_Mn_def by auto + also have "... = prod_encode (e2tail ?e, Suc (hd ls))" + using 2 assms e_hd_def by auto + also have "... = prod_encode (list_encode (map encode_frame fs), Suc (hd ls))" + using assms by simp + also have "... = encode_config (fs, Some (hd ls))" + using encode_config by simp + finally show ?thesis + using 2 assms by simp + next + case 3 + then show ?thesis + using assms encode_frame by (simp add: estep_Mn_def, simp add: encode_config) + qed +qed + +definition "estep e \ + if e2stack e = 0 then prod_encode (0, e2rv e) + else if e2i e = 0 then prod_encode (e2tail e, 1) + else if e2i e = 1 then prod_encode (e2tail e, Suc (Suc (e_hd (e2xs e)))) + else if encode_kind (e2i e) = 2 then + prod_encode (e2tail e, Suc (e_nth (e2xs e) (pdec22 (e2i e)))) + else if encode_kind (e2i e) = 3 then estep_Cn e + else if encode_kind (e2i e) = 4 then estep_Pr e + else if encode_kind (e2i e) = 5 then estep_Mn e + else 0" + +lemma estep_Z: + assumes "c = (((Z, xs, ls) # fs), rv)" + shows "estep (encode_config c) = encode_config (step c)" + using encode_frame by (simp add: assms estep_def, simp add: encode_config assms) + +lemma estep_S: + assumes "c = (((S, xs, ls) # fs), rv)" + and "recfn (length xs) (fst (hd (fst c)))" + shows "estep (encode_config c) = encode_config (step c)" +proof - + let ?e = "encode_config c" + from assms have "length xs > 0" by auto + then have eq: "hd xs = e_hd (e2xs ?e)" + using assms(1) e_hd_def by auto + then have "estep ?e = prod_encode (e2tail ?e, Suc (Suc (e_hd (e2xs ?e))))" + using assms(1) estep_def by simp + moreover have "step c = (fs, Some (Suc (hd xs)))" + using assms(1) by simp + ultimately show ?thesis + using assms(1) eq estep_def encode_config[of fs "Some (Suc (hd xs))"] by simp +qed + +lemma estep_Id: + assumes "c = (((Id m n, xs, ls) # fs), rv)" + and "recfn (length xs) (fst (hd (fst c)))" + shows "estep (encode_config c) = encode_config (step c)" +proof - + let ?e = "encode_config c" + from assms have "length xs = m" and "m > 0" by auto + then have eq: "xs ! n = e_nth (e2xs ?e) n" + using assms e_hd_def by auto + moreover have "encode_kind (e2i ?e) = 2" + using assms(1) encode_kind_2 by auto + ultimately have "estep ?e = + prod_encode (e2tail ?e, Suc (e_nth (e2xs ?e) (pdec22 (e2i ?e))))" + using assms estep_def encode_kind_def by auto + moreover have "step c = (fs, Some (xs ! n))" + using assms(1) by simp + ultimately show ?thesis + using assms(1) eq encode_config[of fs "Some (xs ! n)"] by simp +qed + +lemma estep: + assumes "valid (fst c)" + shows "estep (encode_config c) = encode_config (step c)" +proof (cases "fst c") + case Nil + then show ?thesis + using estep_def + by (metis list_encode.simps(1) e2rv_def e2stack_stack encode_config_def + map_is_Nil_conv prod.collapse prod_encode_inverse snd_conv step.simps(1)) +next + case (Cons s fs) + then obtain f xs ls rv where c: "c = ((f, xs, ls) # fs, rv)" + by (metis prod.exhaust_sel) + with assms valid_def have lenas: "recfn (length xs) f" by simp + show ?thesis + proof (cases f) + case Z + then show ?thesis using estep_Z c by simp + next + case S + then show ?thesis using estep_S c lenas by simp + next + case Id + then show ?thesis using estep_Id c lenas by simp + next + case Cn + then show ?thesis + using estep_Cn c + by (metis e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2 + encode_kind_3 encode_kind_Cn estep_def list.distinct(1) recf.distinct(13) + recf.distinct(19) recf.distinct(5)) + next + case Pr + then show ?thesis + using estep_Pr c lenas + by (metis e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2 + encode_kind_4 encode_kind_Cn encode_kind_Pr estep_def list.distinct(1) recf.distinct(15) + recf.distinct(21) recf.distinct(25) recf.distinct(7)) + next + case Mn + then show ?thesis + using estep_Pr c lenas + by (metis (no_types, lifting) e2i_f e2stack_0_iff_Nil encode.simps(1) + encode.simps(2) encode_kind_2 encode_kind_5 encode_kind_Cn encode_kind_Mn encode_kind_Pr + estep_Mn estep_def list.distinct(1) recf.distinct(17) recf.distinct(23) + recf.distinct(27) recf.distinct(9)) + qed +qed + +subsection \The step function as a partial recursive function\label{s:step_recf}\ + +text \In this section we construct a primitive recursive function +@{term r_step} computing @{term estep}. This will entail defining @{typ +recf}s for many functions defined in the previous section.\ + +definition "r_e2stack \ r_pdec1" + +lemma r_e2stack_prim: "prim_recfn 1 r_e2stack" + unfolding r_e2stack_def using r_pdec1_prim by simp + +lemma r_e2stack [simp]: "eval r_e2stack [e] \= e2stack e" + unfolding r_e2stack_def e2stack_def using r_pdec1_prim by simp + +definition "r_e2rv \ r_pdec2" + +lemma r_e2rv_prim: "prim_recfn 1 r_e2rv" + unfolding r_e2rv_def using r_pdec2_prim by simp + +lemma r_e2rv [simp]: "eval r_e2rv [e] \= e2rv e" + unfolding r_e2rv_def e2rv_def using r_pdec2_prim by simp + +definition "r_e2tail \ Cn 1 r_tl [r_e2stack]" + +lemma r_e2tail_prim: "prim_recfn 1 r_e2tail" + unfolding r_e2tail_def using r_e2stack_prim r_tl_prim by simp + +lemma r_e2tail [simp]: "eval r_e2tail [e] \= e2tail e" + unfolding r_e2tail_def e2tail_def using r_e2stack_prim r_tl_prim by simp + +definition "r_e2frame \ Cn 1 r_hd [r_e2stack]" + +lemma r_e2frame_prim: "prim_recfn 1 r_e2frame" + unfolding r_e2frame_def using r_hd_prim r_e2stack_prim by simp + +lemma r_e2frame [simp]: "eval r_e2frame [e] \= e2frame e" + unfolding r_e2frame_def e2frame_def using r_hd_prim r_e2stack_prim by simp + +definition "r_e2i \ Cn 1 r_pdec1 [r_e2frame]" + +lemma r_e2i_prim: "prim_recfn 1 r_e2i" + unfolding r_e2i_def using r_pdec12_prim r_e2frame_prim by simp + +lemma r_e2i [simp]: "eval r_e2i [e] \= e2i e" + unfolding r_e2i_def e2i_def using r_pdec12_prim r_e2frame_prim by simp + +definition "r_e2xs \ Cn 1 r_pdec12 [r_e2frame]" + +lemma r_e2xs_prim: "prim_recfn 1 r_e2xs" + unfolding r_e2xs_def using r_pdec122_prim r_e2frame_prim by simp + +lemma r_e2xs [simp]: "eval r_e2xs [e] \= e2xs e" + unfolding r_e2xs_def e2xs_def using r_pdec122_prim r_e2frame_prim by simp + +definition "r_e2ls \ Cn 1 r_pdec22 [r_e2frame]" + +lemma r_e2ls_prim: "prim_recfn 1 r_e2ls" + unfolding r_e2ls_def using r_pdec222_prim r_e2frame_prim by simp + +lemma r_e2ls [simp]: "eval r_e2ls [e] \= e2ls e" + unfolding r_e2ls_def e2ls_def using r_pdec222_prim r_e2frame_prim by simp + +definition "r_e2lenls \ Cn 1 r_length [r_e2ls]" + +lemma r_e2lenls_prim: "prim_recfn 1 r_e2lenls" + unfolding r_e2lenls_def using r_length_prim r_e2ls_prim by simp + +lemma r_e2lenls [simp]: "eval r_e2lenls [e] \= e2lenls e" + unfolding r_e2lenls_def e2lenls_def using r_length_prim r_e2ls_prim by simp + +definition "r_kind \ + Cn 1 r_ifz [Id 1 0, Z, Cn 1 r_ifeq [Id 1 0, r_const 1, r_const 1, r_pdec1]]" + +lemma r_kind_prim: "prim_recfn 1 r_kind" + unfolding r_kind_def by simp + +lemma r_kind: "eval r_kind [e] \= encode_kind e" + unfolding r_kind_def encode_kind_def by simp + +lemmas helpers_for_r_step_prim = + r_e2i_prim + r_e2lenls_prim + r_e2ls_prim + r_e2rv_prim + r_e2xs_prim + r_e2stack_prim + r_e2tail_prim + r_e2frame_prim + +text \We define primitive recursive functions @{term r_step_Id}, @{term +r_step_Cn}, @{term r_step_Pr}, and @{term r_step_Mn}. The last three +correspond to @{term estep_Cn}, @{term estep_Pr}, and @{term estep_Mn} from +the previous section.\ + +definition "r_step_Id \ + Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_nth [r_e2xs, Cn 1 r_pdec22 [r_e2i]]]]" + +lemma r_step_Id: + "eval r_step_Id [e] \= prod_encode (e2tail e, Suc (e_nth (e2xs e) (pdec22 (e2i e))))" + unfolding r_step_Id_def using helpers_for_r_step_prim by simp + +abbreviation r_triple_encode :: "recf \ recf \ recf \ recf" where + "r_triple_encode x y z \ Cn 1 r_prod_encode [x, Cn 1 r_prod_encode [y, z]]" + +definition "r_step_Cn \ + Cn 1 r_ifeq + [r_e2lenls, + Cn 1 r_length [Cn 1 r_pdec222 [r_e2i]], + Cn 1 r_ifz + [r_e2rv, + Cn 1 r_prod_encode + [Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec122 [r_e2i]) r_e2ls Z, r_e2stack], + Z], + Cn 1 r_prod_encode [r_e2tail, r_e2rv]], + Cn 1 r_ifz + [r_e2rv, + Cn 1 r_ifless + [r_e2lenls, + Cn 1 r_length [Cn 1 r_pdec222 [r_e2i]], + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode (Cn 1 r_nth [Cn 1 r_pdec222 [r_e2i], r_e2lenls]) r_e2xs Z, + r_e2stack], + Z], + Cn 1 r_prod_encode [r_e2tail, r_e2rv]], + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode r_e2i r_e2xs (Cn 1 r_snoc [r_e2ls, Cn 1 r_dec [r_e2rv]]), + r_e2tail], + Z]]]" + +lemma r_step_Cn_prim: "prim_recfn 1 r_step_Cn" + unfolding r_step_Cn_def using helpers_for_r_step_prim by simp + +lemma r_step_Cn: "eval r_step_Cn [e] \= estep_Cn e" + unfolding r_step_Cn_def estep_Cn_def using helpers_for_r_step_prim by simp + +definition "r_step_Pr \ + Cn 1 r_ifz + [r_e2ls, + Cn 1 r_ifz + [r_e2rv, + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode (Cn 1 r_pdec122 [r_e2i]) (Cn 1 r_tl [r_e2xs]) Z, + r_e2stack], + Z], + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Cn 1 r_dec [r_e2rv]]), + r_e2tail], + Z]], + Cn 1 r_ifeq + [r_e2lenls, + Cn 1 S [Cn 1 r_hd [r_e2xs]], + Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_hd [r_e2ls]]], + Cn 1 r_ifz + [r_e2rv, + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode + (Cn 1 r_pdec222 [r_e2i]) + (Cn 1 r_cons + [Cn 1 r_dec [r_e2lenls], + Cn 1 r_cons [Cn 1 r_hd [r_e2ls], + Cn 1 r_tl [r_e2xs]]]) + Z, + r_e2stack], + Z], + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode r_e2i r_e2xs (Cn 1 r_cons [Cn 1 r_dec [r_e2rv], r_e2ls]), + r_e2tail], + Z]]]]" + +lemma r_step_Pr_prim: "prim_recfn 1 r_step_Pr" + unfolding r_step_Pr_def using helpers_for_r_step_prim by simp + +lemma r_step_Pr: "eval r_step_Pr [e] \= estep_Pr e" + unfolding r_step_Pr_def estep_Pr_def using helpers_for_r_step_prim by simp + +definition "r_step_Mn \ + Cn 1 r_ifz + [r_e2ls, + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode (Cn 1 r_pdec22 [r_e2i]) (Cn 1 r_cons [Z, r_e2xs]) Z, + Cn 1 r_cons + [r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Z]), + r_e2tail]], + Z], + Cn 1 r_ifeq + [r_e2rv, + r_const 1, + Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_hd [r_e2ls]]], + Cn 1 r_prod_encode + [Cn 1 r_cons + [r_triple_encode + (Cn 1 r_pdec22 [r_e2i]) + (Cn 1 r_cons [Cn 1 S [Cn 1 r_hd [r_e2ls]], r_e2xs]) + Z, + Cn 1 r_cons + [r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Cn 1 S [Cn 1 r_hd [r_e2ls]]]), + r_e2tail]], + Z]]]" + +lemma r_step_Mn_prim: "prim_recfn 1 r_step_Mn" + unfolding r_step_Mn_def using helpers_for_r_step_prim by simp + +lemma r_step_Mn: "eval r_step_Mn [e] \= estep_Mn e" + unfolding r_step_Mn_def estep_Mn_def using helpers_for_r_step_prim by simp + +definition "r_step \ + Cn 1 r_ifz + [r_e2stack, + Cn 1 r_prod_encode [Z, r_e2rv], + Cn 1 r_ifz + [r_e2i, + Cn 1 r_prod_encode [r_e2tail, r_const 1], + Cn 1 r_ifeq + [r_e2i, + r_const 1, + Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 S [Cn 1 r_hd [r_e2xs]]]], + Cn 1 r_ifeq + [Cn 1 r_kind [r_e2i], + r_const 2, + Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_nth [r_e2xs, Cn 1 r_pdec22 [r_e2i]]]], + Cn 1 r_ifeq + [Cn 1 r_kind [r_e2i], + r_const 3, + r_step_Cn, + Cn 1 r_ifeq + [Cn 1 r_kind [r_e2i], + r_const 4, + r_step_Pr, + Cn 1 r_ifeq + [Cn 1 r_kind [r_e2i], r_const 5, r_step_Mn, Z]]]]]]]" + +lemma r_step_prim: "prim_recfn 1 r_step" + unfolding r_step_def + using r_kind_prim r_step_Mn_prim r_step_Pr_prim r_step_Cn_prim helpers_for_r_step_prim + by simp + +lemma r_step: "eval r_step [e] \= estep e" + unfolding r_step_def estep_def + using r_kind_prim r_step_Mn_prim r_step_Pr_prim r_step_Cn_prim helpers_for_r_step_prim + r_kind r_step_Cn r_step_Pr r_step_Mn + by simp + +theorem r_step_equiv_step: + assumes "valid (fst c)" + shows "eval r_step [encode_config c] \= encode_config (step c)" + using r_step estep assms by simp + + +subsection \The universal function\label{s:the_universal}\ + +text \The next function computes the configuration after arbitrarily +many steps.\ + +definition "r_leap \ + Pr 2 + (Cn 2 r_prod_encode + [Cn 2 r_singleton_encode + [Cn 2 r_prod_encode [Id 2 0, Cn 2 r_prod_encode [Id 2 1, r_constn 1 0]]], + r_constn 1 0]) + (Cn 4 r_step [Id 4 1])" + +lemma r_leap_prim [simp]: "prim_recfn 3 r_leap" + unfolding r_leap_def using r_step_prim by simp + +lemma r_leap_total: "eval r_leap [t, i, x] \" + using prim_recfn_total[OF r_leap_prim] by simp + +lemma r_leap: + assumes "i = encode f" and "recfn (e_length x) f" + shows "eval r_leap [t, i, x] \= encode_config (iterate t step ([(f, list_decode x, [])], None))" +proof (induction t) + case 0 + then show ?case + unfolding r_leap_def using r_step_prim assms encode_config encode_frame by simp +next + case (Suc t) + let ?c = "([(f, list_decode x, [])], None)" + let ?tc = "iterate t step ?c" + have "valid (fst ?c)" + using valid_def assms by simp + then have valid: "valid (fst ?tc)" + using iterate_step_valid by simp + have "eval r_leap [Suc t, i, x] = + eval (Cn 4 r_step [Id 4 1]) [t, the (eval r_leap [t, i, x]), i, x]" + unfolding r_leap_def using eval_Pr_prim_Suc r_step_prim assms by simp + then have "eval r_leap [Suc t, i, x] = eval (Cn 4 r_step [Id 4 1]) [t, encode_config ?tc, i, x]" + using Suc by simp + then have "eval r_leap [Suc t, i, x] = eval r_step [encode_config ?tc]" + using r_step_prim by simp + then have "eval r_leap [Suc t, i, x] \= encode_config (step ?tc)" + by (simp add: r_step_equiv_step valid) + then show ?case by simp +qed + +lemma step_leaves_empty_stack_empty: + assumes "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" + shows "iterate (t + t') step ([(f, list_decode x, [])], None) = ([], Some v)" + using assms by (induction t') simp_all + +text \The next function is essentially a convenience wrapper around +@{term r_leap}. It returns zero if the configuration returned by @{term +r_leap} is non-final, and @{term "Suc v"} if the configuration is final with +return value $v$.\ + +definition "r_result \ + Cn 3 r_ifz [Cn 3 r_pdec1 [r_leap], Cn 3 r_pdec2 [r_leap], r_constn 2 0]" + +lemma r_result_prim [simp]: "prim_recfn 3 r_result" + unfolding r_result_def using r_leap_prim by simp + +lemma r_result_total: "total r_result" + using r_result_prim by blast + +lemma r_result_empty_stack_None: + assumes "i = encode f" + and "recfn (e_length x) f" + and "iterate t step ([(f, list_decode x, [])], None) = ([], None)" + shows "eval r_result [t, i, x] \= 0" + unfolding r_result_def + using assms r_leap e2stack_0_iff_Nil e2stack_def e2stack_stack r_leap_total r_leap_prim + e2rv_def e2rv_rv + by simp + +lemma r_result_empty_stack_Some: + assumes "i = encode f" + and "recfn (e_length x) f" + and "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" + shows "eval r_result [t, i, x] \= Suc v" + unfolding r_result_def + using assms r_leap e2stack_0_iff_Nil e2stack_def e2stack_stack r_leap_total r_leap_prim + e2rv_def e2rv_rv + by simp + +lemma r_result_empty_stack_stays: + assumes "i = encode f" + and "recfn (e_length x) f" + and "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" + shows "eval r_result [t + t', i, x] \= Suc v" + using assms step_leaves_empty_stack_empty r_result_empty_stack_Some by simp + +lemma r_result_nonempty_stack: + assumes "i = encode f" + and "recfn (e_length x) f" + and "fst (iterate t step ([(f, list_decode x, [])], None)) \ []" + shows "eval r_result [t, i, x] \= 0" +proof - + obtain ss rv where "iterate t step ([(f, list_decode x, [])], None) = (ss, rv)" + by fastforce + moreover from this assms(3) have "ss \ []" by simp + ultimately have "eval r_leap [t, i, x] \= encode_config (ss, rv)" + using assms r_leap by simp + then have "eval (Cn 3 r_pdec1 [r_leap]) [t, i, x] \\ 0" + using `ss \ []` r_leap_prim encode_config r_leap_total list_encode_0 + by (auto, blast) + then show ?thesis unfolding r_result_def using r_leap_prim by auto +qed + +lemma r_result_Suc: + assumes "i = encode f" + and "recfn (e_length x) f" + and "eval r_result [t, i, x] \= Suc v" + shows "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" + (is "?cfg = _") +proof (cases "fst ?cfg") + case Nil + then show ?thesis + using assms r_result_empty_stack_None r_result_empty_stack_Some + by (metis Zero_not_Suc nat.inject option.collapse option.inject prod.exhaust_sel) +next + case Cons + then show ?thesis using assms r_result_nonempty_stack by simp +qed + +lemma r_result_converg: + assumes "i = encode f" + and "recfn (e_length x) f" + and "eval f (list_decode x) \= v" + shows "\t. + (\t'\t. eval r_result [t', i, x] \= Suc v) \ + (\t'= 0)" +proof - + let ?xs = "list_decode x" + let ?stack = "[(f, ?xs, [])]" + have "wellf f" using assms(2) by simp + moreover have "length ?xs = arity f" + using assms(2) by simp + ultimately have "correct (?stack, None)" + using step_correct valid_def by simp + with assms(3) have "reachable (?stack, None) ([], Some v)" + by simp + then obtain t where + "iterate t step (?stack, None) = ([], Some v)" + "\t' []" + using reachable_iterate_step_empty_stack by blast + then have t: + "eval r_result [t, i, x] \= Suc v" + "\t'= 0" + using r_result_empty_stack_Some r_result_nonempty_stack assms(1,2) + by simp_all + then have "eval r_result [t + t', i, x] \= Suc v" for t' + using r_result_empty_stack_stays assms r_result_Suc by simp + then have "\t'\t. eval r_result [t', i, x] \= Suc v" + using le_Suc_ex by blast + with t(2) show ?thesis by auto +qed + +lemma r_result_diverg: + assumes "i = encode f" + and "recfn (e_length x) f" + and "eval f (list_decode x) \" + shows "eval r_result [t, i, x] \= 0" +proof - + let ?xs = "list_decode x" + let ?stack = "[(f, ?xs, [])]" + have "recfn (length ?xs) f" + using assms(2) by auto + then have "correct (?stack, None)" + using step_correct valid_def by simp + with assms(3) have "nonterminating (?stack, None)" + by simp + then show ?thesis + using r_result_nonempty_stack assms(1,2) by simp +qed + +text \Now we can define the universal partial recursive function. This +function executes @{term r_result} for increasing time bounds, waits for it +to reach a final configuration, and then extracts its result value. If no +final configuration is reached, the universal function diverges.\ + +definition "r_univ \ + Cn 2 r_dec [Cn 2 r_result [Mn 2 (Cn 3 r_not [r_result]), Id 2 0, Id 2 1]]" + +lemma r_univ_recfn [simp]: "recfn 2 r_univ" + unfolding r_univ_def by simp + +theorem r_univ: + assumes "i = encode f" and "recfn (e_length x) f" + shows "eval r_univ [i, x] = eval f (list_decode x)" +proof - + let ?cond = "Cn 3 r_not [r_result]" + let ?while = "Mn 2 ?cond" + let ?res = "Cn 2 r_result [?while, Id 2 0, Id 2 1]" + let ?xs = "list_decode x" + have *: "eval ?cond [t, i, x] \= (if eval r_result [t, i, x] \= 0 then 1 else 0)" for t + proof - + have "eval ?cond [t, i, x] = eval r_not [the (eval r_result [t, i, x])]" + using r_result_total by simp + moreover have "eval r_result [t, i, x] \" + by (simp add: r_result_total) + ultimately show ?thesis by auto + qed + show ?thesis + proof (cases "eval f ?xs \") + case True + then show ?thesis + unfolding r_univ_def using * r_result_diverg[OF assms] eval_Mn_diverg by simp + next + case False + then obtain v where v: "eval f ?xs \= v" by auto + then obtain t where t: + "\t'\t. eval r_result [t', i, x] \= Suc v" + "\t'= 0" + using r_result_converg[OF assms] by blast + then have + "\t'\t. eval ?cond [t', i, x] \= 0" + "\t'= 1" + using * by simp_all + then have "eval ?while [i, x] \= t" + using eval_Mn_convergI[of 2 ?cond "[i, x]" t] by simp + then have "eval ?res [i, x] = eval r_result [t, i, x]" + by simp + then have "eval ?res [i, x] \= Suc v" + using t(1) by simp + then show ?thesis + unfolding r_univ_def using v by simp + qed +qed + +theorem r_univ': + assumes "recfn (e_length x) f" + shows "eval r_univ [encode f, x] = eval f (list_decode x)" + using r_univ assms by simp + +text \Universal functions for every arity can be built from @{term "r_univ"}.\ + +definition r_universal :: "nat \ recf" where + "r_universal n \ Cn (Suc n) r_univ [Id (Suc n) 0, r_shift (r_list_encode (n - 1))]" + +lemma r_universal_recfn [simp]: "n > 0 \ recfn (Suc n) (r_universal n)" + unfolding r_universal_def by simp + +lemma r_universal: + assumes "recfn n f" and "length xs = n" + shows "eval (r_universal n) (encode f # xs) = eval f xs" + unfolding r_universal_def using wellf_arity_nonzero assms r_list_encode r_univ' + by fastforce + +text \We will mostly be concerned with computing unary functions. Hence +we introduce separate functions for this case.\ + +definition "r_result1 \ + Cn 3 r_result [Id 3 0, Id 3 1, Cn 3 r_singleton_encode [Id 3 2]]" + +lemma r_result1_prim [simp]: "prim_recfn 3 r_result1" + unfolding r_result1_def by simp + +lemma r_result1_total: "total r_result1" + using Mn_free_imp_total by simp + +lemma r_result1 [simp]: + "eval r_result1 [t, i, x] = eval r_result [t, i, singleton_encode x]" + unfolding r_result1_def by simp + +text \The following function will be our standard Gödel numbering +of all unary partial recursive functions.\ + +definition "r_phi \ r_universal 1" + +lemma r_phi_recfn [simp]: "recfn 2 r_phi" + unfolding r_phi_def by simp + +theorem r_phi: + assumes "i = encode f" and "recfn 1 f" + shows "eval r_phi [i, x] = eval f [x]" + unfolding r_phi_def using r_universal assms by simp + +corollary r_phi': + assumes "recfn 1 f" + shows "eval r_phi [encode f, x] = eval f [x]" + using assms r_phi by simp + +lemma r_phi'': "eval r_phi [i, x] = eval r_univ [i, singleton_encode x]" + unfolding r_universal_def r_phi_def using r_list_encode by simp + + +section \Applications of the universal function\ + +text \In this section we shall see some ways @{term r_univ} and @{term r_result} can +be used.\ + +subsection \Lazy conditional evaluation\ + +text \With the help of @{term r_univ} we can now define a +\hypertarget{p:r_lifz}{lazy variant} of @{term r_ifz}, in which only one +branch is evaluated.\ + +definition r_lazyifzero :: "nat \ nat \ nat \ recf" where + "r_lazyifzero n j\<^sub>1 j\<^sub>2 \ + Cn (Suc (Suc n)) r_univ + [Cn (Suc (Suc n)) r_ifz [Id (Suc (Suc n)) 0, r_constn (Suc n) j\<^sub>1, r_constn (Suc n) j\<^sub>2], + r_shift (r_list_encode n)]" + +lemma r_lazyifzero_recfn: "recfn (Suc (Suc n)) (r_lazyifzero n j\<^sub>1 j\<^sub>2)" + using r_lazyifzero_def by simp + +lemma r_lazyifzero: + assumes "length xs = Suc n" + and "j\<^sub>1 = encode f\<^sub>1" + and "j\<^sub>2 = encode f\<^sub>2" + and "recfn (Suc n) f\<^sub>1" + and "recfn (Suc n) f\<^sub>2" + shows "eval (r_lazyifzero n j\<^sub>1 j\<^sub>2) (c # xs) = (if c = 0 then eval f\<^sub>1 xs else eval f\<^sub>2 xs)" +proof - + let ?a = "r_constn (Suc n) n" + let ?b = "Cn (Suc (Suc n)) r_ifz + [Id (Suc (Suc n)) 0, r_constn (Suc n) j\<^sub>1, r_constn (Suc n) j\<^sub>2]" + let ?c = "r_shift (r_list_encode n)" + have "eval ?a (c # xs) \= n" + using assms(1) by simp + moreover have "eval ?b (c # xs) \= (if c = 0 then j\<^sub>1 else j\<^sub>2)" + using assms(1) by simp + moreover have "eval ?c (c # xs) \= list_encode xs" + using assms(1) r_list_encode r_shift by simp + ultimately have "eval (r_lazyifzero n j\<^sub>1 j\<^sub>2) (c # xs) = + eval r_univ [if c = 0 then j\<^sub>1 else j\<^sub>2, list_encode xs]" + unfolding r_lazyifzero_def using r_lazyifzero_recfn assms(1) by simp + then show ?thesis using assms r_univ by simp +qed + +definition r_lifz :: "recf \ recf \ recf" where + "r_lifz f g \ r_lazyifzero (arity f - 1) (encode f) (encode g)" + +lemma r_lifz_recfn [simp]: + assumes "recfn n f" and "recfn n g" + shows "recfn (Suc n) (r_lifz f g)" + using assms r_lazyifzero_recfn r_lifz_def wellf_arity_nonzero by auto + +lemma r_lifz [simp]: + assumes "length xs = n" and "recfn n f" and "recfn n g" + shows "eval (r_lifz f g) (c # xs) = (if c = 0 then eval f xs else eval g xs)" + using assms r_lazyifzero r_lifz_def wellf_arity_nonzero + by (metis One_nat_def Suc_pred) + + +subsection \Enumerating the domains of partial recursive functions\ + +text \In this section we define a binary function $\mathit{enumdom}$ +such that for all $i$, the domain of $\varphi_i$ equals +$\{\mathit{enumdom}(i, x) \mid \mathit{enumdom}(i, x)\!\downarrow\}$. In +other words, the image of $\mathit{enumdom}_i$ is the domain of $\varphi_i$. + +First we need some more properties of @{term r_leap} and @{term r_result}.\ + +lemma r_leap_Suc: "eval r_leap [Suc t, i, x] = eval r_step [the (eval r_leap [t, i, x])]" +proof - + have "eval r_leap [Suc t, i, x] = + eval (Cn 4 r_step [Id 4 1]) [t, the (eval r_leap [t, i, x]), i, x]" + using r_leap_total eval_Pr_converg_Suc r_leap_def + by (metis length_Cons list.size(3) numeral_2_eq_2 numeral_3_eq_3 r_leap_prim) + then show ?thesis using r_step_prim by auto +qed + +lemma r_leap_Suc_saturating: + assumes "pdec1 (the (eval r_leap [t, i, x])) = 0" + shows "eval r_leap [Suc t, i, x] = eval r_leap [t, i, x]" +proof - + let ?e = "eval r_leap [t, i, x]" + have "eval r_step [the ?e] \= estep (the ?e)" + using r_step by simp + then have "eval r_step [the ?e] \= prod_encode (0, e2rv (the ?e))" + using estep_def assms by (simp add: e2stack_def) + then have "eval r_step [the ?e] \= prod_encode (pdec1 (the ?e), pdec2 (the ?e))" + using assms by (simp add: e2rv_def) + then have "eval r_step [the ?e] \= the ?e" by simp + then show ?thesis using r_leap_total r_leap_Suc by simp +qed + +lemma r_result_Suc_saturating: + assumes "eval r_result [t, i, x] \= Suc v" + shows "eval r_result [Suc t, i, x] \= Suc v" +proof - + let ?r = "\t. eval r_ifz [pdec1 (the (eval r_leap [t, i, x])), pdec2 (the (eval r_leap [t, i, x])), 0]" + have "?r t \= Suc v" + using assms unfolding r_result_def using r_leap_total r_leap_prim by simp + then have "pdec1 (the (eval r_leap [t, i, x])) = 0" + using option.sel by fastforce + then have "eval r_leap [Suc t, i, x] = eval r_leap [t, i, x]" + using r_leap_Suc_saturating by simp + moreover have "eval r_result [t, i, x] = ?r t" + unfolding r_result_def using r_leap_total r_leap_prim by simp + moreover have "eval r_result [Suc t, i, x] = ?r (Suc t)" + unfolding r_result_def using r_leap_total r_leap_prim by simp + ultimately have "eval r_result [Suc t, i, x] = eval r_result [t, i, x]" + by simp + with assms show ?thesis by simp +qed + +lemma r_result_saturating: + assumes "eval r_result [t, i, x] \= Suc v" + shows "eval r_result [t + d, i, x] \= Suc v" + using r_result_Suc_saturating assms by (induction d) simp_all + +lemma r_result_converg': + assumes "eval r_univ [i, x] \= v" + shows "\t. (\t'\t. eval r_result [t', i, x] \= Suc v) \ (\t'= 0)" +proof - + let ?f = "Cn 3 r_not [r_result]" + let ?m = "Mn 2 ?f" + have "recfn 2 ?m" by simp + have eval_m: "eval ?m [i, x] \" + proof + assume "eval ?m [i, x] \" + then have "eval r_univ [i, x] \" + unfolding r_univ_def by simp + with assms show False by simp + qed + then obtain t where t: "eval ?m [i, x] \= t" + by auto + then have f_t: "eval ?f [t, i, x] \= 0" and f_less_t: "\y. y < t \ eval ?f [y, i, x] \\ 0" + using eval_Mn_convergE[of 2 ?f "[i, x]" t] `recfn 2 ?m` + by (metis (no_types, lifting) One_nat_def Suc_1 length_Cons list.size(3))+ + have eval_Cn2: "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \" + proof + assume "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \" + then have "eval r_univ [i, x] \" + unfolding r_univ_def by simp + with assms show False by simp + qed + have "eval r_result [t, i, x] \= Suc v" + proof (rule ccontr) + assume neq_Suc: "\ eval r_result [t, i, x] \= Suc v" + show False + proof (cases "eval r_result [t, i, x] = None") + case True + then show ?thesis using f_t by simp + next + case False + then obtain w where w: "eval r_result [t, i, x] \= w" "w \ Suc v" + using neq_Suc by auto + moreover have "eval r_result [t, i, x] \\ 0" + by (rule ccontr; use f_t in auto) + ultimately have "w \ 0" by simp + have "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] = + eval r_result [the (eval ?m [i, x]), i, x]" + using eval_m by simp + with w t have "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \= w" + by simp + moreover have "eval r_univ [i, x] = + eval r_dec [the (eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x])]" + unfolding r_univ_def using eval_Cn2 by simp + ultimately have "eval r_univ [i, x] = eval r_dec [w]" by simp + then have "eval r_univ [i, x] \= w - 1" by simp + with assms `w \ 0` w show ?thesis by simp + qed + qed + then have "\t'\t. eval r_result [t', i, x] \= Suc v" + using r_result_saturating le_Suc_ex by blast + moreover have "eval r_result [y, i, x] \= 0" if "y < t" for y + proof (rule ccontr) + assume neq0: "eval r_result [y, i, x] \ Some 0" + then show False + proof (cases "eval r_result [y, i, x] = None") + case True + then show ?thesis using f_less_t `y < t` by fastforce + next + case False + then obtain v where "eval r_result [y, i, x] \= v" "v \ 0" + using neq0 by auto + then have "eval ?f [y, i, x] \= 0" by simp + then show ?thesis using f_less_t `y < t` by simp + qed + qed + ultimately show ?thesis by auto +qed + +lemma r_result_diverg': + assumes "eval r_univ [i, x] \" + shows "eval r_result [t, i, x] \= 0" +proof (rule ccontr) + let ?f = "Cn 3 r_not [r_result]" + let ?m = "Mn 2 ?f" + assume "eval r_result [t, i, x] \ Some 0" + with r_result_total have "eval r_result [t, i, x] \\ 0" by simp + then have "eval ?f [t, i, x] \= 0" by auto + moreover have "eval ?f [y, i, x] \" if "y < t" for y + using r_result_total by simp + ultimately have "\z. eval ?f (z # [i, x]) \= 0 \ (\y)" + by blast + then have "eval ?m [i, x] \" by simp + then have "eval r_univ [i, x] \" + unfolding r_univ_def using r_result_total by simp + with assms show False by simp +qed + +lemma r_result_bivalent': + assumes "eval r_univ [i, x] \= v" + shows "eval r_result [t, i, x] \= Suc v \ eval r_result [t, i, x] \= 0" + using r_result_converg'[OF assms] not_less by blast + +lemma r_result_Some': + assumes "eval r_result [t, i, x] \= Suc v" + shows "eval r_univ [i, x] \= v" +proof (rule ccontr) + assume not_v: "\ eval r_univ [i, x] \= v" + show False + proof (cases "eval r_univ [i, x] \") + case True + then show ?thesis + using assms r_result_diverg' by simp + next + case False + then obtain w where w: "eval r_univ [i, x] \= w" "w \ v" + using not_v by auto + then have "eval r_result [t, i, x] \= Suc w \ eval r_result [t, i, x] \= 0" + using r_result_bivalent' by simp + then show ?thesis using assms not_v w by simp + qed +qed + +lemma r_result1_converg': + assumes "eval r_phi [i, x] \= v" + shows "\t. + (\t'\t. eval r_result1 [t', i, x] \= Suc v) \ + (\t'= 0)" + using assms r_result1 r_result_converg' r_phi'' by simp + +lemma r_result1_diverg': + assumes "eval r_phi [i, x] \" + shows "eval r_result1 [t, i, x] \= 0" + using assms r_result1 r_result_diverg' r_phi'' by simp + +lemma r_result1_Some': + assumes "eval r_result1 [t, i, x] \= Suc v" + shows "eval r_phi [i, x] \= v" + using assms r_result1 r_result_Some' r_phi'' by simp + +text \The next function performs dovetailing in order to evaluate +$\varphi_i$ for every argument for arbitrarily many steps. Given $i$ and $z$, +the function decodes $z$ into a pair $(x, t$) and outputs zero (meaning +``true'') iff.\ the computation of $\varphi_i$ on input $x$ halts after at most +$t$ steps. Fixing $i$ and varying $z$ will eventually compute $\varphi_i$ +for every argument in the domain of $\varphi_i$ sufficiently long for it to +converge.\ + +definition "r_dovetail \ + Cn 2 r_not [Cn 2 r_result1 [Cn 2 r_pdec2 [Id 2 1], Id 2 0, Cn 2 r_pdec1 [Id 2 1]]]" + +lemma r_dovetail_prim: "prim_recfn 2 r_dovetail" + by (simp add: r_dovetail_def) + +lemma r_dovetail: + "eval r_dovetail [i, z] \= + (if the (eval r_result1 [pdec2 z, i, pdec1 z]) > 0 then 0 else 1)" + unfolding r_dovetail_def using r_result_total by simp + +text \The function $\mathit{enumdom}$ works as follows in order to +enumerate exactly the domain of $\varphi_i$. Given $i$ and $y$ it searches +for the minimum $z \geq y$ for which the dovetail function returns true. This +$z$ is decoded into $(x, t)$ and the $x$ is output. In this way every value +output by $\mathit{enumdom}$ is in the domain of $\varphi_i$ by construction +of @{term r_dovetail}. Conversely an $x$ in the domain will be output for $y += (x, t)$ where $t$ is such that $\varphi_i$ halts on $x$ within $t$ +steps.\ + +definition "r_dovedelay \ + Cn 3 r_and + [Cn 3 r_dovetail [Id 3 1, Id 3 0], + Cn 3 r_ifle [Id 3 2, Id 3 0, r_constn 2 0, r_constn 2 1]]" + +lemma r_dovedelay_prim: "prim_recfn 3 r_dovedelay" + unfolding r_dovedelay_def using r_dovetail_prim by simp + +lemma r_dovedelay: + "eval r_dovedelay [z, i, y] \= + (if the (eval r_result1 [pdec2 z, i, pdec1 z]) > 0 \ y \ z then 0 else 1)" + by (simp add: r_dovedelay_def r_dovetail r_dovetail_prim) + +definition "r_enumdom \ Cn 2 r_pdec1 [Mn 2 r_dovedelay]" + +lemma r_enumdom_recfn [simp]: "recfn 2 r_enumdom" + by (simp add: r_enumdom_def r_dovedelay_prim) + +lemma r_enumdom [simp]: + "eval r_enumdom [i, y] = + (if \z. eval r_dovedelay [z, i, y] \= 0 + then Some (pdec1 (LEAST z. eval r_dovedelay [z, i, y] \= 0)) + else None)" +proof - + let ?h = "Mn 2 r_dovedelay" + have "total r_dovedelay" + using r_dovedelay_prim by blast + then have "eval ?h [i, y] = + (if (\z. eval r_dovedelay [z, i, y] \= 0) + then Some (LEAST z. eval r_dovedelay [z, i, y] \= 0) + else None)" + using r_dovedelay_prim r_enumdom_recfn eval_Mn_convergI by simp + then show ?thesis + unfolding r_enumdom_def using r_dovedelay_prim by simp +qed + +text \If @{term i} is the code of the empty function, @{term r_enumdom} +has an empty domain, too.\ + +lemma r_enumdom_empty_domain: + assumes "\x. eval r_phi [i, x] \" + shows "\y. eval r_enumdom [i, y] \" + using assms r_result1_diverg' r_dovedelay by simp + +text \If @{term i} is the code of a function with non-empty domain, +@{term r_enumdom} enumerates its domain.\ + +lemma r_enumdom_nonempty_domain: + assumes "eval r_phi [i, x\<^sub>0] \" + shows "\y. eval r_enumdom [i, y] \" + and "\x. eval r_phi [i, x] \ \ (\y. eval r_enumdom [i, y] \= x)" +proof - + show "eval r_enumdom [i, y] \" for y + proof - + obtain t where t: "\t'\t. the (eval r_result1 [t', i, x\<^sub>0]) > 0" + using assms r_result1_converg' by fastforce + let ?z = "prod_encode (x\<^sub>0, max t y)" + have "y \ ?z" + using le_prod_encode_2 max.bounded_iff by blast + moreover have "pdec2 ?z \ t" by simp + ultimately have "the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0" + using t by simp + with `y \ ?z` r_dovedelay have "eval r_dovedelay [?z, i, y] \= 0" + by presburger + then show "eval r_enumdom [i, y] \" + using r_enumdom by auto + qed + show "eval r_phi [i, x] \ = (\y. eval r_enumdom [i, y] \= x)" for x + proof + show "\y. eval r_enumdom [i, y] \= x" if "eval r_phi [i, x] \" for x + proof - + from that obtain v where "eval r_phi [i, x] \= v" by auto + then obtain t where t: "the (eval r_result1 [t, i, x]) > 0" + using r_result1_converg' assms + by (metis Zero_not_Suc dual_order.refl option.sel zero_less_iff_neq_zero) + let ?y = "prod_encode (x, t)" + have "eval r_dovedelay [?y, i, ?y] \= 0" + using r_dovedelay t by simp + moreover from this have "(LEAST z. eval r_dovedelay [z, i, ?y] \= 0) = ?y" + using gr_implies_not_zero r_dovedelay by (intro Least_equality; fastforce) + ultimately have "eval r_enumdom [i, ?y] \= x" + using r_enumdom by auto + then show ?thesis by blast + qed + show "eval r_phi [i, x] \" if "\y. eval r_enumdom [i, y] \= x" for x + proof - + from that obtain y where y: "eval r_enumdom [i, y] \= x" + by auto + then have "eval r_enumdom [i, y] \" + by simp + then have + "\z. eval r_dovedelay [z, i, y] \= 0" and + *: "eval r_enumdom [i, y] \= pdec1 (LEAST z. eval r_dovedelay [z, i, y] \= 0)" + (is "_ \= pdec1 ?z") + using r_enumdom by metis+ + then have z: "eval r_dovedelay [?z, i, y] \= 0" + by (meson wellorder_Least_lemma(1)) + have "the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0" + proof (rule ccontr) + assume "\ (the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0)" + then show False + using r_dovedelay z by simp + qed + then have "eval r_phi [i, pdec1 ?z] \" + using r_result1_diverg' assms by fastforce + then show ?thesis using y * by auto + qed + qed +qed + +text \For every $\varphi_i$ with non-empty domain there is a total +recursive function that enumerates the domain of $\varphi_i$.\ + +lemma nonempty_domain_enumerable: + assumes "eval r_phi [i, x\<^sub>0] \" + shows "\g. recfn 1 g \ total g \ (\x. eval r_phi [i, x] \ \ (\y. eval g [y] \= x))" +proof - + define g where "g \ Cn 1 r_enumdom [r_const i, Id 1 0]" + then have "recfn 1 g" by simp + moreover from this have "total g" + using totalI1[of g] g_def assms r_enumdom_nonempty_domain(1) by simp + moreover have "eval r_phi [i, x] \ \ (\y. eval g [y] \= x)" for x + unfolding g_def using r_enumdom_nonempty_domain(2)[OF assms] by simp + ultimately show ?thesis by auto +qed + + +subsection \Concurrent evaluation of functions\ + +text \We define a function that simulates two @{typ recf}s +``concurrently'' for the same argument and returns the result of the one +converging first. If both diverge, so does the simulation function.\ + +definition "r_both \ + Cn 4 r_ifz + [Cn 4 r_result1 [Id 4 0, Id 4 1, Id 4 3], + Cn 4 r_ifz + [Cn 4 r_result1 [Id 4 0, Id 4 2, Id 4 3], + Cn 4 r_prod_encode [r_constn 3 2, r_constn 3 0], + Cn 4 r_prod_encode + [r_constn 3 1, Cn 4 r_dec [Cn 4 r_result1 [Id 4 0, Id 4 2, Id 4 3]]]], + Cn 4 r_prod_encode + [r_constn 3 0, Cn 4 r_dec [Cn 4 r_result1 [Id 4 0, Id 4 1, Id 4 3]]]]" + +lemma r_both_prim [simp]: "prim_recfn 4 r_both" + unfolding r_both_def by simp + +lemma r_both: + assumes "\x. eval r_phi [i, x] = eval f [x]" + and "\x. eval r_phi [j, x] = eval g [x]" + shows "eval f [x] \ \ eval g [x] \ \ eval r_both [t, i, j, x] \= prod_encode (2, 0)" + and "\eval r_result1 [t, i, x] \= 0; eval r_result1 [t, j, x] \= 0\ \ + eval r_both [t, i, j, x] \= prod_encode (2, 0)" + and "eval r_result1 [t, i, x] \= Suc v \ + eval r_both [t, i, j, x] \= prod_encode (0, the (eval f [x]))" + and "\eval r_result1 [t, i, x] \= 0; eval r_result1 [t, j, x] \= Suc v\ \ + eval r_both [t, i, j, x] \= prod_encode (1, the (eval g [x]))" +proof - + have r_result_total [simp]: "eval r_result [t, k, x] \" for t k x + using r_result_total by simp + { + assume "eval f [x] \ \ eval g [x] \" + then have "eval r_result1 [t, i, x] \= 0" and "eval r_result1 [t, j, x] \= 0" + using assms r_result1_diverg' by auto + then show "eval r_both [t, i, j, x] \= prod_encode (2, 0)" + unfolding r_both_def by simp + next + assume "eval r_result1 [t, i, x] \= 0" and "eval r_result1 [t, j, x] \= 0" + then show "eval r_both [t, i, j, x] \= prod_encode (2, 0)" + unfolding r_both_def by simp + next + assume "eval r_result1 [t, i, x] \= Suc v" + moreover from this have "eval r_result1 [t, i, x] \= Suc (the (eval f [x]))" + using assms r_result1_Some' by fastforce + ultimately show "eval r_both [t, i, j, x] \= prod_encode (0, the (eval f [x]))" + unfolding r_both_def by auto + next + assume "eval r_result1 [t, i, x] \= 0" and "eval r_result1 [t, j, x] \= Suc v" + moreover from this have "eval r_result1 [t, j, x] \= Suc (the (eval g [x]))" + using assms r_result1_Some' by fastforce + ultimately show "eval r_both [t, i, j, x] \= prod_encode (1, the (eval g [x]))" + unfolding r_both_def by auto + } +qed + +definition "r_parallel \ + Cn 3 r_both [Mn 3 (Cn 4 r_le [Cn 4 r_pdec1 [r_both], r_constn 3 1]), Id 3 0, Id 3 1, Id 3 2]" + +lemma r_parallel_recfn [simp]: "recfn 3 r_parallel" + unfolding r_parallel_def by simp + +lemma r_parallel: + assumes "\x. eval r_phi [i, x] = eval f [x]" + and "\x. eval r_phi [j, x] = eval g [x]" + shows "eval f [x] \ \ eval g [x] \ \ eval r_parallel [i, j, x] \" + and "eval f [x] \ \ eval g [x] \ \ + eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x]))" + and "eval g [x] \ \ eval f [x] \ \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" + and "eval f [x] \ \ eval g [x] \ \ + eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x])) \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" +proof - + let ?cond = "Cn 4 r_le [Cn 4 r_pdec1 [r_both], r_constn 3 1]" + define m where "m = Mn 3 ?cond" + then have m: "r_parallel = Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]" + unfolding r_parallel_def by simp + from m_def have "recfn 3 m" by simp + { + assume "eval f [x] \ \ eval g [x] \" + then have "\t. eval r_both [t, i, j, x] \= prod_encode (2, 0)" + using assms r_both by simp + then have "eval ?cond [t, i, j, x] \= 1" for t + by simp + then have "eval m [i, j, x] \" + unfolding m_def using eval_Mn_diverg by simp + then have "eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x] \" + using `recfn 3 m` by simp + then show "eval r_parallel [i, j, x] \" + using m by simp + next + assume "eval f [x] \ \ eval g [x] \" + then obtain vf vg where v: "eval f [x] \= vf" "eval g [x] \= vg" + by auto + then obtain tf where tf: + "\t\tf. eval r_result1 [t, i, x] \= Suc vf" + "\t= 0" + using r_result1_converg' assms by metis + from v obtain tg where tg: + "\t\tg. eval r_result1 [t, j, x] \= Suc vg" + "\t= 0" + using r_result1_converg' assms by metis + show "eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x])) \ + eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" + proof (cases "tf \ tg") + case True + with tg(2) have j0: "\t= 0" + by simp + have *: "eval r_both [tf, i, j, x] \= prod_encode (0, the (eval f [x]))" + using r_both(3) assms tf(1) by simp + have "eval m [i, j, x] \= tf" + unfolding m_def + proof (rule eval_Mn_convergI) + show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp + have "eval (Cn 4 r_pdec1 [r_both]) [tf, i, j, x] \= 0" + using * by simp + then show "eval ?cond [tf, i, j, x] \= 0" by simp + have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < tf" for t + using tf(2) r_both(2) assms that j0 by simp + then have "eval ?cond [t, i, j, x] \= 1" if "t < tf" for t + using that by simp + then show "\y. y < tf \ eval ?cond [y, i, j, x] \\ 0" by simp + qed + moreover have "eval r_parallel [i, j, x] = + eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" + using m by simp + ultimately have "eval r_parallel [i, j, x] = eval r_both [tf, i, j, x]" + using `recfn 3 m` by simp + with * have "eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x]))" + by simp + then show ?thesis by simp + next + case False + with tf(2) have i0: "\t\tg. eval r_result1 [t, i, x] \= 0" + by simp + then have *: "eval r_both [tg, i, j, x] \= prod_encode (1, the (eval g [x]))" + using assms r_both(4) tg(1) by auto + have "eval m [i, j, x] \= tg" + unfolding m_def + proof (rule eval_Mn_convergI) + show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp + have "eval (Cn 4 r_pdec1 [r_both]) [tg, i, j, x] \= 1" + using * by simp + then show "eval ?cond [tg, i, j, x] \= 0" by simp + have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < tg" for t + using tg(2) r_both(2) assms that i0 by simp + then have "eval ?cond [t, i, j, x] \= 1" if "t < tg" for t + using that by simp + then show "\y. y < tg \ eval ?cond [y, i, j, x] \\ 0" by simp + qed + moreover have "eval r_parallel [i, j, x] = + eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" + using m by simp + ultimately have "eval r_parallel [i, j, x] = eval r_both [tg, i, j, x]" + using `recfn 3 m` by simp + with * have "eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" + by simp + then show ?thesis by simp + qed + next + assume eval_fg: "eval g [x] \ \ eval f [x] \" + then have i0: "\t. eval r_result1 [t, i, x] \= 0" + using r_result1_diverg' assms by auto + from eval_fg obtain v where "eval g [x] \= v" + by auto + then obtain t\<^sub>0 where t0: + "\t\t\<^sub>0. eval r_result1 [t, j, x] \= Suc v" + "\t0. eval r_result1 [t, j, x] \= 0" + using r_result1_converg' assms by metis + then have *: "eval r_both [t\<^sub>0, i, j, x] \= prod_encode (1, the (eval g [x]))" + using r_both(4) assms i0 by simp + have "eval m [i, j, x] \= t\<^sub>0" + unfolding m_def + proof (rule eval_Mn_convergI) + show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp + have "eval (Cn 4 r_pdec1 [r_both]) [t\<^sub>0, i, j, x] \= 1" + using * by simp + then show "eval ?cond [t\<^sub>0, i, j, x] \= 0" by simp + have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < t\<^sub>0" for t + using t0(2) r_both(2) assms that i0 by simp + then have "eval ?cond [t, i, j, x] \= 1" if "t < t\<^sub>0" for t + using that by simp + then show "\y. y < t\<^sub>0 \ eval ?cond [y, i, j, x] \\ 0" by simp + qed + moreover have "eval r_parallel [i, j, x] = + eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" + using m by simp + ultimately have "eval r_parallel [i, j, x] = eval r_both [t\<^sub>0, i, j, x]" + using `recfn 3 m` by simp + with * show "eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" + by simp + next + assume eval_fg: "eval f [x] \ \ eval g [x] \" + then have j0: "\t. eval r_result1 [t, j, x] \= 0" + using r_result1_diverg' assms by auto + from eval_fg obtain v where "eval f [x] \= v" + by auto + then obtain t\<^sub>0 where t0: + "\t\t\<^sub>0. eval r_result1 [t, i, x] \= Suc v" + "\t0. eval r_result1 [t, i, x] \= 0" + using r_result1_converg' assms by metis + then have *: "eval r_both [t\<^sub>0, i, j, x] \= prod_encode (0, the (eval f [x]))" + using r_both(3) assms by blast + have "eval m [i, j, x] \= t\<^sub>0" + unfolding m_def + proof (rule eval_Mn_convergI) + show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp + have "eval (Cn 4 r_pdec1 [r_both]) [t\<^sub>0, i, j, x] \= 0" + using * by simp + then show "eval ?cond [t\<^sub>0, i, j, x] \= 0" + by simp + have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < t\<^sub>0" for t + using t0(2) r_both(2) assms that j0 by simp + then have "eval ?cond [t, i, j, x] \= 1" if "t < t\<^sub>0" for t + using that by simp + then show "\y. y < t\<^sub>0 \ eval ?cond [y, i, j, x] \\ 0" by simp + qed + moreover have "eval r_parallel [i, j, x] = + eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" + using m by simp + ultimately have "eval r_parallel [i, j, x] = eval r_both [t\<^sub>0, i, j, x]" + using `recfn 3 m` by simp + with * show "eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x]))" + by simp + } +qed + +end \ No newline at end of file diff --git a/thys/Inductive_Inference/document/root.bib b/thys/Inductive_Inference/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/document/root.bib @@ -0,0 +1,181 @@ +@book{Rogers87, + Author = {Hartley {Rogers, Jr.}}, + Publisher = {The MIT Press}, + Title = {Theory of Recursive Functions and Effective Computability}, + Year = {1987}, + Edition = {2nd} +} + +@Article{Kleene43, + Author = "Stephen Cole Kleene", + Title = "Recursive predicates and quantifiers", + Journal = "Trans. Amer. Math. Soc.", + Volume = "53", + Number = "1", + Year = "1943", + Pages = "41--73", + doi = "10.1090/S0002-9947-1943-0007371-8" +} + +@misc{wiki-krt, + author = "{Wikipedia contributors}", + title = "Kleene's recursion theorem --- {Wikipedia}{,} The Free Encyclopedia", + year = "2020", + url = "https://en.wikipedia.org/w/index.php?title=Kleene%27s_recursion_theorem&oldid=936277979", + note = "[Online; accessed 28-March-2020]" +} + +@InCollection{as-ii-87, + author = "Dana Angluin and Carl H. Smith", + title = "Inductive Inference", + booktitle = "Encyclopedia of Artificial Intelligence", + publisher = "J. Wiley and Sons, New York", + year = "1987", + pages = "409--418", +} + +@Article{fkw-iisde-95, + author = "R\={u}si\c{n}\v{s} Freivalds and Efim B. Kinber and + Rolf Wiehagen", + title = "How Inductive Inference Strategies Discover Their + Errors", + journal = "Inform. Comput.", + volume = "118", + number = "2", + year = "1995", + pages = "208--226", +} + +@Article{cs-cicmii-83, + author = "John Case and Carl H. Smith", + title = "Comparison of Identification Criteria for Machine + Inductive Inference", + journal = "Theoret. Comput. Sci.", + volume = "25", + year = "1983", + pages = "193--220", + annote = "Was in STOC78", +} + +@InCollection{b-ttlsf-74, + author = "J. M. Barzdin", + title = "Two Theorems on the Limiting Synthesis of Functions", + booktitle = "Theory of Algorithms and Programs", + volume = "1", + publisher = "Latvian State University, Riga", + year = "1974", + pages = "82--88", + note = "In Russian", +} + +@Article{jb-cpnii-81, + author = "Klaus P. Jantke and Hans-Rainer Beick", + title = "Combining postulates of naturalness in inductive inference", + journal = "Elektronische Informationsverarbeitung und Kybernetik", + volume = "17", + number = "8/9", + year = "1981", + pages = "465--484", +} + +@Article{w-lerfss-76, + author = "Rolf Wiehagen", + title = "Limes-{E}rkennung rekursiver {F}unktionen durch spezielle {S}trategien", + journal = "Journal of Information Processing and Cybernetics (EIK)", + volume = "12", + year = "1976", + pages = "93--99", +} + +@Article{wz-idmowle-94, + author = "Rolf Wiehagen and Thomas Zeugmann", + title = "Ignoring data may be the only way to learn efficiently", + journal = "J. of Experimental and Theoret. Artif. Intell.", + volume = "6", + number = "1", + year = "1994", + pages = "131--144", +} + +@Article{g-lil-67, + author = "E. Mark Gold", + title = "Language Identification in the Limit", + journal = "Inform. Control", + volume = "10", + number = "5", + year = "1967", + pages = "447--474", + comment = "Classic paper, introducing computer science theory + into learning.", +} + +@Article{g-lr-65, + author = "E. M. Gold", + title = "Limiting Recursion", + journal = "J. Symbolic Logic", + volume = "30", + year = "1965", + pages = "28--48", +} + +@Article{s-ftiip1-64, + author = "R. J. Solomonoff", + title = "A Formal Theory of Inductive Inference: Part 1", + journal = "Inform. Control", + volume = "7", + year = "1964", + pages = "1--22", + comment = "Concerned with extrapolation of sequences. Defines + probability of extension via likelihood random TM + program will generate it.", +} + +@Article{s-ftiip2-64, + author = "R. J. Solomonoff", + title = "A Formal Theory of Inductive Inference: Part 2", + journal = "Inform. Control", + volume = "7", + year = "1964", + pages = "224--254", + comment = "Continues Part I. Inference of probabilities and + grammars.", +} + +@InProceedings{b-iiafp-74, + author = "{Ya}. M. Barzdin", + title = "Inductive Inference of Automata, Functions and Programs", + booktitle = "Proceedings International Congress of Mathematics", + year = "1974", + venue = "Vancouver", + pages = "455--460", +} + +@InProceedings{b-iiafp-77, + author = "J. M. Barzdin", + title = "Inductive Inference of Automata, Functions and Programs", + booktitle = "Amer. Math. Soc. Transl.", + year = "1977", + pages = "107--122", +} + +@Article{bb-tmtii-75, + author = "Leonore Blum and Manuel Blum", + title = "Toward a Mathematical Theory of Inductive Inference", + journal = "Inform. Control", + volume = "28", + number = "2", + month = jun, + year = "1975", + pages = "125--155", +} + +@article{Universal_Turing_Machine-AFP, + author = {Jian Xu and Xingyuan Zhang and Christian Urban and Sebastiaan J. C. Joosten}, + title = {Universal Turing Machine}, + journal = {Archive of Formal Proofs}, + month = feb, + year = 2019, + note = {\url{http://isa-afp.org/entries/Universal_Turing_Machine.html}, + Formal proof development}, + ISSN = {2150-914x}, +} diff --git a/thys/Inductive_Inference/document/root.tex b/thys/Inductive_Inference/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Inductive_Inference/document/root.tex @@ -0,0 +1,71 @@ +\documentclass[11pt,a4paper]{report} +\usepackage{isabelle,isabellesym} + +\usepackage[utf8]{inputenc} + +\usepackage[top=3cm,bottom=3cm]{geometry} + +\usepackage{amssymb} % for \mathbb + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Some classical results in inductive inference of recursive functions} +\author{Frank J. Balbach} +\maketitle + +\begin{abstract} +This entry formalizes some classical concepts and results from inductive +inference of recursive functions. In the basic setting a partial recursive +function (``strategy'') must identify (``learn'') all functions from a set +(``class'') of recursive functions. To that end the strategy receives more and +more values $f(0), f(1), f(2), \ldots$ of some function $f$ from the given class +and in turn outputs descriptions of partial recursive functions, for example, +Gödel numbers. The strategy is considered successful if the sequence of outputs +(``hypotheses'') converges to a description of $f$. A class of functions +learnable in this sense is called ``learnable in the limit''. The set of all +these classes is denoted by LIM. + +Other types of inference considered are finite learning (FIN), behaviorally +correct learning in the limit (BC), and some variants of LIM with restrictions +on the hypotheses: total learning (TOTAL), consistent learning (CONS), and +class-preserving learning (CP). The main results formalized are the proper +inclusions $\mathrm{FIN} \subset \mathrm{CP} \subset \mathrm{TOTAL} \subset +\mathrm{CONS} \subset \mathrm{LIM} \subset \mathrm{BC} \subset 2^{\mathcal{R}}$, +where $\mathcal{R}$ is the set of all total recursive functions. Further +results show that for all these inference types except CONS, strategies can be +assumed to be total recursive functions; that all inference types but CP are +closed under the subset relation between classes; and that no inference type is +closed under the union of classes. + +The above is based on a formalization of recursive functions heavily inspired by +the \emph{Universal Turing Machine} entry by +Xu~et~al.~\cite{Universal_Turing_Machine-AFP}, but different in that it models +partial functions with codomain \emph{nat option}. The formalization contains a +construction of a universal partial recursive function, without resorting to +Turing machines, introduces decidability and recursive enumerability, and proves +some standard results: existence of a Kleene normal form, the $s$-$m$-$n$ +theorem, Rice's theorem, and assorted fixed-point theorems (recursion theorems) +by Kleene, Rogers, and Smullyan. +\end{abstract} + +\tableofcontents + +\newpage + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,553 +1,554 @@ ADS_Functor AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_CC BNF_Operations Banach_Steinhaus Bell_Numbers_Spivey Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructor_Funs Containers CoreC++ Core_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hidden_Markov_Models Higher_Order_Terms Hoare_Time HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality +Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mersenne_Primes MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PCF PLM POPLmark-deBruijn PSemigroupsConvolution PAC_Checker Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Projective_Geometry Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares SuperCalc Surprise_Paradox Symmetric_Polynomials Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation VectorSpace VeriComp Verified-Prover VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL