Skip to content

Commit c3e22cd

Browse files
committed
refactor PostsProblem to only rely on EPF and not other axioms
1 parent 85f7f4c commit c3e22cd

15 files changed

+381
-162
lines changed

theories/Axioms/EA.v

+12-6
Original file line numberDiff line numberDiff line change
@@ -62,11 +62,13 @@ Proof.
6262
- intros H. eapply semi_decidable_enumerable. eauto. eauto.
6363
Qed.
6464

65-
Module Assume_EA.
66-
6765
From SyntheticComputability Require Import reductions.
6866

69-
Axiom EA : EA.
67+
Existing Class EA.
68+
69+
Section AssumeEA.
70+
71+
Context {EA : EA}.
7072

7173
Notation φ := (proj1_sig EA).
7274
Notation EAP := (proj2_sig EA).
@@ -95,7 +97,7 @@ Proof.
9597
eapply W_spec.
9698
Qed.
9799

98-
Global Hint Resolve discrete_nat : core.
100+
Hint Resolve discrete_nat : core.
99101

100102
Lemma EAS' :
101103
forall p : nat -> nat -> Prop, enumerable (fun! ⟨x,y⟩ => p x y) ->
@@ -231,7 +233,7 @@ Proof.
231233
exists (fun '(n,m) => ⟨n,m⟩). intros [n m]. now rewrite embedP.
232234
Qed.
233235

234-
Global Hint Resolve discrete_prod discrete_nat : core.
236+
Hint Resolve discrete_prod discrete_nat : core.
235237

236238
Lemma W_not_enumerable : ~ enumerable (compl (uncurry W)).
237239
Proof.
@@ -302,4 +304,8 @@ Proof.
302304
intros ? (? & <- & ?) % in_map_iff. rewrite embedP. firstorder.
303305
Qed.
304306

305-
End Assume_EA.
307+
End AssumeEA.
308+
309+
Notation "m-complete p" := (forall q : nat -> Prop, enumerable q -> q ⪯ₘ p) (at level 10).
310+
311+
Global Hint Resolve discrete_prod discrete_nat : core.

theories/Basic/Rice.v

+8-1
Original file line numberDiff line numberDiff line change
@@ -2,7 +2,12 @@ From SyntheticComputability Require Import Equivalence EA equiv_on DecidabilityF
22

33
Import EmbedNatNotations.
44

5-
Import Assume_EA.
5+
Section Assume_EA.
6+
7+
Context {EA : EA}.
8+
9+
Notation φ := (proj1_sig EA).
10+
Notation EAP := (proj2_sig EA).
611

712
Set Default Goal Selector "!".
813

@@ -136,3 +141,5 @@ Proof.
136141
firstorder.
137142
- destruct H as [f Hf]. exists (fun c => f (θ c)). firstorder.
138143
Qed.
144+
145+
End Assume_EA.

theories/CRM/principles.v

+6-3
Original file line numberDiff line numberDiff line change
@@ -713,9 +713,12 @@ Proof.
713713
eapply MP.
714714
Qed.
715715

716-
Module assume_EA.
716+
Section Assume_EA.
717717

718-
Import Assume_EA. (* assumes EA as an axiom, i.e. a parametrically universal enumerator *)
718+
Context {EA : EA}.
719+
720+
Notation φ := (proj1_sig EA).
721+
Notation EAP := (proj2_sig EA).
719722

720723
Lemma MP_iff_stable_W :
721724
MP <-> Definitions.stable (uncurry W).
@@ -760,7 +763,7 @@ Module assume_EA.
760763
eapply Hm; eauto.
761764
Qed.
762765

763-
End assume_EA.
766+
End Assume_EA.
764767

765768
Lemma DNE_sdec_to_cosdec :
766769
DNE ->

theories/PostsProblem/limit_computability.v

+27-15
Original file line numberDiff line numberDiff line change
@@ -72,25 +72,35 @@ Convention:
7272
(* Naming the halting problem as K *)
7373
Notation K := (­{0}^(1)).
7474

75+
Section AssumePartiality.
76+
77+
Context {Part : partiality}.
78+
79+
Context {enc : encoding ()}.
80+
81+
Context {EPF_assm : EPF.EPF}.
82+
83+
Section LimitLemma1.
7584

76-
Section LimitLemma1.
7785
(* Limit computable predicate P is reduciable to K *)
7886

79-
Variable vec_to_nat : ∀ k, vec nat k → nat.
80-
Variable nat_to_vec : ∀ k, nat → vec nat k.
81-
Variable vec_nat_inv : ∀ k v, nat_to_vec k (vec_to_nat v) = v.
82-
Variable nat_vec_inv : ∀ k n, vec_to_nat (nat_to_vec k n) = n.
87+
Context {vec_datatype : datatype (vec nat)}.
8388

84-
Variable list_vec_to_nat : ∀ k, list (vec nat k) → nat.
85-
Variable nat_to_list_vec : ∀ k, nat → list (vec nat k).
86-
Variable list_vec_nat_inv : ∀ k v, nat_to_list_vec k (list_vec_to_nat v) = v.
87-
Variable nat_list_vec_inv : ∀ k n, list_vec_to_nat (nat_to_list_vec k n) = n.
89+
Notation vec_to_nat := (@X_to_nat (vec nat) _ _).
90+
Notation nat_to_vec := (@nat_to_X (vec nat) _ _).
91+
Notation vec_nat_inv := (@X_nat_inv (vec nat) _ _).
8892

89-
Variable nat_to_list_bool : nat → list bool.
90-
Variable list_bool_to_nat : list bool → nat.
91-
Variable list_bool_nat_inv : ∀ l, nat_to_list_bool (list_bool_to_nat l) = l.
92-
Variable nat_list_bool_inv : ∀ n, list_bool_to_nat (nat_to_list_bool n) = n.
93+
Context {list_vec_datatype : datatype (fun k => list (vec nat k))}.
9394

95+
Notation list_vec_to_nat := (@X_to_nat (fun k => list (vec nat k)) _ _).
96+
Notation nat_to_list_vec := (@nat_to_X (fun k => list (vec nat k)) _).
97+
Notation list_vec_nat_inv := (@X_nat_inv (fun k => list (vec nat k)) _ _).
98+
99+
Context {list_bool_datatype : datatype (fun _ => list bool)}.
100+
101+
Notation list_bool_to_nat := (@X_to_nat (fun _ => list bool) _ 0).
102+
Notation nat_to_list_bool := (@nat_to_X (fun _ => list bool) _ 0).
103+
Notation list_bool_nat_inv := (@X_nat_inv (fun _ => list bool) _ 0).
94104

95105
Section def_K.
96106

@@ -111,9 +121,9 @@ Section LimitLemma1.
111121
Lemma def_K: definite K.
112122
Proof.
113123
apply semi_dec_def.
114-
assert (isΣsem 1 (@jumpNK _ 1 1)).
124+
assert (isΣsem 1 (@jumpNK _ _ _ 1 1)).
115125
eapply jump_in_Σn; eauto.
116-
assert (@jumpNK _ 1 1 ≡ₘ ­{0}^(1)).
126+
assert (@jumpNK _ _ _ 1 1 ≡ₘ ­{0}^(1)).
117127
apply jumpNKspec.
118128
rewrite <- semi_dec_iff_Σ1 in H.
119129
destruct H0 as [_ [f Hf]].
@@ -397,3 +407,5 @@ Section LimitLemma2.
397407
Qed.
398408

399409
End LimitLemma2.
410+
411+
End AssumePartiality.

theories/PostsProblem/low_simple_predicates.v

+123-34
Original file line numberDiff line numberDiff line change
@@ -32,26 +32,39 @@ Section Facts.
3232
End Facts.
3333

3434

35+
Notation "(¬¬Σ⁰₁)-LEM" :=
36+
((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ¬¬ definite p)) (at level 0).
3537

36-
(* Definition of low *)
37-
Definition low (P: nat → Prop) := P´ ⪯ᴛ K.
38+
Section AssumePartiality.
39+
40+
Context {Part : partial.partiality}.
41+
42+
Context {enc : encoding ()}.
43+
44+
Context {EPF_assm : EPF.EPF}.
45+
46+
(* Definition of low *)
47+
Definition low (P: nat → Prop) := P´ ⪯ᴛ K.
3848

3949
Section LowFacts.
40-
41-
Variable vec_to_nat : ∀ k, vec nat k → nat.
42-
Variable nat_to_vec : ∀ k, nat → vec nat k.
43-
Variable vec_nat_inv : ∀ k v, nat_to_vec k (vec_to_nat v) = v.
44-
Variable nat_vec_inv : ∀ k n, vec_to_nat (nat_to_vec k n) = n.
45-
46-
Variable list_vec_to_nat : ∀ k, list (vec nat k) → nat.
47-
Variable nat_to_list_vec : ∀ k, nat → list (vec nat k).
48-
Variable list_vec_nat_inv : ∀ k v, nat_to_list_vec k (list_vec_to_nat v) = v.
49-
Variable nat_list_vec_inv : ∀ k n, list_vec_to_nat (nat_to_list_vec k n) = n.
50-
51-
Variable nat_to_list_bool : nat → list bool.
52-
Variable list_bool_to_nat : list bool → nat.
53-
Variable list_bool_nat_inv : ∀ l, nat_to_list_bool (list_bool_to_nat l) = l.
54-
Variable nat_list_bool_inv : ∀ n, list_bool_to_nat (nat_to_list_bool n) = n.
50+
51+
Context {vec_datatype : datatype (vec nat)}.
52+
53+
Notation vec_to_nat := (@X_to_nat (vec nat) _ _).
54+
Notation nat_to_vec := (@nat_to_X (vec nat) _ _).
55+
Notation vec_nat_inv := (@X_nat_inv (vec nat) _ _).
56+
57+
Context {list_vec_datatype : datatype (fun k => list (vec nat k))}.
58+
59+
Notation list_vec_to_nat := (@X_to_nat (fun k => list (vec nat k)) _ _).
60+
Notation nat_to_list_vec := (@nat_to_X (fun k => list (vec nat k)) _).
61+
Notation list_vec_nat_inv := (@X_nat_inv (fun k => list (vec nat k)) _ _).
62+
63+
Context {list_bool_datatype : datatype (fun _ => list bool)}.
64+
65+
Notation list_bool_to_nat := (@X_to_nat (fun _ => list bool) _ 0).
66+
Notation nat_to_list_bool := (@nat_to_X (fun _ => list bool) _ 0).
67+
Notation list_bool_nat_inv := (@X_nat_inv (fun _ => list bool) _ 0).
5568

5669
Lemma lowness (P: nat → Prop) :
5770
low P → ¬ K ⪯ᴛ P.
@@ -77,7 +90,7 @@ Section LowFacts.
7790
Proof.
7891
intros LEM defK H IH.
7992
apply lowness with (P := A); [|apply IH].
80-
eapply limit_turing_red_K; eauto. exact 42.
93+
eapply limit_turing_red_K; eauto. Unshelve. exact 42.
8194
Qed.
8295

8396
Definition low_simple P := low P ∧ simple P.
@@ -92,43 +105,34 @@ Section LowFacts.
92105
split; [destruct H2 as [H2 _]; eauto| now apply lowness].
93106
Qed.
94107

95-
96-
97-
(*** Instance of low simple predicate ***)
108+
(*** Instance of low simple predicate ***)
98109

99110
Section LowSimplePredicate.
100111

101-
Variable η: nat → nat → option nat.
102-
Hypothesis EA:
103-
∀ P, semi_decidable P → exists e, ∀ x, P x ↔ ∃ n, η e n = Some x.
104-
105112
Hypothesis LEM_Σ_1: LEM_Σ 1.
106113
Hypothesis def_K: definite K.
107114

108115
Theorem a_sol_Post's_problem: ∃ P, sol_Post's_problem P.
109116
Proof.
110117
eexists. eapply low_simple_correct; split.
111-
- eapply limit_turing_red_K; eauto. exact 42.
118+
- eapply limit_turing_red_K; eauto.
112119
apply jump_P_limit; eauto.
113120
- eapply P_simple.
114121
intros. intros d. apply d.
115122
apply wall_convergence. by unfold wall.
116-
assumption.
123+
assumption. Unshelve. exact 42.
117124
Qed.
118125

119126
End LowSimplePredicate.
120127

121-
Notation "(¬¬Σ⁰₁)-LEM" :=
122-
((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ¬¬ definite p)) (at level 0).
123-
124128
Lemma m_red_K_semi_decidable {n} (P: vec nat n → Prop):
125129
semi_decidable P → P ⪯ₘ K.
126130
Proof.
127131
intros H. unfold K. rewrite <- red_m_iff_semidec_jump_vec.
128-
by apply semi_decidable_OracleSemiDecidable. eauto.
132+
by apply semi_decidable_OracleSemiDecidable.
129133
Qed.
130134

131-
Theorem PostProblem_from_neg_negLPO :
135+
Lemma PostProblem_from_neg_negLPO_aux :
132136
∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM → ¬ K ⪯ᴛ p).
133137
Proof.
134138
exists (P wall).
@@ -159,13 +163,97 @@ Section LowFacts.
159163
by eapply m_red_K_semi_decidable.
160164
}
161165
revert G. apply lowness. red.
162-
eapply limit_turing_red_K; eauto. exact 42.
166+
eapply limit_turing_red_K; eauto. Unshelve. 2: exact 42.
163167
apply jump_P_limit; eauto.
164168
Qed.
165169

166170
End LowFacts.
167171

168-
Check PostProblem_from_neg_negLPO.
172+
End AssumePartiality.
173+
174+
From SyntheticComputability Require Import EnumerabilityFacts ListEnumerabilityFacts.
175+
176+
Theorem PostProblem_from_neg_negLPO {Part : partial.partiality} :
177+
(exists θ, EPF.EPF_for θ) ->
178+
∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM → forall K : nat -> Prop, (forall q : nat -> Prop, semi_decidable q -> q ⪯ₘ K) -> ~ K ⪯ᴛ p).
179+
Proof.
180+
intros [θ EPF].
181+
destruct (EnumerabilityFacts.datatype_retract (nat * list bool)) as [(I & R & HIR) _].
182+
{
183+
split. eapply discrete_iff. econstructor. exact _.
184+
apply enumerableᵗ_prod.
185+
eapply enumerableᵗ_nat.
186+
apply enum_enumT.
187+
apply enumerable_list. apply enum_enumT. eapply enumerableᵗ_bool.
188+
}
189+
destruct (EnumerabilityFacts.datatype_retract (list bool)) as [(I2 & R2 & HIR2) _].
190+
{
191+
split. eapply discrete_iff. econstructor. exact _.
192+
apply enum_enumT.
193+
apply enumerable_list. apply enum_enumT. eapply enumerableᵗ_bool.
194+
}
195+
unshelve edestruct @PostProblem_from_neg_negLPO_aux as (p & undec & semidec & H).
196+
- assumption.
197+
- unshelve econstructor.
198+
exact I. exact (fun x => match x with inl n => S n | inr _ => 0 end).
199+
exact (fun n => match R n with None => (0, []) | Some x => x end).
200+
exact (fun v => match v with 0 => inr tt | S n => inl n end).
201+
cbn. intros n.
202+
now destruct (HIR n) as [-> _].
203+
intros []. reflexivity. now destruct u.
204+
- exists θ. assumption.
205+
- unshelve econstructor.
206+
3: eapply VectorEmbedding.vec_nat_inv.
207+
- eassert (forall k, enumeratorᵗ _ (list (vec nat k))).
208+
{
209+
intros k. eapply list_enumeratorᵗ_enumeratorᵗ.
210+
eapply enumeratorᵗ_list.
211+
eapply enumeratorᵗ_of_list.
212+
Unshelve.
213+
2:{ intros n. apply Some. apply (VectorEmbedding.nat_to_vec k n). }
214+
red. intros. exists (VectorEmbedding.vec_to_nat x).
215+
now rewrite VectorEmbedding.vec_nat_inv.
216+
}
217+
eassert (forall k, decider _ (eq_on (list (vec nat k)))).
218+
{
219+
intros k. apply decider_eq_list.
220+
Unshelve.
221+
2:{ intros [l1 l2].
222+
exact (VectorEmbedding.vec_to_nat l1 =? VectorEmbedding.vec_to_nat l2). }
223+
cbn.
224+
red. intros [l1 l2].
225+
red.
226+
split. intros ->.
227+
eapply Nat.eqb_refl.
228+
intros Heq.
229+
destruct (Nat.eqb_spec (VectorEmbedding.vec_to_nat l1) (VectorEmbedding.vec_to_nat l2)).
230+
eapply (f_equal (VectorEmbedding.nat_to_vec k)) in e.
231+
now rewrite !VectorEmbedding.vec_nat_inv in e.
232+
congruence.
233+
}
234+
unshelve econstructor.
235+
+ intros k.
236+
specialize (H k). eapply enumerator_retraction in H as [I3 HI].
237+
exact I3. eapply (H0 k).
238+
+ intros k. specialize (H k). eapply enumerator_retraction in H as [I3 HI].
239+
set (fun! ⟨ n, m ⟩ => nth_error (L_list (λ n0 : nat, [VectorEmbedding.nat_to_vec k n0]) n) m) as R3.
240+
refine (fun n => match R3 n with None => [] | Some x => x end). eapply (H0 k).
241+
+ cbn. intros. destruct enumerator_retraction.
242+
red in r. now rewrite r.
243+
- unshelve econstructor. exact (fun _ => I2).
244+
exact (fun _ n => match R2 n with None => [] | Some x => x end).
245+
cbn. intros _ n.
246+
now destruct (HIR2 n) as [-> _].
247+
- exists p. repeat split. assumption. assumption.
248+
intros lpo K HK Hp.
249+
apply H. assumption.
250+
eapply Turing_transitive.
251+
eapply red_m_impl_red_T.
252+
eapply HK. eapply semi_dec_halting.
253+
assumption.
254+
Qed.
255+
256+
Check @PostProblem_from_neg_negLPO.
169257
Print Assumptions PostProblem_from_neg_negLPO.
170258

171259
(* general proof that (¬¬Σ⁰₁)-LEM <-> ¬¬(Σ⁰₁)-LEM under many-one complete Σ⁰₁ predicate *)
@@ -194,3 +282,4 @@ Section assume.
194282
Qed.
195283

196284
End assume.
285+

0 commit comments

Comments
 (0)