|
| 1 | +/- |
| 2 | +Copyright (c) 2018 Mario Carneiro. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Author: Mario Carneiro |
| 5 | +
|
| 6 | +More partial recursive functions using a universal program; |
| 7 | +Rice's theorem and the halting problem. |
| 8 | +-/ |
| 9 | +import data.computability.partrec_code |
| 10 | + |
| 11 | +open encodable denumerable |
| 12 | + |
| 13 | +namespace nat.partrec |
| 14 | +open computable roption |
| 15 | + |
| 16 | +theorem merge' {f g} |
| 17 | + (hf : nat.partrec f) (hg : nat.partrec g) : |
| 18 | + ∃ h, nat.partrec h ∧ ∀ a, |
| 19 | + (∀ x ∈ h a, x ∈ f a ∨ x ∈ g a) ∧ |
| 20 | + ((h a).dom ↔ (f a).dom ∨ (g a).dom) := |
| 21 | +begin |
| 22 | + rcases code.exists_code.1 hf with ⟨cf, rfl⟩, |
| 23 | + rcases code.exists_code.1 hg with ⟨cg, rfl⟩, |
| 24 | + have : nat.partrec (λ n, |
| 25 | + (nat.rfind_opt (λ k, cf.evaln k n <|> cg.evaln k n))) := |
| 26 | + partrec.nat_iff.1 (partrec.rfind_opt $ |
| 27 | + primrec.option_orelse.to_comp.comp |
| 28 | + (code.evaln_prim.to_comp.comp $ (snd.pair (const cf)).pair fst) |
| 29 | + (code.evaln_prim.to_comp.comp $ (snd.pair (const cg)).pair fst)), |
| 30 | + refine ⟨_, this, λ n, _⟩, |
| 31 | + suffices, refine ⟨this, |
| 32 | + ⟨λ h, (this _ ⟨h, rfl⟩).imp Exists.fst Exists.fst, _⟩⟩, |
| 33 | + { intro h, rw nat.rfind_opt_dom, |
| 34 | + simp [dom_iff_mem, code.evaln_complete] at h, |
| 35 | + rcases h with ⟨x, k, e⟩ | ⟨x, k, e⟩, |
| 36 | + { refine ⟨k, x, _⟩, simp [e] }, |
| 37 | + { refine ⟨k, _⟩, |
| 38 | + cases cf.evaln k n with y, |
| 39 | + { exact ⟨x, by simp [e]⟩ }, |
| 40 | + { exact ⟨y, by simp⟩ } } }, |
| 41 | + { intros x h, |
| 42 | + rcases nat.rfind_opt_spec h with ⟨k, e⟩, |
| 43 | + revert e, |
| 44 | + simp; cases e' : cf.evaln k n with y; simp; intro, |
| 45 | + { exact or.inr (code.evaln_sound e) }, |
| 46 | + { subst y, |
| 47 | + exact or.inl (code.evaln_sound e') } } |
| 48 | +end |
| 49 | + |
| 50 | +end nat.partrec |
| 51 | + |
| 52 | +namespace partrec |
| 53 | +variables {α : Type*} {β : Type*} {γ : Type*} {σ : Type*} |
| 54 | +variables [primcodable α] [primcodable β] [primcodable γ] [primcodable σ] |
| 55 | + |
| 56 | +open computable roption nat.partrec (code) nat.partrec.code |
| 57 | + |
| 58 | +theorem merge' {f g : α →. σ} |
| 59 | + (hf : partrec f) (hg : partrec g) : |
| 60 | + ∃ k : α →. σ, partrec k ∧ ∀ a, |
| 61 | + (∀ x ∈ k a, x ∈ f a ∨ x ∈ g a) ∧ |
| 62 | + ((k a).dom ↔ (f a).dom ∨ (g a).dom) := |
| 63 | +let ⟨k, hk, H⟩ := |
| 64 | + nat.partrec.merge' (bind_decode2_iff.1 hf) (bind_decode2_iff.1 hg) in |
| 65 | +begin |
| 66 | + let k' := λ a, (k (encode a)).bind (λ n, decode σ n), |
| 67 | + refine ⟨k', ((nat_iff.2 hk).comp computable.encode).bind |
| 68 | + (computable.decode.of_option.comp snd).to₂, λ a, _⟩, |
| 69 | + suffices, refine ⟨this, |
| 70 | + ⟨λ h, (this _ ⟨h, rfl⟩).imp Exists.fst Exists.fst, _⟩⟩, |
| 71 | + { intro h, simp [k'], |
| 72 | + have hk : (k (encode a)).dom := |
| 73 | + (H _).2.2 (by simpa [encodek2] using h), |
| 74 | + existsi hk, |
| 75 | + cases (H _).1 _ ⟨hk, rfl⟩ with h h; |
| 76 | + { simp at h, |
| 77 | + rcases h with ⟨a', ha', y, hy, e⟩, |
| 78 | + simp [e.symm, encodek] } }, |
| 79 | + { intros x h', simp [k'] at h', |
| 80 | + rcases h' with ⟨n, hn, hx⟩, |
| 81 | + have := (H _).1 _ hn, simp [mem_decode2] at this, |
| 82 | + cases this with h h; |
| 83 | + { rcases h with ⟨a', ⟨ha₁, ha₂⟩, y, hy, rfl⟩, |
| 84 | + rw encodek at hx ha₁, simp at hx ha₁, substs y a', |
| 85 | + simp [hy] } }, |
| 86 | +end |
| 87 | + |
| 88 | +theorem merge {f g : α →. σ} |
| 89 | + (hf : partrec f) (hg : partrec g) |
| 90 | + (H : ∀ a (x ∈ f a) (y ∈ g a), x = y) : |
| 91 | + ∃ k : α →. σ, partrec k ∧ ∀ a x, x ∈ k a ↔ x ∈ f a ∨ x ∈ g a := |
| 92 | +let ⟨k, hk, K⟩ := merge' hf hg in |
| 93 | +⟨k, hk, λ a x, ⟨(K _).1 _, λ h, begin |
| 94 | + have : (k a).dom := (K _).2.2 (h.imp Exists.fst Exists.fst), |
| 95 | + refine ⟨this, _⟩, |
| 96 | + cases h with h h; cases (K _).1 _ ⟨this, rfl⟩ with h' h', |
| 97 | + { exact mem_unique h' h }, |
| 98 | + { exact (H _ _ h _ h').symm }, |
| 99 | + { exact H _ _ h' _ h }, |
| 100 | + { exact mem_unique h' h } |
| 101 | +end⟩⟩ |
| 102 | + |
| 103 | +theorem cond {c : α → bool} {f : α →. σ} {g : α →. σ} |
| 104 | + (hc : computable c) (hf : partrec f) (hg : partrec g) : |
| 105 | + partrec (λ a, cond (c a) (f a) (g a)) := |
| 106 | +let ⟨cf, ef⟩ := exists_code.1 hf, |
| 107 | + ⟨cg, eg⟩ := exists_code.1 hg in |
| 108 | +((eval_part.comp |
| 109 | + (computable.cond hc (const cf) (const cg)) computable.id).bind |
| 110 | + ((@computable.decode σ _).comp snd).of_option.to₂).of_eq $ |
| 111 | +λ a, by cases c a; simp [ef, eg, encodek] |
| 112 | + |
| 113 | +theorem sum_cases |
| 114 | + {f : α → β ⊕ γ} {g : α → β →. σ} {h : α → γ →. σ} |
| 115 | + (hf : computable f) (hg : partrec₂ g) (hh : partrec₂ h) : |
| 116 | + @partrec _ σ _ _ (λ a, sum.cases_on (f a) (g a) (h a)) := |
| 117 | +option_some_iff.1 $ (cond |
| 118 | + (sum_cases hf (const tt).to₂ (const ff).to₂) |
| 119 | + (sum_cases_left hf (option_some_iff.2 hg).to₂ (const option.none).to₂) |
| 120 | + (sum_cases_right hf (const option.none).to₂ (option_some_iff.2 hh).to₂)) |
| 121 | +.of_eq $ λ a, by cases f a; simp |
| 122 | + |
| 123 | +end partrec |
| 124 | + |
| 125 | +def computable_pred {α} [primcodable α] (p : α → Prop) := |
| 126 | +∃ [D : decidable_pred p], |
| 127 | +by exactI computable (λ a, to_bool (p a)) |
| 128 | + |
| 129 | +/- recursively enumerable predicate -/ |
| 130 | +def re_pred {α} [primcodable α] (p : α → Prop) := |
| 131 | +partrec (λ a, roption.assert (p a) (λ _, roption.some ())) |
| 132 | + |
| 133 | +theorem computable_pred.of_eq {α} [primcodable α] |
| 134 | + {p q : α → Prop} |
| 135 | + (hp : computable_pred p) (H : ∀ a, p a ↔ q a) : computable_pred q := |
| 136 | +(funext (λ a, propext (H a)) : p = q) ▸ hp |
| 137 | + |
| 138 | +namespace computable_pred |
| 139 | +variables {α : Type*} {σ : Type*} |
| 140 | +variables [primcodable α] [primcodable σ] |
| 141 | +open nat.partrec (code) nat.partrec.code computable |
| 142 | + |
| 143 | +theorem rice (C : set (ℕ →. ℕ)) |
| 144 | + (h : computable_pred (λ c, eval c ∈ C)) |
| 145 | + {f g} (hf : nat.partrec f) (hg : nat.partrec g) |
| 146 | + (fC : f ∈ C) : g ∈ C := |
| 147 | +begin |
| 148 | + cases h with _ h, resetI, |
| 149 | + rcases fixed_point₂ (partrec.cond (h.comp fst) |
| 150 | + ((partrec.nat_iff.2 hg).comp snd).to₂ |
| 151 | + ((partrec.nat_iff.2 hf).comp snd).to₂).to₂ with ⟨c, e⟩, |
| 152 | + simp at e, |
| 153 | + by_cases eval c ∈ C, |
| 154 | + { simp [h] at e, rwa ← e }, |
| 155 | + { simp at h, simp [h] at e, |
| 156 | + rw e at h, contradiction } |
| 157 | +end |
| 158 | + |
| 159 | +theorem rice₂ (C : set code) |
| 160 | + (H : ∀ cf cg, eval cf = eval cg → (cf ∈ C ↔ cg ∈ C)) : |
| 161 | + computable_pred (λ c, c ∈ C) ↔ C = ∅ ∨ C = set.univ := |
| 162 | +by haveI := classical.dec; exact |
| 163 | +have hC : ∀ f, f ∈ C ↔ eval f ∈ eval '' C, |
| 164 | +from λ f, ⟨set.mem_image_of_mem _, λ ⟨g, hg, e⟩, (H _ _ e).1 hg⟩, |
| 165 | +⟨λ h, or_iff_not_imp_left.2 $ λ C0, |
| 166 | + set.eq_univ_of_forall $ λ cg, |
| 167 | + let ⟨cf, fC⟩ := set.ne_empty_iff_exists_mem.1 C0 in |
| 168 | + (hC _).2 $ rice (eval '' C) (h.of_eq hC) |
| 169 | + (partrec.nat_iff.1 $ eval_part.comp (const cf) computable.id) |
| 170 | + (partrec.nat_iff.1 $ eval_part.comp (const cg) computable.id) |
| 171 | + ((hC _).1 fC), |
| 172 | +λ h, by rcases h with rfl | rfl; simp [computable_pred]; |
| 173 | + exact ⟨by apply_instance, computable.const _⟩⟩ |
| 174 | + |
| 175 | +theorem halting_problem (n) : ¬ computable_pred (λ c, (eval c n).dom) |
| 176 | +| h := rice {f | (f n).dom} h nat.partrec.zero nat.partrec.none trivial |
| 177 | + |
| 178 | +end computable_pred |
0 commit comments