Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit d62bf56

Browse files
committed
feat(computability/halting): halting problem
1 parent f0bcba5 commit d62bf56

File tree

7 files changed

+352
-81
lines changed

7 files changed

+352
-81
lines changed

data/computability/halting.lean

Lines changed: 178 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,178 @@
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

data/computability/partrec.lean

Lines changed: 66 additions & 38 deletions
Original file line numberDiff line numberDiff line change
@@ -100,6 +100,40 @@ eq_none_iff.2 $ λ a h,
100100
let ⟨n, h₁, h₂⟩ := rfind_dom'.1 h.fst in
101101
(p0 ▸ h₂ (zero_le _) : (@roption.none bool).dom)
102102

103+
def rfind_opt {α} (f : ℕ → option α) : roption α :=
104+
(rfind (λ n, (f n).is_some)).bind (λ n, f n)
105+
106+
theorem rfind_opt_spec {α} {f : ℕ → option α} {a}
107+
(h : a ∈ rfind_opt f) : ∃ n, a ∈ f n :=
108+
let ⟨n, h₁, h₂⟩ := mem_bind_iff.1 h in ⟨n, mem_coe.1 h₂⟩
109+
110+
theorem rfind_opt_dom {α} {f : ℕ → option α} :
111+
(rfind_opt f).dom ↔ ∃ n a, a ∈ f n :=
112+
⟨λ h, (rfind_opt_spec ⟨h, rfl⟩).imp (λ n h, ⟨_, h⟩),
113+
λ h, begin
114+
have h' : ∃ n, (f n).is_some :=
115+
h.imp (λ n, option.is_some_iff_exists.2),
116+
have s := nat.find_spec h',
117+
have fd : (rfind (λ n, (f n).is_some)).dom :=
118+
⟨nat.find h', by simpa using s.symm, λ _ _, trivial⟩,
119+
refine ⟨fd, _⟩,
120+
have := rfind_spec (get_mem fd),
121+
simp at this ⊢,
122+
cases option.is_some_iff_exists.1 this.symm with a e,
123+
rw e, trivial
124+
end
125+
126+
theorem rfind_opt_mono {α} {f : ℕ → option α}
127+
(H : ∀ {a m n}, m ≤ n → a ∈ f m → a ∈ f n)
128+
{a} : a ∈ rfind_opt f ↔ ∃ n, a ∈ f n :=
129+
⟨rfind_opt_spec, λ ⟨n, h⟩, begin
130+
have h' := rfind_opt_dom.2 ⟨_, _, h⟩,
131+
cases rfind_opt_spec ⟨h', rfl⟩ with k hk,
132+
have := (H (le_max_left _ _) h).symm.trans
133+
(H (le_max_right _ _) hk),
134+
simp at this, simp [this, get_mem]
135+
end
136+
103137
inductive partrec : (ℕ →. ℕ) → Prop
104138
| zero : partrec (pure 0)
105139
| succ : partrec succ
@@ -193,6 +227,9 @@ theorem primrec.to_comp {α σ} [primcodable α] [primcodable σ]
193227
(nat.partrec.ppred.comp (nat.partrec.of_primrec hf)).of_eq $
194228
λ n, by simp; cases decode α n; simp [option.map, option.bind]
195229

230+
theorem primrec₂.to_comp {α β σ} [primcodable α] [primcodable β] [primcodable σ]
231+
{f : α → β → σ} (hf : primrec₂ f) : computable₂ f := hf.to_comp
232+
196233
theorem computable.part {α σ} [primcodable α] [primcodable σ]
197234
{f : α → σ} (hf : computable f) : partrec (f : α →. σ) := hf
198235

@@ -247,6 +284,21 @@ theorem list_append : computable₂ ((++) : list α → list α → list α) :=
247284
theorem list_concat : computable₂ (λ l (a:α), l ++ [a]) := primrec.list_concat.to_comp
248285
theorem list_length : computable (@list.length α) := primrec.list_length.to_comp
249286

287+
protected theorem encode : computable (@encode α _) :=
288+
primrec.encode.to_comp
289+
290+
protected theorem decode : computable (decode α) :=
291+
primrec.decode.to_comp
292+
293+
protected theorem of_nat (α) [denumerable α] : computable (of_nat α) :=
294+
(primrec.of_nat _).to_comp
295+
296+
theorem encode_iff {f : α → σ} : computable (λ a, encode (f a)) ↔ computable f :=
297+
iff.rfl
298+
299+
theorem option_some : computable (@option.some α) :=
300+
primrec.option_some.to_comp
301+
250302
end computable
251303

252304
namespace partrec
@@ -306,6 +358,9 @@ theorem comp {f : β →. σ} {g : α → β}
306358
theorem nat_iff {f : ℕ →. ℕ} : partrec f ↔ nat.partrec f :=
307359
by simp [partrec, map_id']
308360

361+
theorem map_encode_iff {f : α →. σ} : partrec (λ a, (f a).map encode) ↔ partrec f :=
362+
iff.rfl
363+
309364
end partrec
310365

311366
namespace partrec₂
@@ -381,6 +436,11 @@ theorem rfind {p : α → ℕ →. bool} (hp : partrec₂ p) :
381436
cases b; refl
382437
end
383438

439+
theorem rfind_opt {f : α → ℕ → option σ} (hf : computable₂ f) :
440+
partrec (λ a, nat.rfind_opt (f a)) :=
441+
(rfind (primrec.option_is_some.to_comp.comp hf).part.to₂).bind
442+
(of_option hf)
443+
384444
theorem nat_cases_right
385445
{f : α → ℕ} {g : α → σ} {h : α → ℕ →. σ}
386446
(hf : computable f) (hg : computable g) (hh : partrec₂ h) :
@@ -396,51 +456,19 @@ theorem nat_cases_right
396456
exact ⟨⟨this n, H.fst⟩, H.snd⟩ }
397457
end
398458

399-
/-
400-
theorem cond {c : α → bool} {f : α →. σ} {g : α →. σ}
401-
(hc : computable c) (hf : partrec f) (hg : partrec g) :
402-
partrec (λ a, cond (c a) (f a) (g a)) :=
403-
(nat_cases (encode_iff.2 hc) hg (hf.comp fst).to₂).of_eq $
404-
λ a, by cases c a; refl
405-
406-
theorem sum_cases
407-
{f : α → β ⊕ γ} {g : α → β →. σ} {h : α → γ →. σ}
408-
(hf : computable f) (hg : partrec₂ g) (hh : partrec₂ h) :
409-
@partrec _ σ _ _ (λ a, sum.cases_on (f a) (g a) (h a)) :=
410-
(cond (nat_bodd.comp $ encode_iff.2 hf)
411-
(option_map (primrec.decode.comp $ nat_div2.comp $ encode_iff.2 hf) hh)
412-
(option_map (primrec.decode.comp $ nat_div2.comp $ encode_iff.2 hf) hg)).of_eq $
413-
λ a, by cases f a with b c;
414-
simp [nat.div2_bit, nat.bodd_bit, encodek]; refl
415-
416-
theorem fix {α σ} [primcodable α] [primcodable σ]
417-
{f : α →. σ ⊕ α} (hf : partrec f) : partrec (pfun.fix f) :=
418-
begin
419-
have := nat_elim snd fst _,
420-
end
421-
-/
459+
theorem bind_decode2_iff {f : α →. σ} : partrec f ↔
460+
nat.partrec (λ n, roption.bind (decode2 α n) (λ a, (f a).map encode)) :=
461+
⟨λ hf, nat_iff.1 $ (of_option primrec.decode2.to_comp).bind $
462+
(map hf (computable.encode.comp snd).to₂).comp snd,
463+
λ h, map_encode_iff.1 $ by simpa [encodek2]
464+
using (nat_iff.2 h).comp (@computable.encode α _)⟩
422465

423466
end partrec
424467

425468
namespace computable
426469
variables {α : Type*} {β : Type*} {γ : Type*} {σ : Type*}
427470
variables [primcodable α] [primcodable β] [primcodable γ] [primcodable σ]
428471

429-
protected theorem encode : computable (@encode α _) :=
430-
primrec.encode.to_comp
431-
432-
protected theorem decode : computable (decode α) :=
433-
primrec.decode.to_comp
434-
435-
protected theorem of_nat (α) [denumerable α] : computable (of_nat α) :=
436-
(primrec.of_nat _).to_comp
437-
438-
theorem encode_iff {f : α → σ} : computable (λ a, encode (f a)) ↔ computable f :=
439-
iff.rfl
440-
441-
theorem option_some : computable (@option.some α) :=
442-
primrec.option_some.to_comp
443-
444472
theorem option_some_iff {f : α → σ} : computable (λ a, some (f a)) ↔ computable f :=
445473
⟨λ h, encode_iff.1 $ primrec.pred.to_comp.comp $ encode_iff.2 h,
446474
option_some.comp⟩

0 commit comments

Comments
 (0)