@@ -32,26 +32,39 @@ Section Facts.
32
32
End Facts.
33
33
34
34
35
+ Notation "(¬¬Σ⁰₁)-LEM" :=
36
+ ((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ¬¬ definite p)) (at level 0).
35
37
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.
38
48
39
49
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).
55
68
56
69
Lemma lowness (P: nat → Prop ) :
57
70
low P → ¬ K ⪯ᴛ P.
@@ -77,7 +90,7 @@ Section LowFacts.
77
90
Proof .
78
91
intros LEM defK H IH.
79
92
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.
81
94
Qed .
82
95
83
96
Definition low_simple P := low P ∧ simple P.
@@ -92,43 +105,34 @@ Section LowFacts.
92
105
split; [destruct H2 as [H2 _]; eauto| now apply lowness].
93
106
Qed .
94
107
95
-
96
-
97
- (*** Instance of low simple predicate ** *)
108
+ (*** Instance of low simple predicate ** *)
98
109
99
110
Section LowSimplePredicate.
100
111
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
-
105
112
Hypothesis LEM_Σ_1: LEM_Σ 1.
106
113
Hypothesis def_K: definite K.
107
114
108
115
Theorem a_sol_Post's_problem: ∃ P, sol_Post's_problem P.
109
116
Proof .
110
117
eexists. eapply low_simple_correct; split.
111
- - eapply limit_turing_red_K; eauto. exact 42.
118
+ - eapply limit_turing_red_K; eauto.
112
119
apply jump_P_limit; eauto.
113
120
- eapply P_simple.
114
121
intros. intros d. apply d.
115
122
apply wall_convergence. by unfold wall.
116
- assumption.
123
+ assumption. Unshelve. exact 42.
117
124
Qed .
118
125
119
126
End LowSimplePredicate.
120
127
121
- Notation "(¬¬Σ⁰₁)-LEM" :=
122
- ((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ¬¬ definite p)) (at level 0).
123
-
124
128
Lemma m_red_K_semi_decidable {n} (P: vec nat n → Prop ):
125
129
semi_decidable P → P ⪯ₘ K.
126
130
Proof .
127
131
intros H. unfold K. rewrite <- red_m_iff_semidec_jump_vec.
128
- by apply semi_decidable_OracleSemiDecidable. eauto.
132
+ by apply semi_decidable_OracleSemiDecidable.
129
133
Qed .
130
134
131
- Theorem PostProblem_from_neg_negLPO :
135
+ Lemma PostProblem_from_neg_negLPO_aux :
132
136
∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM → ¬ K ⪯ᴛ p).
133
137
Proof .
134
138
exists (P wall).
@@ -159,13 +163,97 @@ Section LowFacts.
159
163
by eapply m_red_K_semi_decidable.
160
164
}
161
165
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.
163
167
apply jump_P_limit; eauto.
164
168
Qed .
165
169
166
170
End LowFacts.
167
171
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.
169
257
Print Assumptions PostProblem_from_neg_negLPO.
170
258
171
259
(* general proof that (¬¬Σ⁰₁)-LEM <-> ¬¬(Σ⁰₁)-LEM under many-one complete Σ⁰₁ predicate *)
@@ -194,3 +282,4 @@ Section assume.
194
282
Qed .
195
283
196
284
End assume.
285
+
0 commit comments