Skip to content

Commit c20f773

Browse files
Merge pull request #4 from HermesMarc/formulas
Tennenbaum's Theorem
2 parents 0fe6a74 + 578300f commit c20f773

16 files changed

+1734
-1827
lines changed
Lines changed: 68 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,68 @@
1+
Require Import Lia.
2+
3+
From FOL.Tennenbaum Require Import DN_Utils.
4+
5+
Section Coding.
6+
7+
(** We want to capture more abstractly what is needed for the coding
8+
theorem to hold. Originally we used divisibility for numbers, but we
9+
can more generally assume that there is some binary predicate satisfying
10+
the following two axioms.
11+
*)
12+
Variable In : nat -> nat -> Prop.
13+
14+
Notation "u ∈ c" := (In u c) (at level 45).
15+
16+
Hypothesis H_empty : exists e, forall x, ~ x ∈ e.
17+
Hypothesis H_extend : forall n c, exists c', forall x, x ∈ c' <-> x = n \/ x ∈ c.
18+
19+
20+
(** We can now prove that this binary relation can be used for encoding predicates.
21+
*)
22+
Lemma bounded_coding p n :
23+
(forall x, x < n -> p x \/ ~ p x) ->
24+
exists c, forall u, (u < n -> (p u <-> u ∈ c)) /\ (u ∈ c -> u < n).
25+
Proof.
26+
assert (forall u n, u < S n -> u < n \/ u = n) as E by lia.
27+
induction n.
28+
{ intros. destruct H_empty as [e He].
29+
exists e. intros u. split; [lia|]. intros []%He. }
30+
intros Dec_p.
31+
destruct IHn as [c Hc], (Dec_p n ltac:(lia)) as [pn|pn'].
32+
1, 2: intros; apply Dec_p; lia.
33+
- destruct (H_extend n c) as [c' Hc']. exists c'.
34+
intros u. split.
35+
+ intros [| ->]%E. split.
36+
* intros p_u%Hc; auto. apply Hc'; auto.
37+
* intros [->|]%Hc'; auto. now apply Hc.
38+
* split; eauto. intros _. apply Hc'; auto.
39+
+ intros [->|H%Hc]%Hc'; auto.
40+
- exists c.
41+
intros u; split.
42+
+ intros [| ->]%E; [now apply Hc|].
43+
split; [now intros ?%pn'|].
44+
intros H%Hc. lia.
45+
+ intros H%Hc. lia.
46+
Qed.
47+
48+
Lemma DN_bounded_lem p n :
49+
~ ~ (forall x, x < n -> p x \/ ~ p x).
50+
Proof.
51+
induction n as [|n IH].
52+
{ DN.ret. lia. }
53+
DN.lem (p n). intros Hn.
54+
DN.bind IH. DN.ret. intros x Hx.
55+
assert (x < n \/ x = n) as [| ->] by lia.
56+
all: auto.
57+
Qed.
58+
59+
Corollary DN_coding p n :
60+
~~ exists c, forall u, (u < n -> (p u <-> u ∈ c)) /\ (u ∈ c -> u < n).
61+
Proof.
62+
eapply DN.bind_.
63+
- apply DN_bounded_lem.
64+
- intro. apply DN.ret_, (bounded_coding p n); eauto.
65+
Qed.
66+
67+
68+
End Coding.

theories/Tennenbaum/CantorPairing.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ Proof.
1818
intros. apply IHn. lia.
1919
Defined.
2020

21-
21+
(** We define the Cantor pairing function and show that it is a Bijection. *)
2222
Section Cantor.
2323

2424
Definition next '(x,y) :=

theories/Tennenbaum/Church.v

Lines changed: 80 additions & 152 deletions
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,7 @@
11
From FOL Require Import FullSyntax Arithmetics Theories.
22
From Undecidability.Shared Require Import ListAutomation.
3-
From FOL.Tennenbaum Require Import NumberUtils Formulas SyntheticInType Peano.
3+
From FOL.Tennenbaum Require Import NumberUtils DN_Utils Formulas SyntheticInType Peano.
4+
From FOL.Incompleteness Require Import qdec sigma1 ctq.
45

56
(* Require Import FOL Tarski Deduction Peano Formulas NumberTheory Synthetic DecidabilityFacts. *)
67
Require Import Lia.
@@ -15,96 +16,107 @@ Notation "x ∣ y" := (exists k, x * k = y) (at level 50).
1516
Definition unary α := bounded 1 α.
1617
Definition binary α := bounded 2 α.
1718

18-
19-
2019
Section ChurchThesis.
20+
Existing Instance PA_preds_signature.
21+
Existing Instance PA_funcs_signature.
2122

23+
Context {peirce : peirce}.
2224
Instance ff : falsity_flag := falsity_on.
23-
Context {Δ1 : Delta1}.
2425

2526
(** * Church's Thesis *)
2627

2728
(** ** CT_Q *)
2829

29-
(* CT_Q internalizes computability, by claiming that every function
30+
(** CT_Q internalizes computability, by claiming that every function
3031
nat -> nat is computable in a model of computation. In this case,
3132
the model of computaion is based on arithmetic: It represents
32-
functions by formulas in the language of PA in such a way that a
33-
weak fragment (Q <<= PA) can prove the agreement of the formula
34-
with the function on every input.
33+
functions by formulas in the language of PA in such a way that a
34+
weak fragment (Q PA) can prove the agreement of the formula
35+
with the function on every input.
3536
*)
3637

37-
Definition represents ϕ f := forall x, Q ⊢I ∀ ϕ[up (num x)..] ↔ $0 == num (f x).
38-
Definition CT_Q := forall f : nat -> nat, exists ϕ, bounded 3 ϕ /\ inhabited(delta1 ϕ) /\ represents (∃ ϕ) f.
38+
Notation Q := Qeq.
3939

40-
(* Weaker Version of CT_Q, where the existence of the formula
41-
is only given potentially (i.e. behing a double negation).
40+
Definition represents φ f := forall x, Q ⊢ ∀ φ[num x .: $0..] ↔ $0 == num (f x).
41+
Definition CT_Q := forall f : nat -> nat, exists φ, bounded 2 φ /\ Σ1 φ /\ represents φ f.
42+
(** WCT_Q is a weaker Version of CT_Q, where the existence of the formula
43+
is only given potentially (i.e. behing a double negation).
4244
*)
43-
Definition WCT_Q := forall f : nat -> nat, ~ ~ exists ϕ, bounded 3 ϕ /\ inhabited(delta1 ϕ) /\ represents (∃ ϕ) f.
45+
Definition WCT_Q := forall f : nat -> nat, ~~exists φ, bounded 2 φ /\ Σ1 φ /\ represents φ f.
46+
47+
Lemma prime_div_relation :
48+
CT_Q -> exists Π, bounded 2 Π /\ Σ1 Π /\
49+
forall x, Q ⊢ ∀ Π[num x .: $0..] ↔ $0 == num (Irred x).
50+
Proof.
51+
intros ct. apply ct.
52+
Qed.
4453

4554
(** Strong Representability *)
4655

47-
Definition strong_repr ϕ (p : nat -> Prop) := (forall x, p x -> Q ⊢I ϕ[(num x)..]) /\ (forall x, ~ p x -> Q ⊢I ¬ϕ[(num x)..]).
48-
Definition RT_strong := forall p : nat -> Prop, Dec p -> exists ϕ, bounded 2 ϕ /\ inhabited(delta1 ϕ) /\ strong_repr (∃ ϕ) p.
49-
Definition WRT_strong := forall p : nat -> Prop, Dec p -> ~ ~ exists ϕ, bounded 2 ϕ /\ inhabited(delta1 ϕ) /\ strong_repr (∃ ϕ) p.
56+
Definition strong_repr ϕ (p : nat -> Prop) :=
57+
(forall x, p x -> Q ⊢ ϕ[(num x)..]) /\ (forall x, ~ p x -> Q ⊢ ¬ϕ[(num x)..]).
58+
Definition RT_strong :=
59+
forall p : nat -> Prop, Dec p -> exists ϕ, bounded 1 ϕ /\ Σ1 ϕ /\ strong_repr ϕ p.
60+
Definition WRT_strong :=
61+
forall p : nat -> Prop, Dec p -> ~ ~ exists ϕ, bounded 1 ϕ /\ Σ1 ϕ /\ strong_repr ϕ p.
5062

5163
(** Weak Representability *)
5264

53-
Definition weak_repr ϕ (p : nat -> Prop) := (forall x, p x <-> Q ⊢I ϕ[(num x)..]).
54-
Definition RT_weak := forall p : nat -> Prop, enumerable p -> exists ϕ, bounded 3 ϕ /\ inhabited(delta1 ϕ) /\ weak_repr (∃∃ ϕ) p.
55-
Definition WRT_weak := forall p : nat -> Prop, enumerable p -> ~ ~ exists ϕ, bounded 3 ϕ /\ inhabited(delta1 ϕ) /\ weak_repr (∃∃ ϕ) p.
65+
Definition weak_repr ϕ (p : nat -> Prop) := (forall x, p x <-> Q ⊢ ϕ[(num x)..]).
66+
Definition RT_weak :=
67+
forall p : nat -> Prop, enumerable p -> exists ϕ, bounded 1 ϕ /\ Σ1 ϕ /\ weak_repr ϕ p.
68+
Definition WRT_weak :=
69+
forall p : nat -> Prop, enumerable p -> ~ ~ exists ϕ, bounded 1 ϕ /\ Σ1 ϕ /\ weak_repr ϕ p.
5670

5771
Definition RT := RT_strong /\ RT_weak.
5872

59-
Lemma prv_split {p : peirce} α β Gamma :
60-
Gamma ⊢ α ↔ β <-> Gamma ⊢ (α → β) /\ Gamma ⊢ (β → α).
73+
Lemma prv_split α β Γ :
74+
Γ ⊢ α ↔ β <-> Γ ⊢ (α → β) /\ Γ ⊢ (β → α).
6175
Proof.
62-
split; intros H.
76+
split; intros H.
6377
- split.
6478
+ eapply CE1. apply H.
6579
+ eapply CE2. apply H.
6680
- now constructor.
6781
Qed.
6882

69-
Lemma up_switch ϕ s t :
70-
ϕ[up (num s)..][up (num t)..] = ϕ[up (up (num t)..)][up (num s)..].
71-
Proof.
72-
rewrite !subst_comp. apply subst_ext.
73-
intros [|[|[]]]; cbn; now rewrite ?num_subst.
74-
Qed.
75-
7683
Lemma CT_WCT :
7784
CT_Q -> WCT_Q.
7885
Proof. intros ct f. specialize (ct f). tauto. Qed.
7986

87+
Hint Resolve CT_WCT : core.
88+
8089
(** ** Strong part of the representability theorem. *)
8190
Lemma CT_RTs :
8291
CT_Q -> RT_strong.
8392
Proof.
8493
intros ct p Dec_p.
8594
destruct (Dec_decider_nat _ Dec_p) as [f Hf].
86-
destruct (ct f) as [ϕ [b2 [[s1] H]]].
87-
pose (Φ := ϕ[up (num 0)..]).
88-
exists Φ. split; unfold Φ.
89-
{ asimpl. eapply subst_bound; eauto.
90-
intros [|[|[]]] G; cbn; try lia. all: solve_bounds. }
91-
split.
92-
{ constructor. now apply delta1_subst. }
95+
destruct (ct f) as [ϕ (bϕ2 & Sϕ & Hϕ)].
96+
exists ϕ[$0 .: (num 0) ..].
97+
assert (forall a b c, ϕ[a .: b .: c..] = ϕ[a .: b..]) as Help.
98+
{ intros *. eapply bounded_subst; eauto.
99+
intros [|[]]; simpl; lia || reflexivity. }
93100
repeat split.
94-
all: intros x; specialize (H x).
95-
all: eapply AllE with (t := num 0) in H; cbn -[Q] in H.
96-
all: apply prv_split in H; destruct H as [H1 H2].
101+
{ eapply subst_bound; eauto.
102+
intros [|[]] ?; try lia; simpl; solve_bounds. }
103+
{ now apply Σ1_subst. }
104+
all: intros x; specialize (Hϕ x).
105+
all: eapply AllE with (t := num 0) in Hϕ; cbn -[Q] in Hϕ.
106+
all: asimpl; asimpl in Hϕ.
107+
all: rewrite !num_subst in *; rewrite Help in *.
108+
all: apply prv_split in Hϕ; destruct Hϕ as [H1 H2].
97109
- intros px%Hf. symmetry in px.
98-
eapply num_eq with (Gamma := Q)(p := intu) in px; [|firstorder].
99-
eapply IE. cbn -[Q num]. rewrite up_switch.
100-
apply H2. now rewrite num_subst.
110+
eapply num_eq with (Gamma := Q) in px; [|firstorder].
111+
eapply IE. cbn -[Qeq num].
112+
apply H2. apply px.
101113
- intros npx. assert (0 <> f x) as E by firstorder.
102-
apply num_neq with (Gamma := Q)(p := intu) in E; [|firstorder].
114+
apply num_neq with (Gamma := Q)in E; [|firstorder].
103115
apply II. eapply IE.
104116
{eapply Weak; [apply E|now right]. }
105117
eapply IE; [|apply Ctx; now left].
106-
rewrite num_subst in H1. eapply Weak.
107-
+ cbn -[Q num]. rewrite up_switch. apply H1.
118+
eapply Weak.
119+
+ cbn -[Q num]. apply H1.
108120
+ now right.
109121
Qed.
110122

@@ -113,128 +125,44 @@ Lemma WCT_WRTs :
113125
WCT_Q -> WRT_strong.
114126
Proof.
115127
intros wct p Dec_p.
116-
destruct (Dec_decider_nat _ Dec_p) as [f Hf].
117-
apply (DN_chaining (wct f)), DN. intros [ϕ [b2 [[s1] H]]].
118-
pose (Φ := ϕ[up (num 0)..]).
119-
exists Φ. split; unfold Φ.
120-
{ eapply subst_bound; eauto.
121-
intros [|[|[]]] ?; cbn; try lia. all: solve_bounds. }
122-
split.
123-
{ constructor. now apply delta1_subst. }
128+
destruct (Dec_decider_nat _ Dec_p) as [f Hf].
129+
apply (DN.bind_ (wct f)).
130+
intros [ϕ (bϕ2 & Sϕ & Hϕ)]. DN.ret.
131+
exists ϕ[$0 .: (num 0) ..].
132+
assert (forall a b c, ϕ[a .: b .: c..] = ϕ[a .: b..]) as Help.
133+
{ intros *. eapply bounded_subst; eauto.
134+
intros [|[]]; simpl; lia || reflexivity. }
124135
repeat split.
125-
all: intros x; specialize (H x).
126-
all: eapply AllE with (t := num 0) in H; cbn -[Q] in H.
127-
all: apply prv_split in H; destruct H as [H1 H2].
136+
{ eapply subst_bound; eauto.
137+
intros [|[]] ?; try lia; simpl; solve_bounds. }
138+
{ now apply Σ1_subst. }
139+
all: intros x; specialize (Hϕ x).
140+
all: eapply AllE with (t := num 0) in Hϕ; cbn -[Q] in Hϕ.
141+
all: asimpl; asimpl in Hϕ.
142+
all: rewrite !num_subst in *; rewrite Help in *.
143+
all: apply prv_split in Hϕ; destruct Hϕ as [H1 H2].
128144
- intros px%Hf. symmetry in px.
129-
eapply num_eq with (Gamma := Q)(p := intu) in px; [|firstorder].
130-
eapply IE. cbn -[Q num]. rewrite up_switch.
131-
apply H2. now rewrite num_subst.
145+
eapply num_eq with (Gamma := Q) in px; [|firstorder].
146+
eapply IE. cbn -[Q num].
147+
apply H2. apply px.
132148
- intros npx. assert (0 <> f x) as E by firstorder.
133-
apply num_neq with (Gamma := Q)(p := intu) in E; [|firstorder].
149+
apply num_neq with (Gamma := Q) in E; [|firstorder].
134150
apply II. eapply IE.
135151
{eapply Weak; [apply E|now right]. }
136152
eapply IE; [|apply Ctx; now left].
137-
rewrite num_subst in H1. eapply Weak.
138-
+ cbn -[Q num]. rewrite up_switch. apply H1.
153+
eapply Weak.
154+
+ cbn -[Q num]. apply H1.
139155
+ now right.
140156
Qed.
141157

142158
(** ** Weak part of the representability theorem. *)
143159
Lemma CT_RTw :
144160
CT_Q -> RT_weak.
145161
Proof.
146-
intros ct p [f Hf]%enumerable_nat.
147-
destruct (ct f) as [ϕ [b2 [[s1] H]]].
148-
pose (Φ := ϕ[up (σ $1)..] ).
149-
exists Φ; split. 2: split.
150-
- unfold Φ. eapply subst_bound; eauto.
151-
intros [|[|[]]] ?; try lia; repeat solve_bounds.
152-
- constructor. now apply delta1_subst.
153-
- intros x. rewrite Hf. split.
154-
+ intros [n Hn]. symmetry in Hn.
155-
apply sigma1_ternary_complete.
156-
{ unfold Φ. eapply subst_bound; eauto.
157-
intros [|[|[]]] ?; try lia; repeat solve_bounds. }
158-
{unfold Φ. now apply delta1_subst. }
159-
exists n. specialize (H n).
160-
apply soundness in H.
161-
unshelve refine (let H := (H nat (extensional_model interp_nat) (fun _ => 0)) _ in _ ).
162-
{apply Q_std_axioms. }
163-
cbn in H. specialize (H (S x)) as [_ H2].
164-
rewrite eval_num, inu_nat_id in *.
165-
apply H2 in Hn. destruct Hn as [d Hd].
166-
exists d.
167-
unfold Φ. rewrite !sat_comp in *.
168-
eapply bound_ext with (N:=3).
169-
apply b2. 2: apply Hd.
170-
intros [|[|[]]]; cbn; rewrite ?num_subst, ?eval_num, ?inu_nat_id in *.
171-
all: try lia.
172-
+ intros HQ%soundness. specialize (HQ nat (extensional_model interp_nat) (fun _ => 0)). cbn in HQ.
173-
destruct HQ as [n [d Hnd]].
174-
{apply Q_std_axioms. }
175-
exists n. specialize (H n).
176-
apply soundness in H.
177-
unshelve refine (let H := (H nat (extensional_model interp_nat) (fun _ => 0)) _ in _ ).
178-
apply Q_std_axioms.
179-
cbn in H. specialize (H (S x)) as [H1 _].
180-
rewrite eval_num, inu_nat_id in *.
181-
symmetry. apply H1. exists d.
182-
rewrite !sat_comp in Hnd.
183-
unfold Φ in Hnd. rewrite sat_comp in *.
184-
eapply bound_ext.
185-
apply b2. 2: apply Hnd.
186-
intros [|[|[]]]; cbn; rewrite ?num_subst, ?eval_num, ?inu_nat_id in *.
187-
all: try lia; reflexivity.
162+
intros ct p Hp. unfold weak_repr.
163+
apply ctq_weak_repr; auto.
188164
Qed.
189165

190166

191-
Lemma WCT_WRTw :
192-
WCT_Q -> WRT_weak.
193-
Proof.
194-
intros wct p [f Hf]%enumerable_nat.
195-
apply (DN_chaining (wct f)), DN.
196-
intros [ϕ [b2 [[s1] H]]].
197-
pose (Φ := ϕ[up (σ $1)..] ).
198-
exists Φ; split. 2: split.
199-
- unfold Φ. eapply subst_bound; eauto.
200-
intros [|[|[]]] ?; try lia; repeat solve_bounds.
201-
- constructor. now apply delta1_subst.
202-
- intros x. rewrite Hf. split.
203-
+ intros [n Hn]. symmetry in Hn.
204-
apply sigma1_ternary_complete.
205-
{ unfold Φ. eapply subst_bound; eauto.
206-
intros [|[|[]]] ?; try lia; repeat solve_bounds. }
207-
{unfold Φ. now apply delta1_subst. }
208-
exists n. specialize (H n).
209-
apply soundness in H.
210-
unshelve refine (let H := (H nat (extensional_model interp_nat) (fun _ => 0)) _ in _ ).
211-
{apply Q_std_axioms. }
212-
cbn in H. specialize (H (S x)) as [_ H2].
213-
rewrite eval_num, inu_nat_id in *.
214-
apply H2 in Hn. destruct Hn as [d Hd].
215-
exists d.
216-
unfold Φ. rewrite !sat_comp in *.
217-
eapply bound_ext with (N:=3).
218-
apply b2. 2: apply Hd.
219-
intros [|[|[]]]; cbn; rewrite ?num_subst, ?eval_num, ?inu_nat_id in *.
220-
all: try lia; reflexivity.
221-
+ intros HQ%soundness. specialize (HQ nat (extensional_model interp_nat) (fun _ => 0)). cbn in HQ.
222-
destruct HQ as [n [d Hnd]].
223-
{apply Q_std_axioms. }
224-
exists n. specialize (H n).
225-
apply soundness in H.
226-
unshelve refine (let H := (H nat (extensional_model interp_nat) (fun _ => 0)) _ in _ ).
227-
apply Q_std_axioms.
228-
cbn in H. specialize (H (S x)) as [H1 _].
229-
rewrite eval_num, inu_nat_id in *.
230-
symmetry. apply H1. exists d.
231-
rewrite !sat_comp in Hnd.
232-
unfold Φ in Hnd. rewrite sat_comp in *.
233-
eapply bound_ext.
234-
apply b2. 2: apply Hnd.
235-
intros [|[|[]]]; cbn; rewrite ?num_subst, ?eval_num, ?inu_nat_id in *.
236-
all: try lia; reflexivity.
237-
Qed.
238-
239167

240168
End ChurchThesis.

0 commit comments

Comments
 (0)