From 069b6ff000529bd40cc2d38e9f88c5cec7323276 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 3 Oct 2023 16:34:07 +0200 Subject: [PATCH 01/54] Adds proof of Limit lemma 1 --- theories/Basic/Limit.v | 221 ++++++++++++++++++++++++++ theories/PostsTheorem/TuringJump.v | 12 +- theories/TuringReducibility/SemiDec.v | 19 +++ 3 files changed, 249 insertions(+), 3 deletions(-) create mode 100644 theories/Basic/Limit.v diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v new file mode 100644 index 0000000..43761c8 --- /dev/null +++ b/theories/Basic/Limit.v @@ -0,0 +1,221 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic PostsTheorem reductions SemiDec TuringJump OracleComputability Definitions. + +Require Import Vectors.VectorDef Arith.Compare_dec Lia. +Import Vector.VectorNotations. + +Local Notation vec := Vector.t. + + +(* ########################################################################## *) +(** * Limit Lemma *) +(* ########################################################################## *) + +(** This file contains the definition of limit computable and proves the +Limit Lemma, i.e., limit computable is equivalence to reduciable to halting +problem. + +Convention: + + limit: limit computable + semi_dec(_K): semi decidable (on K) + turing_red_K: turing reduciable to halting problem + char[_X]: extensionality of X + +**) + + +(* definition of limit ciomputable *) + + Definition limit_computable {X} (P: X -> Prop) := + exists f: X -> nat -> bool, forall x, + (P x <-> exists N, forall n, n > N -> f x n = true) /\ + (~ P x <-> exists N, forall n, n > N -> f x n = false). + +(* Naming the halting problem as K *) + Notation K := (­{0}^(1)). + + +Section LimitLemma1. + (* Limit computable predicate P is reduciable to K *) + + Variable vec_to_nat : forall k, vec nat k -> nat. + Variable nat_to_vec : forall k, nat -> vec nat k. + Variable vec_nat_inv : forall k v, nat_to_vec k (vec_to_nat v) = v. + Variable nat_vec_inv : forall k n, vec_to_nat (nat_to_vec k n) = n. + + Variable list_vec_to_nat : forall k, list (vec nat k) -> nat. + Variable nat_to_list_vec : forall k, nat -> list (vec nat k). + Variable list_vec_nat_inv : forall k v, nat_to_list_vec k (list_vec_to_nat v) = v. + Variable nat_list_vec_inv : forall k n, list_vec_to_nat (nat_to_list_vec k n) = n. + + Variable nat_to_list_bool : nat -> list bool. + Variable list_bool_to_nat : list bool -> nat. + Variable list_bool_nat_inv : forall l, nat_to_list_bool (list_bool_to_nat l) = l. + Variable nat_list_bool_inv : forall n, list_bool_to_nat (nat_to_list_bool n) = n. + + (* Extensionality of Σ2, i.e. P t iff ∃ x. ∀ y. f(x, y, t) = true *) + + Lemma char_Σ2 {k: nat} (P: vec nat k -> Prop) : + (exists f: nat -> nat -> vec nat k -> bool, forall x, P x <-> (exists n, forall m, f n m x = true)) -> + isΣsem 2 P. + Proof. + intros [f H]. + eapply isΣsemS_ with (p := fun v => forall y, f (hd v) y (tl v) = true). + eapply isΠsemS_ with (p := fun v => f (hd (tl v)) (hd v) (tl (tl v)) = true). + eapply isΣsem0. all: easy. + Qed. + + Lemma limit_Σ2 {k: nat} (P: vec nat k -> Prop) : + limit_computable P -> isΣsem 2 P /\ isΣsem 2 (compl P). + Proof. + intros [f H]; split; eapply char_Σ2. + - exists (fun N n x => if le_dec n N then true else f x n). + intro w. destruct (H w) as [-> _]; split; intros [N Hn]; exists N. + + intro m. destruct (le_dec m N); try apply Hn; lia. + + intros n He. specialize (Hn n); destruct (le_dec n N); auto; lia. + - exists (fun N n x => if le_dec n N then true else negb (f x n)). + intro w. destruct (H w) as [_ ->]; split; intros [N Hn]; exists N. + + intro m. destruct (le_dec m N); [auto| rewrite (Hn m); lia]. + + intros n He. specialize (Hn n). + destruct (le_dec n N); auto; [lia|destruct (f w n); easy]. + Qed. + + Lemma limit_semi_dec_K {k: nat} (P: vec nat k -> Prop) : + LEM_Σ 1 -> + limit_computable P -> + OracleSemiDecidable K P /\ + OracleSemiDecidable K (compl P). + Proof. + intros LEM H%limit_Σ2. + rewrite <- !(Σ_semi_decidable_jump). + all: eauto. + Qed. + + (** TODO: LEM_Σ 1 <-> definite K **) + (* First part of limit lemma *) + + Lemma limit_turing_red_K {k: nat} (P: vec nat k -> Prop) : + LEM_Σ 1 -> + definite K -> + limit_computable P -> + P ⪯ᴛ K. + Proof. + intros LEM D H % (limit_semi_dec_K LEM); destruct H as [h1 h2]. + apply PT; try assumption. + apply Dec.nat_eq_dec. + Qed. + +(* + (** Move to lowsimple.v **) + + Definition low (P: nat -> Prop) := P´ ⪯ᴛ ­{0}^(1). + + Lemma lowness (P: nat -> Prop) : + low P -> ~ ­{0}^(1) ⪯ᴛ P. + Proof. + intros H IH. + eapply not_turing_red_J with (Q := P). + eapply Turing_transitive; [apply H| easy]. + Qed. + + Lemma limit_jump_lowness (A: nat -> Prop) : + LEM_Σ 1 -> + definite (­{0}^(1)) -> + LimitComputable (A´) -> ~ ­{0}^(1) ⪯ᴛ A. + Proof. + intros LEM defK H IH. + apply lowness with (P := A); [|apply IH]. + pose (P := fun (v: vec nat 1) => A´ (hd v)). + eapply Turing_transitive; [|apply (@limit_turing_red _ P LEM defK)]. + Admitted. + *) + +End LimitLemma1. + + +Section Σ1Approximation. + + (* Turing jump of a trivial decidable problem is semi decidable *) + + Lemma semi_dec_halting : semi_decidable K. + Proof. + eapply OracleSemiDecidable_semi_decidable with (q := ­{0}). + - exists (fun n => match n with | O => true | _ => false end); intros [|n]; easy. + - eapply semidecidable_J. + Qed. + + + (* Stabilizing the semi decider allows the semi decider + to be used as a Σ1 approximation *) + + Definition stable (f: nat -> bool) := forall n m, n <= m -> f n = true -> f m = true. + + Fixpoint stabilize_step {X} (f: X -> nat -> bool) x n := + match n with + | O => false + | S n => if f x n then true else stabilize_step f x n + end. + + Lemma stabilize {X} (P: X -> Prop) : + semi_decidable P -> exists f, semi_decider f P /\ forall x, stable (f x). + Proof. + intros [f Hf]. + exists (fun x n => stabilize_step f x n); split. + - intro x; split; intro h. + rewrite (Hf x) in h. + destruct h as [c Hc]. + now exists (S c); cbn; rewrite Hc. + rewrite (Hf x). + destruct h as [c Hc]. + induction c; cbn in Hc; [congruence|]. + destruct (f x c) eqn: E; [now exists c|now apply IHc]. + - intros x n m Lenm Hn. + induction Lenm; [trivial|]. + cbn; destruct (f x m) eqn: E; [trivial|assumption]. + Qed. + + (* The Σ1 approximation output correct answers for arbitray list of questions *) + + Lemma semi_dec_char_Σ1 {X} (P: X -> Prop) : + semi_decidable P -> + exists K_: X -> nat -> bool, + forall l: list X, + definite P -> + exists c: nat, forall e, List.In e l -> K_ e c = true <-> P e. + Proof. + intros semi_P. + destruct (stabilize semi_P) as [h [Hh HS]]. + exists h; intros l defP. induction l as [|a l [c Hc]]. + - exists 42; eauto. + - destruct (defP a) as [h1| h2]. + + destruct (Hh a) as [H _]. + destruct (H h1) as [N HN]. + exists (max c N); intros e [->| He]. + split; [easy|intros _]. + eapply HS; [|eapply HN]; lia. + rewrite <- Hc; [|assumption]. + split; intro h'. + rewrite Hc; [|exact He]. + unfold semi_decider in Hh. + rewrite Hh; now exists (Nat.max c N). + eapply HS; [|apply h']; lia. + + exists c; intros e [->| He]. + split; [intro h'|easy]. + unfold semi_decider in Hh. + now rewrite Hh; exists c. + rewrite Hc; eauto. + Qed. + +End Σ1Approximation. + + +Section LimitLemma2. + + (* A predicate P is reduciable to K if P is limit computable *) + + Theorem turing_red_K_lim (P: nat -> Prop) : + P ⪯ᴛ K -> limit_computable P. + Proof. + Admitted. + +End LimitLemma2. diff --git a/theories/PostsTheorem/TuringJump.v b/theories/PostsTheorem/TuringJump.v index 4b597c5..5a37607 100644 --- a/theories/PostsTheorem/TuringJump.v +++ b/theories/PostsTheorem/TuringJump.v @@ -275,16 +275,16 @@ Section jump. - unfold J. reflexivity. Qed. + (** Complement not semi-decidable ***) + Lemma not_semidecidable_compl_J Q : ~ oracle_semi_decidable Q (compl (J Q)). Proof. - intros (F & Hcomp & H). + intros (F & Hcomp & H). specialize (surjective Hcomp) as [c Hc]. unfold J in H. specialize (H c). rewrite <- Hc in H. tauto. Qed. - (** Complement not semi-decidable ***) - Definition 𝒥 Q := fun! ⟨c, x⟩ => Ξ c (char_rel Q) x. Lemma jump_gt Q : @@ -308,6 +308,12 @@ Section jump. eapply not_semidecidable_compl_J; eassumption. Qed. + Lemma not_turing_red_J Q: ~ (J Q ⪯ᴛ Q). + Proof. + intros H % Turing_to_sdec_compl. + eapply not_semidecidable_compl_J; eassumption. + Qed. + (** # #*) Lemma J_self_𝒥_m_red: forall Q, (J Q) ⪯ₘ (𝒥 Q). diff --git a/theories/TuringReducibility/SemiDec.v b/theories/TuringReducibility/SemiDec.v index f7128d1..4b8b912 100644 --- a/theories/TuringReducibility/SemiDec.v +++ b/theories/TuringReducibility/SemiDec.v @@ -786,6 +786,25 @@ Proof. - firstorder. Qed. + +Lemma Turing_to_sdec_compl {Part : partiality} {X Y} (q : Y -> Prop) (p : X -> Prop) : + red_Turing p q -> + OracleSemiDecidable q (compl p). +Proof. + intros [F [HF H2]]. + exists (fun R x o => F R x false). split. + - eapply OracleComputable_ext. + eapply computable_bind. eapply HF. + eapply computable_if with (test := snd). + eapply computable_nothing. + eapply computable_ret with (v := tt). + cbn; split. + + intros [[]]; firstorder. + + destruct o. exists false; firstorder. + - firstorder. +Qed. + + Lemma OracleSemiDecidable_refl {Part : partiality} X (Q : X -> Prop) : OracleSemiDecidable Q Q. Proof. From 4458ab845c226ee763752e60cbca960912ec0745 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 3 Oct 2023 16:45:47 +0200 Subject: [PATCH 02/54] Add lowsimple.v --- theories/ReducibilityDegrees/lowsimple.v | 42 ++++++++++++++++++++++++ 1 file changed, 42 insertions(+) create mode 100644 theories/ReducibilityDegrees/lowsimple.v diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v new file mode 100644 index 0000000..cd0523a --- /dev/null +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -0,0 +1,42 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit. + +Require Import Vectors.VectorDef Arith.Compare_dec Lia. +Import Vector.VectorNotations. + +Local Notation vec := Vector.t. + + +(* ########################################################################## *) +(** * Low Simple Set *) +(* ########################################################################## *) + +(** This file contains the definition of low simple set and proves the +essential property of low simple, i.e. Low simple as a solution to +Post's Problem in Turing degree. **) + + (* Definition of low *) + Definition low (P: nat -> Prop) := P´ ⪯ᴛ K. + + +Section LowFacts. + (* If the turing jump of A is low, then A is not reduciable to K *) + + Lemma lowness (P: nat -> Prop) : + low P -> ~ K ⪯ᴛ P. + Proof. + intros H IH. + eapply not_turing_red_J with (Q := P). + eapply Turing_transitive; [apply H| easy]. + Qed. + + Lemma limit_jump_lowness (A: nat -> Prop) : + LEM_Σ 1 -> + definite K -> + limit_computable (A´) -> ~ K ⪯ᴛ A. + Proof. + intros LEM defK H IH. + apply lowness with (P := A); [|apply IH]. + pose (P := fun (v: vec nat 1) => A´ (hd v)). + Admitted. + +End LowFacts. From 7164566c29b0ebd015b5e724ecefd5db98c771ee Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 3 Oct 2023 16:58:28 +0200 Subject: [PATCH 03/54] delete dup comments --- theories/Basic/Limit.v | 27 +-------------------------- theories/PostsTheorem/TuringJump.v | 4 ++-- 2 files changed, 3 insertions(+), 28 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index 43761c8..23cab9f 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -24,7 +24,7 @@ Convention: **) -(* definition of limit ciomputable *) +(* Definition of limit ciomputable *) Definition limit_computable {X} (P: X -> Prop) := exists f: X -> nat -> bool, forall x, @@ -105,31 +105,6 @@ Section LimitLemma1. apply Dec.nat_eq_dec. Qed. -(* - (** Move to lowsimple.v **) - - Definition low (P: nat -> Prop) := P´ ⪯ᴛ ­{0}^(1). - - Lemma lowness (P: nat -> Prop) : - low P -> ~ ­{0}^(1) ⪯ᴛ P. - Proof. - intros H IH. - eapply not_turing_red_J with (Q := P). - eapply Turing_transitive; [apply H| easy]. - Qed. - - Lemma limit_jump_lowness (A: nat -> Prop) : - LEM_Σ 1 -> - definite (­{0}^(1)) -> - LimitComputable (A´) -> ~ ­{0}^(1) ⪯ᴛ A. - Proof. - intros LEM defK H IH. - apply lowness with (P := A); [|apply IH]. - pose (P := fun (v: vec nat 1) => A´ (hd v)). - eapply Turing_transitive; [|apply (@limit_turing_red _ P LEM defK)]. - Admitted. - *) - End LimitLemma1. diff --git a/theories/PostsTheorem/TuringJump.v b/theories/PostsTheorem/TuringJump.v index 5a37607..1164a96 100644 --- a/theories/PostsTheorem/TuringJump.v +++ b/theories/PostsTheorem/TuringJump.v @@ -330,7 +330,7 @@ Section jump. reflexivity. Qed. - Lemma red_m_iff_semidec_jump (P : nat -> Prop) (Q : nat -> Prop): + Lemma red_m_iff_semidec_jump (P : nat -> Prop) (Q : nat -> Prop): oracle_semi_decidable Q P <-> P ⪯ₘ (J Q). Proof. split. @@ -404,4 +404,4 @@ Fixpoint jump_n {Part : partiality} Q n := | 0 => Q | S n => J (jump_n Q n) end. -Notation "A '^(' n ')'" := (jump_n A n) (at level 25, format "A ^( n )"). +Notation "A '^(' n ')'" := (jump_n A n) (at level 25, format "A ^( n )"). \ No newline at end of file From 5b36e60b017061bdf53fb4d820d2d5593d145f64 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sun, 29 Oct 2023 20:36:13 +0100 Subject: [PATCH 04/54] Ongoing Limit Lemma2 --- theories/Basic/Limit.v | 210 ++++++++++++++++++++++++++++++++++++----- 1 file changed, 185 insertions(+), 25 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index 23cab9f..61b977f 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -1,4 +1,4 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic PostsTheorem reductions SemiDec TuringJump OracleComputability Definitions. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic PostsTheorem reductions SemiDec TuringJump OracleComputability Definitions partial. Require Import Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. @@ -21,17 +21,75 @@ Convention: turing_red_K: turing reduciable to halting problem char[_X]: extensionality of X -**) + **) + Definition limit_computable' {X} (P: X -> bool -> Prop) := + exists f: X -> nat ↛ bool, forall x y, P x y <-> exists N, forall n, n > N -> f x n =! y. -(* Definition of limit ciomputable *) + + (* Definition of limit ciomputable *) Definition limit_computable {X} (P: X -> Prop) := exists f: X -> nat -> bool, forall x, (P x <-> exists N, forall n, n > N -> f x n = true) /\ (~ P x <-> exists N, forall n, n > N -> f x n = false). -(* Naming the halting problem as K *) + Definition limit_computable'' {X} (P: X -> bool -> Prop) := + exists f: X -> nat -> bool, forall x y, P x y <-> exists N, forall n, n > N -> f x n = y. + + Lemma limit''_limit {X} (P: X -> Prop): limit_computable'' (char_rel P) -> limit_computable P. + Proof. + intros [f Hf]. + exists f; intros x; split. + - now rewrite (Hf x true). + - now rewrite (Hf x false). + Qed. + + Definition Post := forall X (p : X -> Prop), semi_decidable p -> semi_decidable (compl p) -> decidable p. + + Lemma try {X} (f: X ↛ bool) (P: X -> Prop): Post -> + (forall x y, (char_rel P) x y <-> f x =! y) -> + (exists f, forall x y, (char_rel P) x y <-> f x = y). + Proof. + intros PT H2. + assert (decidable P). + { apply PT. + unfold semi_decidable, semi_decider. + exists (fun x n => if seval (f x ) n is Some t then t else false). + intro x. rewrite (H2 x true); split. + - intros H'%seval_hasvalue. destruct H' as [n1 Hn1]. + exists n1. now rewrite Hn1. + - intros [n1 Hn1]; rewrite seval_hasvalue. + exists n1. destruct (seval (f x) n1) as [|[]]; try congruence. + - exists (fun x n => if seval (f x) n is Some t then negb t else false). + intro x. rewrite (H2 x false); split. + intros H'%seval_hasvalue. destruct H' as [n1 Hn1]. + exists n1. now rewrite Hn1. + intros [n1 Hn1]; rewrite seval_hasvalue. + exists n1. destruct (seval (f x) n1) as [|[]]; try congruence. + destruct b; easy. } + destruct H as [g Hg]. + exists g; intros x []. cbn; apply Hg. + cbn. split; intro h1; destruct (g x) eqn: E; try easy. + now apply Hg in E. intro E'. apply Hg in E'. congruence. + Qed. + + Lemma limit_impl {X} (P: X -> Prop) : limit_computable' (char_rel P) -> limit_computable P. + Proof. + intros [f Hf]. + unfold limit_computable''. + Admitted. + + + + + + + + + + + (* Naming the halting problem as K *) Notation K := (­{0}^(1)). @@ -150,37 +208,38 @@ Section Σ1Approximation. Qed. (* The Σ1 approximation output correct answers for arbitray list of questions *) + Definition approximation_list {A} (P: A -> Prop) (f: A -> bool) L := + forall i, List.In i L -> P i <-> f i = true. + + Definition approximation_Σ1 {A} (P: A -> Prop) := + exists P_ : nat -> A -> bool, + forall L, exists c, forall c', c' >= c -> approximation_list P (P_ c') L. - Lemma semi_dec_char_Σ1 {X} (P: X -> Prop) : - semi_decidable P -> - exists K_: X -> nat -> bool, - forall l: list X, + Lemma semi_dec_approximation_Σ1 {X} (P: X -> Prop) : definite P -> - exists c: nat, forall e, List.In e l -> K_ e c = true <-> P e. + semi_decidable P -> approximation_Σ1 P. Proof. - intros semi_P. - destruct (stabilize semi_P) as [h [Hh HS]]. - exists h; intros l defP. induction l as [|a l [c Hc]]. + intros defP semiP; unfold approximation_Σ1, approximation_list. + destruct (stabilize semiP) as [h [Hh HS]]. + exists (fun n x => h x n). intro l. induction l as [|a l [c Hc]]. - exists 42; eauto. - destruct (defP a) as [h1| h2]. + destruct (Hh a) as [H _]. destruct (H h1) as [N HN]. - exists (max c N); intros e [->| He]. - split; [easy|intros _]. + exists (max c N); intros c' Hc' e [->| He]. + split; [intros _|easy]. eapply HS; [|eapply HN]; lia. - rewrite <- Hc; [|assumption]. - split; intro h'. - rewrite Hc; [|exact He]. + rewrite <- (Hc c'); [trivial|lia | assumption]. + + exists c; intros c' Hc' e [->| He]. + split; [easy| intros h']. unfold semi_decider in Hh. - rewrite Hh; now exists (Nat.max c N). - eapply HS; [|apply h']; lia. - + exists c; intros e [->| He]. - split; [intro h'|easy]. - unfold semi_decider in Hh. - now rewrite Hh; exists c. + now rewrite Hh; exists c'. rewrite Hc; eauto. Qed. + Lemma approximation_Σ1_halting : definite K -> approximation_Σ1 K. + Proof. now intros H; apply semi_dec_approximation_Σ1, semi_dec_halting. Qed. + End Σ1Approximation. @@ -188,9 +247,110 @@ Section LimitLemma2. (* A predicate P is reduciable to K if P is limit computable *) + Section Construction. + + Variable f : nat -> nat -> bool. + Variable tau : nat -> tree nat bool bool. + Hypothesis Hf: forall L, exists c, forall c', c' >= c -> approximation_list K (f c') L. + + Definition K_ n := fun i o => f n i = o. + Definition char_K_ n := fun i => ret (f n i). + + Lemma dec_K_ n : decidable (fun i => f n i = true). + Proof. + exists (f n). easy. + Qed. + + Lemma pcomputes_K_ n: pcomputes (char_K_ n) (fun i o => f n i = o). + Proof. + intros i o; split; intro H. + now apply ret_hasvalue_inv. + now apply ret_hasvalue'. + Qed. + + Variable F: nat ↛ bool -> nat ↛ bool. + + Definition Phi x n := if (seval (F (char_K_ n) x) n) is Some t then t else true. + + Check cont_to_cont. + + End Construction. + Theorem turing_red_K_lim (P: nat -> Prop) : - P ⪯ᴛ K -> limit_computable P. + P ⪯ᴛ K -> + definite K -> + definite P -> + limit_computable'' (char_rel P). Proof. - Admitted. + intros [F [H HF]] defK defP. + remember H as H'; remember defK as defK'. + destruct (approximation_Σ1_halting defK') as [k_ Hk_]. + destruct H' as [tau Htau]. + specialize (Turing_transports_computable F H) as [F' HPhi]. + pose (char_K_ n := char_K_ k_ n). + pose (K_ n := K_ k_ n). + pose (Phi := Phi k_ F'). + specialize (fun n => HPhi (char_K_ n) (K_ n) (pcomputes_K_ _ _)). + unfold pcomputes in HPhi. +(* + assert ({g : nat -> nat -> bool | + forall n x y, (forall s, s >= n -> g x s = y) <-> }) as Lemma1. + { + exists (fun x s => if (seval (Phi (char_K_ s) x) s) is Some t then t else true). + intros n x y; split. + - intros. apply HPhi. + rewrite seval_hasvalue. admit. + - intros. + } + + + apply limit_impl. + exists (fun x n => Phi (char_K_ n) x). +*) + assert (forall x y, char_rel P x y -> exists N : nat, forall n : nat, n > N -> Phi x n = y) as HL. + { intros x y H1. rewrite HF in H1. + destruct (cont_to_cont _ H _ _ _ H1) as [lm Hlm]. + destruct (Hk_ lm) as [nth_K Hnth_K]. + assert (forall n, n > nth_K -> F (K_ n) x y); [intros n1 Hn1|]. + apply Hlm; intros i' o' h1 h2. + assert (approximation_list K (k_ n1) lm). + apply Hnth_K; lia. unfold approximation_list in H0. + specialize (H0 i' h1). + cbn in h2. destruct o'. + now apply H0. destruct (k_ n1 i') eqn: E; [|easy]. + now exfalso; apply h2; rewrite H0. + (** TODO **) + (* + exists nth_K; intros n1 Hn1. + unfold pcomputes in HPhi. + rewrite (HPhi n1 x y). + now apply H0. } *) + admit. + } + exists Phi. intros x y; split; intro H1. + - now apply HL. + - destruct y; cbn. + destruct (defP x); [easy|]. + assert (char_rel P x false); [easy|]. + apply HL in H2. + destruct H1 as [N1 HN1]. + destruct H2 as [N2 HN2]. + specialize (HN1 (S (max N1 N2))). + specialize (HN2 (S (max N1 N2))). + enough (true = false) by congruence. + rewrite <- HN1, HN2; lia. + destruct (defP x); [|easy]. + assert (char_rel P x true); [easy|]. + apply HL in H2. + destruct H1 as [N1 HN1]. + destruct H2 as [N2 HN2]. + specialize (HN1 (S (max N1 N2))). + specialize (HN2 (S (max N1 N2))). + enough (true = false) by congruence. + rewrite <- HN2, HN1; lia. + Qed. + + + End LimitLemma2. From 83e54b17df176d3e57ad0eb1c3b3c92f70c270ae Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 31 Oct 2023 02:52:59 +0100 Subject: [PATCH 05/54] Finshed proof outline --- theories/Basic/Limit.v | 148 +++++++-------- .../TuringReducibility/OracleComputability.v | 177 +++++++++++++++++- 2 files changed, 246 insertions(+), 79 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index 61b977f..f92a69d 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -31,11 +31,11 @@ Convention: Definition limit_computable {X} (P: X -> Prop) := exists f: X -> nat -> bool, forall x, - (P x <-> exists N, forall n, n > N -> f x n = true) /\ - (~ P x <-> exists N, forall n, n > N -> f x n = false). + (P x <-> exists N, forall n, n >= N -> f x n = true) /\ + (~ P x <-> exists N, forall n, n >= N -> f x n = false). Definition limit_computable'' {X} (P: X -> bool -> Prop) := - exists f: X -> nat -> bool, forall x y, P x y <-> exists N, forall n, n > N -> f x n = y. + exists f: X -> nat -> bool, forall x y, P x y <-> exists N, forall n, n >= N -> f x n = y. Lemma limit''_limit {X} (P: X -> Prop): limit_computable'' (char_rel P) -> limit_computable P. Proof. @@ -74,18 +74,6 @@ Convention: now apply Hg in E. intro E'. apply Hg in E'. congruence. Qed. - Lemma limit_impl {X} (P: X -> Prop) : limit_computable' (char_rel P) -> limit_computable P. - Proof. - intros [f Hf]. - unfold limit_computable''. - Admitted. - - - - - - - @@ -127,15 +115,15 @@ Section LimitLemma1. limit_computable P -> isΣsem 2 P /\ isΣsem 2 (compl P). Proof. intros [f H]; split; eapply char_Σ2. - - exists (fun N n x => if le_dec n N then true else f x n). + - exists (fun N n x => if lt_dec n N then true else f x n). intro w. destruct (H w) as [-> _]; split; intros [N Hn]; exists N. - + intro m. destruct (le_dec m N); try apply Hn; lia. - + intros n He. specialize (Hn n); destruct (le_dec n N); auto; lia. - - exists (fun N n x => if le_dec n N then true else negb (f x n)). + + intro m. destruct (lt_dec m N); try apply Hn; lia. + + intros n He. specialize (Hn n); destruct (lt_dec n N); auto; lia. + - exists (fun N n x => if lt_dec n N then true else negb (f x n)). intro w. destruct (H w) as [_ ->]; split; intros [N Hn]; exists N. - + intro m. destruct (le_dec m N); [auto| rewrite (Hn m); lia]. + + intro m. destruct (lt_dec m N); [auto| rewrite (Hn m); lia]. + intros n He. specialize (Hn n). - destruct (le_dec n N); auto; [lia|destruct (f w n); easy]. + destruct (lt_dec n N); auto; [lia|destruct (f w n); easy]. Qed. Lemma limit_semi_dec_K {k: nat} (P: vec nat k -> Prop) : @@ -215,6 +203,11 @@ Section Σ1Approximation. exists P_ : nat -> A -> bool, forall L, exists c, forall c', c' >= c -> approximation_list P (P_ c') L. + Definition approximation_Σ1_strong {A} (P: A -> Prop) := + exists P_ : nat -> A -> bool, + (forall L, exists c, forall c', c' >= c -> approximation_list P (P_ c') L) /\ + (forall tau q a, @interrogation _ _ _ bool tau (char_rel P) q a -> exists n, forall m, m >= n -> interrogation tau (fun q a => P_ m q = a) q a). + Lemma semi_dec_approximation_Σ1 {X} (P: X -> Prop) : definite P -> semi_decidable P -> approximation_Σ1 P. @@ -237,9 +230,32 @@ Section Σ1Approximation. rewrite Hc; eauto. Qed. + Lemma semi_dec_approximation_Σ1_strong {X} (P: X -> Prop) : + definite P -> + semi_decidable P -> approximation_Σ1_strong P. + Proof. + intros defP semiP. + destruct (semi_dec_approximation_Σ1 defP semiP) as [P_ HP_]. + exists P_; split; [apply HP_|]. + intros tau q ans Htau. + destruct (HP_ q) as [w Hw]. + exists w. intros m Hm. rewrite interrogation_ext. + exact Htau. eauto. + intros q_ a H1. + specialize (Hw m Hm q_ H1). + unfold char_rel; cbn. + destruct a; eauto; split; intro h2. + intro h. rewrite Hw in h. congruence. + firstorder. + Qed. + Lemma approximation_Σ1_halting : definite K -> approximation_Σ1 K. Proof. now intros H; apply semi_dec_approximation_Σ1, semi_dec_halting. Qed. + Lemma approximation_Σ1_halting_strong: definite K -> approximation_Σ1_strong K. + Proof. now intros H; apply semi_dec_approximation_Σ1_strong, semi_dec_halting. Qed. + + End Σ1Approximation. @@ -268,12 +284,6 @@ Section LimitLemma2. now apply ret_hasvalue'. Qed. - Variable F: nat ↛ bool -> nat ↛ bool. - - Definition Phi x n := if (seval (F (char_K_ n) x) n) is Some t then t else true. - - Check cont_to_cont. - End Construction. Theorem turing_red_K_lim (P: nat -> Prop) : @@ -283,67 +293,53 @@ Section LimitLemma2. limit_computable'' (char_rel P). Proof. intros [F [H HF]] defK defP. - remember H as H'; remember defK as defK'. - destruct (approximation_Σ1_halting defK') as [k_ Hk_]. - destruct H' as [tau Htau]. - specialize (Turing_transports_computable F H) as [F' HPhi]. + destruct (approximation_Σ1_halting_strong defK) as [k_ [Hk_1 Hk_2]]. + destruct H as [tau Htau]. pose (char_K_ n := char_K_ k_ n). pose (K_ n := K_ k_ n). - pose (Phi := Phi k_ F'). - specialize (fun n => HPhi (char_K_ n) (K_ n) (pcomputes_K_ _ _)). - unfold pcomputes in HPhi. -(* - assert ({g : nat -> nat -> bool | - forall n x y, (forall s, s >= n -> g x s = y) <-> }) as Lemma1. + pose (Phi x n := evalt_comp (tau x) (k_ n) n n). + assert (forall x y, char_rel P x y -> exists N : nat, forall n : nat, n >= N -> (evalt_comp (tau x) (k_ n)) n n = Some (inr y)) as HL. { - exists (fun x s => if (seval (Phi (char_K_ s) x) s) is Some t then t else true). - intros n x y; split. - - intros. apply HPhi. - rewrite seval_hasvalue. admit. - - intros. + intros x y H. + rewrite HF in H. + rewrite Htau in H. + destruct H as (qs & ans & Hint & Out). + specialize (Hk_2 (tau x) qs ans Hint). + destruct Hk_2 as [nth Hnth]. + assert (interrogation (tau x) + (fun (q : nat) (a : bool) => (k_ nth) q = a) qs ans) as Hnthbase. + eapply Hnth. lia. + edestruct (interrogation_evalt_comp_limit (tau x) k_ qs ans y) as [L Hlimt]. + exists nth. intros. eapply Hnth. easy. + eapply Out. + exists L. intros. now apply Hlimt. } - - - apply limit_impl. - exists (fun x n => Phi (char_K_ n) x). -*) - assert (forall x y, char_rel P x y -> exists N : nat, forall n : nat, n > N -> Phi x n = y) as HL. - { intros x y H1. rewrite HF in H1. - destruct (cont_to_cont _ H _ _ _ H1) as [lm Hlm]. - destruct (Hk_ lm) as [nth_K Hnth_K]. - assert (forall n, n > nth_K -> F (K_ n) x y); [intros n1 Hn1|]. - apply Hlm; intros i' o' h1 h2. - assert (approximation_list K (k_ n1) lm). - apply Hnth_K; lia. unfold approximation_list in H0. - specialize (H0 i' h1). - cbn in h2. destruct o'. - now apply H0. destruct (k_ n1 i') eqn: E; [|easy]. - now exfalso; apply h2; rewrite H0. - (** TODO **) - (* - exists nth_K; intros n1 Hn1. - unfold pcomputes in HPhi. - rewrite (HPhi n1 x y). - now apply H0. } *) - admit. + assert (exists f, forall x y, char_rel P x y -> exists N : nat, forall n : nat, n >= N -> f x n = y) as [f HL']. + { + exists (fun x n => match (Phi x n) with + | Some (inr y) => y | _ => false end). + intros x y Hxy%HL. + destruct (Hxy) as [N HN]. + exists N; intros. + unfold Phi. rewrite HN; eauto. } - exists Phi. intros x y; split; intro H1. - - now apply HL. - - destruct y; cbn. + exists f. intros x y; split. + - now intros; apply HL'. + - intro H0. destruct y; cbn. destruct (defP x); [easy|]. assert (char_rel P x false); [easy|]. - apply HL in H2. - destruct H1 as [N1 HN1]. - destruct H2 as [N2 HN2]. + apply HL' in H1. + destruct H0 as [N1 HN1]. + destruct H1 as [N2 HN2]. specialize (HN1 (S (max N1 N2))). specialize (HN2 (S (max N1 N2))). enough (true = false) by congruence. rewrite <- HN1, HN2; lia. destruct (defP x); [|easy]. assert (char_rel P x true); [easy|]. - apply HL in H2. - destruct H1 as [N1 HN1]. - destruct H2 as [N2 HN2]. + apply HL' in H1. + destruct H0 as [N1 HN1]. + destruct H1 as [N2 HN2]. specialize (HN1 (S (max N1 N2))). specialize (HN2 (S (max N1 N2))). enough (true = false) by congruence. @@ -351,6 +347,4 @@ Section LimitLemma2. Qed. - - End LimitLemma2. diff --git a/theories/TuringReducibility/OracleComputability.v b/theories/TuringReducibility/OracleComputability.v index 7aa08ed..48a3827 100644 --- a/theories/TuringReducibility/OracleComputability.v +++ b/theories/TuringReducibility/OracleComputability.v @@ -812,7 +812,9 @@ Qed. Lemma interrogation_equiv_evalt Q A I O : forall (τ : I -> list A ↛ (Q + O)) (f : Q ↛ A) (i : I) (o : O), - (exists (qs : list Q) (ans : list A), interrogation (τ i) (fun x y => f x =! y) qs ans /\ τ i ans =! inr o) <-> (exists n : nat, evalt (τ i) f n =! inr o). + (exists (qs : list Q) (ans : list A), + interrogation (τ i) (fun x y => f x =! y) qs ans /\ τ i ans =! inr o) <-> + (exists n : nat, evalt (τ i) f n =! inr o). Proof. intros τ f i o. split. + intros (qs & ans & Hi & Hout). @@ -821,6 +823,176 @@ Proof. + intros [n H]. eapply evalt_to_interrogation in H as (? & ? & ? & ? & ?); eauto. Qed. + +Fixpoint evalt_comp {Q A O} (tau : list A ↛ (Q + O)) + (f : Q -> A) (n : nat) (step: nat): option (Q + O) := + match (seval (tau []) step) with + | Some x => match n, x with + | 0, inl q => Some (inl q) + | S n, inl q => evalt_comp (fun l => tau (f q :: l)) f n step + | _, inr o => Some (inr o) end + | None => None end. + + +Lemma evalt_comp_ext {Q A O} τ τ' f n m: + (forall l n, seval (τ l) n = seval (τ' l) n) -> + evalt_comp τ f n m = @evalt_comp Q A O τ' f n m. +Proof. + intro Heq; induction n in τ, τ', Heq |- *. + - cbn. now rewrite <- Heq. + - cbn. rewrite <- Heq. + destruct (seval (A:=Q + O) (τ []) m); cbn. + destruct s; eauto. easy. +Qed. + +Lemma list_cons_or {X} (l l1: list X) a : + l `prefix_of` l1 ++ [a] -> + l `prefix_of` l1 \/ l = l1 ++ [a]. +Proof. +Admitted. + +Lemma interrogation_ter {Q A O} tau f l lv v: + @interrogation Q A O tau (fun x y => f x = y) l lv -> + tau lv =! v -> + exists m, forall ans_, prefix ans_ lv -> exists v, seval (tau ans_) m = Some v. +Proof. + intros H1 H2. + induction H1 in H2, v |-*. + - rewrite seval_hasvalue in H2. + destruct H2 as [m Hm]. exists m. + intros ans_ Hans_. exists v. + apply prefix_nil_inv in Hans_. + rewrite Hans_. easy. + - rewrite seval_hasvalue in H2. + destruct H2 as [m Hm]. + destruct (IHinterrogation (Ask q) H) as [m' Hm']. + exists (max m m'). + intros ans_ Hans_. + destruct (list_cons_or _ _ _ Hans_) as [h| ->]. + + destruct (Hm' ans_ h) as [v' Hv']. + exists v'. eapply seval_mono. + eauto. lia. + + exists v. eapply seval_mono; [eauto| lia]. +Qed. + +Lemma evalt_comp_n_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : + evalt_comp tau f n m = Some (Output o) -> + forall n', n' >= n -> evalt_comp tau f n' m = Some (Output o). +Proof. + intros H n' Hn'. + induction n. + +Admitted. + +Lemma evalt_comp_step_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : + evalt_comp tau f n m = Some o -> + forall m', m' >= m -> evalt_comp tau f n m' = Some o. +Proof. + intros H m' Hm'. + induction n in H, tau, o |-*. + - cbn in *. + destruct (seval (A:=Q + O) (tau []) m) eqn: E; try congruence. + assert (seval (A:=Q + O) (tau []) m' = (Some s)) as ->. + eapply seval_mono. exact E. lia. + easy. + - cbn in *. + destruct (seval (A:=Q + O) (tau []) m) eqn: E. + + assert (seval (A:=Q + O) (tau []) m' = Some s) as ->. + eapply seval_mono. exact E. lia. + destruct s. + now apply IHn. + exact H. + + congruence. +Qed. + +Require Import Coq.Program.Equality. + +Lemma interrogation_plus_comp {Q A O} tau f n m l lv v2: + @interrogation Q A O tau (fun x y => f x = y) l lv -> + (forall ans_, prefix ans_ lv -> exists v, seval (tau ans_) m = Some v) -> + evalt_comp (fun l' => tau (lv ++ l')) f n m = Some v2 -> + evalt_comp tau f (length l + n) m = Some v2. +Proof. + intros H H1. revert n. dependent induction H. + - cbn. eauto. + - intros. + cbn -[evalt]. rewrite app_length. cbn -[evalt]. + replace (length qs + 1 + n) with (length qs + (S n)) by lia. + eapply IHinterrogation. intros; apply H2. + etransitivity. exact H4. + now eapply prefix_app_r. + cbn. rewrite app_nil_r. + destruct (H2 ans). + now eapply prefix_app_r. + assert (exists m, seval (tau ans) m = Some x). + now exists m. + rewrite <- seval_hasvalue in H5. + assert (x = Ask q). + eapply hasvalue_det; eauto. + rewrite H4, H6, H1. + rewrite <- H3. eapply evalt_comp_ext. + intros; now list_simplifier. +Qed. + +Lemma interrogation_evalt_comp {Q A O} tau f l lv v: + @interrogation Q A O tau (fun x y => f x = y) l lv -> + tau lv =! v -> + exists n step, evalt_comp tau f n step = Some v. +Proof. + intros H1 h2. + destruct (interrogation_ter _ _ _ _ _ H1 h2) as [step Hstep]. + exists (length l + 0). + exists step. eapply interrogation_plus_comp; eauto. + cbn. rewrite app_nil_r. + destruct (Hstep lv) as [v' Hv']. + reflexivity. + assert (exists k, seval (A:=Q + O) (tau lv) k = Some v'). + exists step. easy. + rewrite <- seval_hasvalue in H. + assert (v' = v). + eapply hasvalue_det; eauto. + rewrite Hv', H0. + destruct v; done. +Qed. + +Lemma evalt_comp_index_mono {Q A O} tau f l lv v: + @interrogation Q A O tau (fun x y => f x = y) l lv -> + tau lv =! v -> + exists a b, + forall g, @interrogation Q A O tau (fun x y => g x = y) l lv -> + evalt_comp tau g a b = Some v. +Proof. + intros H h. + destruct (interrogation_evalt_comp _ _ _ _ _ H h) as (a&b&Hab). + exists a, b. intros g Hg. + induction a. + - cbn in *. destruct (seval (A:=Q + O) (tau []) b). + + now destruct s. + + discriminate. + - cbn in *. + destruct (seval (tau []) b); [|done]. + destruct s; [|done]. +Admitted. + +Lemma interrogation_evalt_comp_limit {Q A O} tau f l lv v: + (exists K, forall k, k >= K -> + @interrogation Q A O tau (fun x y => (f k) x = y) l lv) -> + tau lv =! Output v -> + exists N, forall n, n >= N -> evalt_comp tau (f n) n n = Some (Output v). +Proof. + intros [k h1] h2. + assert (interrogation tau (fun x y => f k x = y) l lv) as H. + apply h1. lia. + destruct (evalt_comp_index_mono _ _ _ _ _ H h2) as (a & b & Hab). + exists (max (max a b) k). + intros n Hn. + eapply evalt_comp_step_mono. + eapply evalt_comp_n_mono. + eapply Hab. apply h1. + all: lia. +Qed. + + (** ** Closure properties of Oracle computability *) (** Computability of precomposition *) @@ -862,7 +1034,6 @@ Proof. cbn. firstorder subst; psimpl. Qed. -(** Computability of constant functions *) Lemma computable_ret A Q I O v : @OracleComputable Q A I O (fun f i o => o = v). Proof. @@ -1276,6 +1447,7 @@ Proof. + intros []; cbn; firstorder. Qed. + (** Turing reduction transports partial computability - relying on the evalt function from above *) Lemma Turing_transports_computable_strong {Q A I O} F tau : (∀ (R : Q → A → Prop) (x : I) (o : O), F R x o ↔ (∃ (qs : list Q) (ans : list A), interrogation (tau x) R qs ans ∧ tau x ans =! Output o)) -> @@ -1346,6 +1518,7 @@ Proof. destruct (Turing_transports_computable_strong F tau) as [F' ]; eauto. Qed. + (** Transport of decidability -- which is equivalent to Markov's principle *) Definition char_rel_fun {X Y} (f : X -> Y) := (fun x b => f x = b). From 4e3f02a98de28a99cd701420a2aabd125eb80cc6 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 31 Oct 2023 17:29:30 +0100 Subject: [PATCH 06/54] Proof mono_m --- .../TuringReducibility/OracleComputability.v | 92 ++++++++++++++----- 1 file changed, 71 insertions(+), 21 deletions(-) diff --git a/theories/TuringReducibility/OracleComputability.v b/theories/TuringReducibility/OracleComputability.v index 48a3827..880eaf4 100644 --- a/theories/TuringReducibility/OracleComputability.v +++ b/theories/TuringReducibility/OracleComputability.v @@ -849,7 +849,22 @@ Lemma list_cons_or {X} (l l1: list X) a : l `prefix_of` l1 ++ [a] -> l `prefix_of` l1 \/ l = l1 ++ [a]. Proof. -Admitted. + induction l in l1 |-*; intros. + - left. eapply prefix_nil. + - destruct l1. + + right. list_simplifier. + set (H' := H). + apply prefix_cons_inv_2, prefix_nil_inv in H' as ->. + apply prefix_cons_inv_1 in H as ->. + done. + + list_simplifier. + set (H' := H). + apply prefix_cons_inv_2 in H'. + apply prefix_cons_inv_1 in H as ->. + destruct (IHl _ H') as [h1 | ->]. + ++ left. by eapply prefix_cons. + ++ now right. +Qed. Lemma interrogation_ter {Q A O} tau f l lv v: @interrogation Q A O tau (fun x y => f x = y) l lv -> @@ -875,14 +890,6 @@ Proof. + exists v. eapply seval_mono; [eauto| lia]. Qed. -Lemma evalt_comp_n_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : - evalt_comp tau f n m = Some (Output o) -> - forall n', n' >= n -> evalt_comp tau f n' m = Some (Output o). -Proof. - intros H n' Hn'. - induction n. - -Admitted. Lemma evalt_comp_step_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : evalt_comp tau f n m = Some o -> @@ -934,6 +941,43 @@ Proof. intros; now list_simplifier. Qed. +Lemma interrogation_plus_comp {Q A O} tau f n m l lv v2: + @interrogation Q A O tau (fun x y => f x = y) l lv -> + (forall ans_, prefix ans_ lv -> exists v, seval (tau ans_) m = Some v) -> + evalt_comp (fun l' => tau (lv ++ l')) f n m = Some v2 -> + evalt_comp tau f (length l + n) m = Some v2. +Proof. + intros H H1. revert n. dependent induction H. + - cbn. eauto. + - intros. + cbn -[evalt]. rewrite app_length. cbn -[evalt]. + replace (length qs + 1 + n) with (length qs + (S n)) by lia. + eapply IHinterrogation. intros; apply H2. + etransitivity. exact H4. + now eapply prefix_app_r. + cbn. rewrite app_nil_r. + destruct (H2 ans). + now eapply prefix_app_r. + assert (exists m, seval (tau ans) m = Some x). + now exists m. + rewrite <- seval_hasvalue in H5. + assert (x = Ask q). + eapply hasvalue_det; eauto. + rewrite H4, H6, H1. + rewrite <- H3. eapply evalt_comp_ext. + intros; now list_simplifier. +Qed. + +Lemma evalt_comp_n_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : + evalt_comp tau f n m = Some (Output o) -> + forall n', n' >= n -> evalt_comp tau f n' m = Some (Output o). +Proof. + intros H n' Hn'. + induction n. + - admit. + - cbn in H. +Admitted. + Lemma interrogation_evalt_comp {Q A O} tau f l lv v: @interrogation Q A O tau (fun x y => f x = y) l lv -> tau lv =! v -> @@ -943,7 +987,7 @@ Proof. destruct (interrogation_ter _ _ _ _ _ H1 h2) as [step Hstep]. exists (length l + 0). exists step. eapply interrogation_plus_comp; eauto. - cbn. rewrite app_nil_r. + cbn. rewrite app_nil_r. destruct (Hstep lv) as [v' Hv']. reflexivity. assert (exists k, seval (A:=Q + O) (tau lv) k = Some v'). @@ -962,17 +1006,23 @@ Lemma evalt_comp_index_mono {Q A O} tau f l lv v: forall g, @interrogation Q A O tau (fun x y => g x = y) l lv -> evalt_comp tau g a b = Some v. Proof. - intros H h. - destruct (interrogation_evalt_comp _ _ _ _ _ H h) as (a&b&Hab). - exists a, b. intros g Hg. - induction a. - - cbn in *. destruct (seval (A:=Q + O) (tau []) b). - + now destruct s. - + discriminate. - - cbn in *. - destruct (seval (tau []) b); [|done]. - destruct s; [|done]. -Admitted. + intros H1 h2. + destruct (interrogation_ter _ _ _ _ _ H1 h2) as [step Hstep]. + exists (length l + 0). + exists step. + intros. + eapply interrogation_plus_comp; eauto. + cbn. rewrite app_nil_r. + destruct (Hstep lv) as [v' Hv']. + reflexivity. + assert (exists k, seval (A:=Q + O) (tau lv) k = Some v'). + exists step. easy. + rewrite <- seval_hasvalue in H0. + assert (v' = v). + eapply hasvalue_det; eauto. + rewrite Hv', H2. + destruct v; done. +Qed. Lemma interrogation_evalt_comp_limit {Q A O} tau f l lv v: (exists K, forall k, k >= K -> From e8551f69d8f80fa7449f52a8dc991d83ca1d9a41 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sun, 5 Nov 2023 15:53:24 +0100 Subject: [PATCH 07/54] Proof monotonic of m --- .../TuringReducibility/OracleComputability.v | 33 ++----------------- 1 file changed, 2 insertions(+), 31 deletions(-) diff --git a/theories/TuringReducibility/OracleComputability.v b/theories/TuringReducibility/OracleComputability.v index 880eaf4..dfdc689 100644 --- a/theories/TuringReducibility/OracleComputability.v +++ b/theories/TuringReducibility/OracleComputability.v @@ -941,41 +941,11 @@ Proof. intros; now list_simplifier. Qed. -Lemma interrogation_plus_comp {Q A O} tau f n m l lv v2: - @interrogation Q A O tau (fun x y => f x = y) l lv -> - (forall ans_, prefix ans_ lv -> exists v, seval (tau ans_) m = Some v) -> - evalt_comp (fun l' => tau (lv ++ l')) f n m = Some v2 -> - evalt_comp tau f (length l + n) m = Some v2. -Proof. - intros H H1. revert n. dependent induction H. - - cbn. eauto. - - intros. - cbn -[evalt]. rewrite app_length. cbn -[evalt]. - replace (length qs + 1 + n) with (length qs + (S n)) by lia. - eapply IHinterrogation. intros; apply H2. - etransitivity. exact H4. - now eapply prefix_app_r. - cbn. rewrite app_nil_r. - destruct (H2 ans). - now eapply prefix_app_r. - assert (exists m, seval (tau ans) m = Some x). - now exists m. - rewrite <- seval_hasvalue in H5. - assert (x = Ask q). - eapply hasvalue_det; eauto. - rewrite H4, H6, H1. - rewrite <- H3. eapply evalt_comp_ext. - intros; now list_simplifier. -Qed. - Lemma evalt_comp_n_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : evalt_comp tau f n m = Some (Output o) -> forall n', n' >= n -> evalt_comp tau f n' m = Some (Output o). Proof. - intros H n' Hn'. - induction n. - - admit. - - cbn in H. + Admitted. Lemma interrogation_evalt_comp {Q A O} tau f l lv v: @@ -2002,3 +1972,4 @@ End part. Notation "P ⪯ᴛ Q" := (red_Turing P Q) (at level 50). +Search evalt. From cac6b7e86f73e86b27372d1785a35aad39083375 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 7 Nov 2023 17:23:50 +0100 Subject: [PATCH 08/54] Done Limit Lemma --- theories/Basic/Limit.v | 51 +++------------ .../TuringReducibility/OracleComputability.v | 63 +++++++++++++++---- 2 files changed, 59 insertions(+), 55 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index f92a69d..64fcea0 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -23,9 +23,6 @@ Convention: **) - Definition limit_computable' {X} (P: X -> bool -> Prop) := - exists f: X -> nat ↛ bool, forall x y, P x y <-> exists N, forall n, n > N -> f x n =! y. - (* Definition of limit ciomputable *) @@ -34,49 +31,17 @@ Convention: (P x <-> exists N, forall n, n >= N -> f x n = true) /\ (~ P x <-> exists N, forall n, n >= N -> f x n = false). - Definition limit_computable'' {X} (P: X -> bool -> Prop) := + Definition char_rel_limit_computable {X} (P: X -> bool -> Prop) := exists f: X -> nat -> bool, forall x y, P x y <-> exists N, forall n, n >= N -> f x n = y. - Lemma limit''_limit {X} (P: X -> Prop): limit_computable'' (char_rel P) -> limit_computable P. - Proof. - intros [f Hf]. - exists f; intros x; split. - - now rewrite (Hf x true). - - now rewrite (Hf x false). - Qed. - - Definition Post := forall X (p : X -> Prop), semi_decidable p -> semi_decidable (compl p) -> decidable p. - - Lemma try {X} (f: X ↛ bool) (P: X -> Prop): Post -> - (forall x y, (char_rel P) x y <-> f x =! y) -> - (exists f, forall x y, (char_rel P) x y <-> f x = y). + Lemma char_rel_limit_equv {X} (P: X -> Prop): + char_rel_limit_computable (char_rel P) <-> limit_computable P. Proof. - intros PT H2. - assert (decidable P). - { apply PT. - unfold semi_decidable, semi_decider. - exists (fun x n => if seval (f x ) n is Some t then t else false). - intro x. rewrite (H2 x true); split. - - intros H'%seval_hasvalue. destruct H' as [n1 Hn1]. - exists n1. now rewrite Hn1. - - intros [n1 Hn1]; rewrite seval_hasvalue. - exists n1. destruct (seval (f x) n1) as [|[]]; try congruence. - - exists (fun x n => if seval (f x) n is Some t then negb t else false). - intro x. rewrite (H2 x false); split. - intros H'%seval_hasvalue. destruct H' as [n1 Hn1]. - exists n1. now rewrite Hn1. - intros [n1 Hn1]; rewrite seval_hasvalue. - exists n1. destruct (seval (f x) n1) as [|[]]; try congruence. - destruct b; easy. } - destruct H as [g Hg]. - exists g; intros x []. cbn; apply Hg. - cbn. split; intro h1; destruct (g x) eqn: E; try easy. - now apply Hg in E. intro E'. apply Hg in E'. congruence. + split; intros [f Hf]; exists f; intros x. + - split; firstorder. + - intros []; destruct (Hf x) as [h1 h2]; eauto. Qed. - - - (* Naming the halting problem as K *) Notation K := (­{0}^(1)). @@ -290,9 +255,10 @@ Section LimitLemma2. P ⪯ᴛ K -> definite K -> definite P -> - limit_computable'' (char_rel P). + limit_computable P. Proof. intros [F [H HF]] defK defP. + rewrite <- char_rel_limit_equv. destruct (approximation_Σ1_halting_strong defK) as [k_ [Hk_1 Hk_2]]. destruct H as [tau Htau]. pose (char_K_ n := char_K_ k_ n). @@ -346,5 +312,4 @@ Section LimitLemma2. rewrite <- HN2, HN1; lia. Qed. - End LimitLemma2. diff --git a/theories/TuringReducibility/OracleComputability.v b/theories/TuringReducibility/OracleComputability.v index dfdc689..e59ed86 100644 --- a/theories/TuringReducibility/OracleComputability.v +++ b/theories/TuringReducibility/OracleComputability.v @@ -1,5 +1,6 @@ From SyntheticComputability Require Import partial Dec. From stdpp Require Import list. +Require Import Coq.Program.Equality. Import PartialTactics. Lemma list_find_repeat_not {Y} P D x n : @@ -824,6 +825,8 @@ Proof. Qed. +(** A total computable version of evalt **) + Fixpoint evalt_comp {Q A O} (tau : list A ↛ (Q + O)) (f : Q -> A) (n : nat) (step: nat): option (Q + O) := match (seval (tau []) step) with @@ -890,6 +893,7 @@ Proof. + exists v. eapply seval_mono; [eauto| lia]. Qed. +(** Basic property of evalt_comp **) Lemma evalt_comp_step_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : evalt_comp tau f n m = Some o -> @@ -912,15 +916,13 @@ Proof. + congruence. Qed. -Require Import Coq.Program.Equality. - Lemma interrogation_plus_comp {Q A O} tau f n m l lv v2: @interrogation Q A O tau (fun x y => f x = y) l lv -> (forall ans_, prefix ans_ lv -> exists v, seval (tau ans_) m = Some v) -> - evalt_comp (fun l' => tau (lv ++ l')) f n m = Some v2 -> + evalt_comp (fun l' => tau (lv ++ l')) f n m = Some v2 <-> evalt_comp tau f (length l + n) m = Some v2. Proof. - intros H H1. revert n. dependent induction H. + intros H H1. split; revert n; dependent induction H. - cbn. eauto. - intros. cbn -[evalt]. rewrite app_length. cbn -[evalt]. @@ -939,14 +941,50 @@ Proof. rewrite H4, H6, H1. rewrite <- H3. eapply evalt_comp_ext. intros; now list_simplifier. + - cbn. eauto. + - intros. + rewrite app_length in H3. cbn in H3. + replace (length qs + 1 + n) with (length qs + (S n)) in H3 by lia. + eapply IHinterrogation in H3. + 2: { + intros; apply H2. + etransitivity. exact H4. + now eapply prefix_app_r. + } + cbn in H3. + rewrite app_nil_r in H3. + destruct (H2 ans). + now eapply prefix_app_r. + assert (exists m, seval (tau ans) m = Some x). + now exists m. + rewrite <- seval_hasvalue in H5. + assert (x = Ask q). + eapply hasvalue_det; eauto. + rewrite H4, H6, H1 in H3. + rewrite <- H3. eapply evalt_comp_ext. + intros; now list_simplifier. Qed. -Lemma evalt_comp_n_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : - evalt_comp tau f n m = Some (Output o) -> - forall n', n' >= n -> evalt_comp tau f n' m = Some (Output o). +Lemma evalt_comp_step_length_mono {Q A O} (tau: (list A) ↛ (Q + O)) f qs ans o: + @interrogation Q A O tau (fun x y => f x = y) qs ans -> + tau ans =! Output o -> + exists step n, + forall g, @interrogation Q A O tau (fun x y => g x = y) qs ans -> + forall n', n <= n' -> evalt_comp tau g n' step = Some (inr o). Proof. - -Admitted. + intros H1 H2. + destruct (interrogation_ter _ _ _ _ _ H1 H2) as [step Hstep]. + exists step. exists (length qs). intros g Hg n' Hn'. + assert (exists v, seval (tau ans) step = Some v) as [v Hv]. + { eapply Hstep; naive_solver. } + assert (v = Output o). + { eapply hasvalue_det; [|eapply H2]. rewrite seval_hasvalue. eauto. } + eapply Nat.le_exists_sub in Hn' as [k [-> _]]. + rewrite Nat.add_comm. + eapply interrogation_plus_comp; eauto. + induction k. + all: cbn; rewrite app_nil_r; by rewrite Hv, H. +Qed. Lemma interrogation_evalt_comp {Q A O} tau f l lv v: @interrogation Q A O tau (fun x y => f x = y) l lv -> @@ -1003,12 +1041,13 @@ Proof. intros [k h1] h2. assert (interrogation tau (fun x y => f k x = y) l lv) as H. apply h1. lia. + destruct (evalt_comp_step_length_mono _ _ _ _ _ H h2) as (a' & b' & Hs). destruct (evalt_comp_index_mono _ _ _ _ _ H h2) as (a & b & Hab). - exists (max (max a b) k). + exists (max b'(max a' (max (max a b) k))). intros n Hn. eapply evalt_comp_step_mono. - eapply evalt_comp_n_mono. - eapply Hab. apply h1. + eapply (Hs (f n)); eauto. + eapply h1. all: lia. Qed. From 4a1eb99f99b66ae857470dbdd09b72ae71a6b82a Mon Sep 17 00:00:00 2001 From: Haoyi Date: Thu, 1 Feb 2024 14:03:40 +0100 Subject: [PATCH 09/54] Try SS --- theories/ReducibilityDegrees/lowsimple.v | 59 +++++++++++++++++++++++- 1 file changed, 57 insertions(+), 2 deletions(-) diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index cd0523a..6df7620 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -1,4 +1,4 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. Require Import Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. @@ -36,7 +36,62 @@ Section LowFacts. Proof. intros LEM defK H IH. apply lowness with (P := A); [|apply IH]. - pose (P := fun (v: vec nat 1) => A´ (hd v)). Admitted. End LowFacts. + + +Section LowSimplePredicate. + +Definition low_simple P := low P /\ simple P. + +Definition sol_Post's_problem (P: nat -> Prop) := + (~ decidable P) /\ (enumerable P) /\ ~ (K ⪯ᴛ P). + +Fact low_simple_correct: + forall P, low_simple P -> sol_Post's_problem P. +Proof. + intros P [H1 H2]; split; [now apply simple_undecidable|]. + split; [destruct H2 as [H2 _]; eauto| now apply lowness]. +Qed. + +End LowSimplePredicate. + + +Section Construction. + + Definition mu (P: nat -> Prop) n := P n /\ forall m, P m -> n <= m. + + Variable W: nat -> nat -> nat -> Prop. + + Definition W_spec (c n x: nat) : Prop := forall P, semi_decidable P -> exists c, forall x, P x <-> exists n, W c n x. + + Definition disj (A B: nat -> Prop) := forall x, A x <-> ~ B x. + Notation "A ⊎ B" := (disj A B) (at level 30). + + (** ** Is this a correct definition for inifity ? **) + + Definition infity (P: nat -> Prop) := forall x, exists y, y > x -> P y. + + Fixpoint P n : nat -> Prop := + match n with + | O => fun _ => False + | S n => fun x => + P n x \/ exists e, mu (fun e => e < n /\ ((W e n) ⊎ (P n)) /\ mu (fun x => W e n x /\ 2 * e < x) x) e + end. + + Definition W' e x := exists n, W e n x. + Definition P' x := exists n, P n x. + + Definition R P e := infity (W' e) -> exists w, (W' e w) /\ P w. + + Lemma P_meet_R : forall n, R P' n. + Proof. + Admitted. + + Lemma P_semi_decidable : semi_decidable P'. + Proof. + Admitted. + + +End Construction. From 08bde05e658902e603f4ccfdf576e4f027a95b7a Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 6 Feb 2024 01:30:52 +0100 Subject: [PATCH 10/54] Add F_ --- theories/ReducibilityDegrees/lowsimple.v | 164 ++++++++++++++++++++--- 1 file changed, 148 insertions(+), 16 deletions(-) diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index 6df7620..c93da3f 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -58,40 +58,172 @@ Qed. End LowSimplePredicate. +Require Import List. +Import ListNotations. +Notation "'Σ' x .. y , p" := + (sigT (fun x => .. (sigT (fun y => p)) ..)) + (at level 200, x binder, right associativity, + format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") + : type_scope. + +Require Import SyntheticComputability.Shared.FinitenessFacts. + Section Construction. - Definition mu (P: nat -> Prop) n := P n /\ forall m, P m -> n <= m. + Definition mu n (P: nat -> Prop) := P n /\ forall m, P m -> n <= m. - Variable W: nat -> nat -> nat -> Prop. + Variable W_: nat -> nat -> nat -> Prop. + Hypothesis W_spec: forall P, semi_decidable P -> exists c, forall x, P x <-> exists n, W_ c n x. - Definition W_spec (c n x: nat) : Prop := forall P, semi_decidable P -> exists c, forall x, P x <-> exists n, W c n x. + Definition disj (A B: nat -> Prop) := forall x, A x -> B x -> False. + Definition disj__l (A: list nat) (B: nat -> Prop) := forall x, In x A -> B x -> False. - Definition disj (A B: nat -> Prop) := forall x, A x <-> ~ B x. - Notation "A ⊎ B" := (disj A B) (at level 30). + Notation "A # B" := (disj A B) (at level 30). + Notation "A l# B" := (disj__l A B) (at level 30). - (** ** Is this a correct definition for inifity ? **) + Record Extendsion := + { + extendP: list nat -> nat -> option nat -> Prop; + extend_dec: forall l x, Σ y, extendP l x y; + extend_uni: forall l x y1 y2, extendP l x y1 -> extendP l x y2 -> y1 = y2 + }. - Definition infity (P: nat -> Prop) := forall x, exists y, y > x -> P y. + Fixpoint P_templete (E: (nat -> Prop) -> nat -> nat -> Prop) n x: Prop := + match n with + | O => False + | S n => P_templete E n x \/ E (P_templete E n) n x + end. - Fixpoint P n : nat -> Prop := + Fixpoint F_ (E: Extendsion) n l: Prop := match n with - | O => fun _ => False - | S n => fun x => - P n x \/ exists e, mu (fun e => e < n /\ ((W e n) ⊎ (P n)) /\ mu (fun x => W e n x /\ 2 * e < x) x) e + | O => l = [] + | S n => exists ls a, F_ E n ls /\ (extendP E ls n a) /\ l = if a is Some a then a :: ls else ls end. - Definition W' e x := exists n, W e n x. - Definition P' x := exists n, P n x. + Definition F_with E x := exists l n, In x l /\ (F_ E n l). + + Lemma F_uni E : forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . + Proof. + induction n; simpl; [intros ?? -> ->; easy |]. + intros l1 l2 (ls & a & H1 & H2 & ->) (ls' & a' & H1' & H2' & ->). + rewrite !(IHn _ _ H1 H1') in *. + rewrite !(extend_uni H2 H2') in *. + easy. + Qed. + + Lemma F_computable E : exists f: nat -> list nat, forall n, F_ E n (f n). + Proof. + set (F := fix f x := + match x with + | O => [] + | S x => match projT1 ((extend_dec E) (f x) x) with + | None => f x + | Some a => a :: (f x) + end + end). + exists F. induction n; simpl. + - easy. + - exists (F n). destruct (extend_dec E (F n) n). + exists x. cbn. split; [eauto|]. split; [eauto|]. + easy. + Qed. + + Definition decider: nat -> list nat -> bool. + Proof. + intros x l. + destruct (ListDec.In_dec EqDec.nat_eq_dec x l). + - exact true. - exact false. + Defined. + + Fact decider_spec: forall x l, (decider x l) = true <-> In x l. + Proof. + intros x l. unfold decider. + destruct (ListDec.In_dec EqDec.nat_eq_dec x l); firstorder. + Qed. + + Lemma F_with_semi_decidable E: semi_decidable (F_with E). + Proof. + unfold semi_decidable, semi_decider. + destruct (F_computable E) as [f Hf]. + exists (fun x n => decider x (f n)). + intros x. split. + - intros (l & n & Hxl & Hl). + exists n. rewrite decider_spec. + now rewrite (F_uni (Hf n) Hl). + - intros (n & Hn%decider_spec). + exists (f n), n; split; eauto. + Qed. + + Section SimpleSet. + + Definition extend_simple l__n n x := + exists e, mu e (fun alpha => e < n /\ (l__n l# W_ alpha n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). + + Definition simple_extendsion: Extendsion. + Proof. + exists (fun l n a => forall x, a = Some x <-> extend_simple l n x). + - intros l x. admit. + - intros l x y1 y2 Hy1 Hy2. + admit. + Admitted. + + Definition E_simple P__n n x := + exists e, mu e (fun alpha => e < n /\ (W_ alpha n # P__n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). + + +(* Definition P_ := P_templete E_simple. *) + Definition P_ := F_ simple_extendsion. + Definition W e x := exists n, W_ e n x. + Definition P := F_with simple_extendsion. + + Definition R_simple P e := (~ exhaustible (W e)) -> ~ (W e) # P. - Definition R P e := infity (W' e) -> exists w, (W' e w) /\ P w. + Section StrongInduction. - Lemma P_meet_R : forall n, R P' n. + Definition strong_induction (p: nat -> Type) : + (forall x, (forall y, y < x -> p y) -> p x) -> forall x, p x. + Proof. + intros H x; apply H. + induction x; [intros; lia| ]. + intros; apply H; intros; apply IHx; lia. + Defined. + + End StrongInduction. + + Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. + + Lemma P_meet_R_simple : forall e, R_simple P e. Proof. + strong induction e. + intros Hex Hdisj. + unfold disj in Hdisj. Admitted. - Lemma P_semi_decidable : semi_decidable P'. + Lemma P_semi_decidable : semi_decidable P. + Proof. + apply F_with_semi_decidable. + Qed. + + Lemma P_coinfinite : ~ exhaustible (compl P). Proof. Admitted. + Lemma P_simple : simple P. + Proof. + unfold simple; repeat split. + - rewrite EA.enum_iff. now apply P_semi_decidable. + - apply P_coinfinite. + - intros [q (Hqenum & Hqinf & Hq)]. + rewrite EA.enum_iff in Hqenum. + destruct (W_spec Hqenum) as [c HWq]. + apply (@P_meet_R_simple c). + intros [l Hqe]; apply Hqinf. + exists l. intros x Hqx. apply (Hqe x). + now rewrite HWq in Hqx. + intros x HWcx HPx. unfold W in HWcx. + rewrite <- HWq in HWcx. apply (Hq x HWcx HPx). + Qed. + + End SimpleSet. End Construction. From 12e8cfc583464a4da5beab9b2a5091fb286362f8 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 6 Feb 2024 09:03:04 +0100 Subject: [PATCH 11/54] typo --- theories/ReducibilityDegrees/lowsimple.v | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index c93da3f..ed53ba4 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -81,7 +81,7 @@ Section Construction. Notation "A # B" := (disj A B) (at level 30). Notation "A l# B" := (disj__l A B) (at level 30). - Record Extendsion := + Record Extension := { extendP: list nat -> nat -> option nat -> Prop; extend_dec: forall l x, Σ y, extendP l x y; @@ -94,7 +94,7 @@ Section Construction. | S n => P_templete E n x \/ E (P_templete E n) n x end. - Fixpoint F_ (E: Extendsion) n l: Prop := + Fixpoint F_ (E: Extension) n l: Prop := match n with | O => l = [] | S n => exists ls a, F_ E n ls /\ (extendP E ls n a) /\ l = if a is Some a then a :: ls else ls @@ -157,9 +157,11 @@ Section Construction. Section SimpleSet. Definition extend_simple l__n n x := - exists e, mu e (fun alpha => e < n /\ (l__n l# W_ alpha n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). + exists e, mu e (fun alpha => e < n /\ + (l__n l# W_ alpha n) /\ + mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). - Definition simple_extendsion: Extendsion. + Definition simple_extendsion: Extension. Proof. exists (fun l n a => forall x, a = Some x <-> extend_simple l n x). - intros l x. admit. From ac9a3275bd5704cc1fc344024bdd71b12742bc03 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Fri, 15 Mar 2024 10:16:55 +0100 Subject: [PATCH 12/54] update --- theories/ReducibilityDegrees/lowsimple.v | 566 +++++++++++++++++++---- 1 file changed, 488 insertions(+), 78 deletions(-) diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index ed53ba4..690b1a6 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -1,5 +1,9 @@ From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. - +Require Import SyntheticComputability.Synthetic.DecidabilityFacts. +Require Export SyntheticComputability.Shared.FinitenessFacts. +Require Export SyntheticComputability.Shared.Pigeonhole. +Require Export SyntheticComputability.Shared.ListAutomation. +Require Import Arith. Require Import Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. @@ -35,7 +39,8 @@ Section LowFacts. limit_computable (A´) -> ~ K ⪯ᴛ A. Proof. intros LEM defK H IH. - apply lowness with (P := A); [|apply IH]. + apply lowness with (P := A); [|apply IH]. + (* specialize (limit_turing_red_K _ _ _ _ defK H). *) Admitted. End LowFacts. @@ -66,149 +71,554 @@ Notation "'Σ' x .. y , p" := format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") : type_scope. +Section StrongInduction. + + Definition strong_induction (p: nat -> Type) : + (forall x, (forall y, y < x -> p y) -> p x) -> forall x, p x. + Proof. + intros H x; apply H. + induction x; [intros; lia| ]. + intros; apply H; intros; apply IHx; lia. + Defined. + +End StrongInduction. + +Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. + Require Import SyntheticComputability.Shared.FinitenessFacts. -Section Construction. - Definition mu n (P: nat -> Prop) := P n /\ forall m, P m -> n <= m. +Section ComplToBound. +Definition complToBound L b : list nat + := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). + +Lemma complToBound_Bound L b : + forall x, In x (complToBound L b) -> x <= b. +Proof. + intros x [H % in_seq ?] % in_filter_iff. lia. +Qed. +Lemma filter_length {X} f (l : list X) : + length l = length (filter f l) + length (filter (fun x => (negb (f x))) l). +Proof. + induction l; cbn. + - reflexivity. + - destruct f; cbn; lia. +Qed. +Lemma filter_NoDup {X} f (l : list X) : + NoDup l -> NoDup (filter f l). +Proof. + induction 1; cbn. + - econstructor. + - destruct f; eauto. econstructor; auto. + intros ? % in_filter_iff. firstorder. +Qed. +Lemma complToBound_length L b: + length (complToBound L b) + length L >= S b. +Proof. + rewrite <- (seq_length (S b) 0). + erewrite filter_length with (l := seq 0 (S b)). + unfold complToBound. + eapply Plus.plus_le_compat_l_stt. + generalize (seq_NoDup (S b) 0). + generalize (seq 0 (S b)). clear. + intros. erewrite filter_ext with (g := fun x => Dec (In x L)). + 2:{ intros a. destruct Dec; cbn; destruct Dec; firstorder congruence. } + eapply NoDup_incl_length. now eapply filter_NoDup. + clear. induction l; cbn. + - firstorder. + - destruct Dec; cbn. 2: eauto. + intros ? [-> | ]; eauto. +Qed. +Lemma complToBound_NoDup L b: + NoDup (complToBound L b). +Proof. + eapply filter_NoDup, seq_NoDup. +Qed. +Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) -> In x l. +Proof. + induction n in x, l |- *; destruct l; cbn; firstorder. +Qed. + +Lemma firstn_NoDup {X} (l : list X) n : NoDup l -> NoDup (firstn n l). +Proof. + induction 1 in n |- *; destruct n; cbn; try now econstructor. + econstructor; eauto. + now intros ? % firstn_In. +Qed. + +End ComplToBound. + +Section Construction. Variable W_: nat -> nat -> nat -> Prop. Hypothesis W_spec: forall P, semi_decidable P -> exists c, forall x, P x <-> exists n, W_ c n x. - Definition disj (A B: nat -> Prop) := forall x, A x -> B x -> False. - Definition disj__l (A: list nat) (B: nat -> Prop) := forall x, In x A -> B x -> False. - + Definition disj {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. + Definition intersect {X} (A B: X -> Prop) := forall x, A x -> B x -> False. Notation "A # B" := (disj A B) (at level 30). - Notation "A l# B" := (disj__l A B) (at level 30). Record Extension := { - extendP: list nat -> nat -> option nat -> Prop; - extend_dec: forall l x, Σ y, extendP l x y; + extendP: list nat -> nat -> nat -> Prop; + extend_dec: forall l x, (Σ y, extendP l x y) + (forall y, ~ extendP l x y); extend_uni: forall l x y1 y2, extendP l x y1 -> extendP l x y2 -> y1 = y2 }. - Fixpoint P_templete (E: (nat -> Prop) -> nat -> nat -> Prop) n x: Prop := - match n with - | O => False - | S n => P_templete E n x \/ E (P_templete E n) n x - end. - - Fixpoint F_ (E: Extension) n l: Prop := - match n with - | O => l = [] - | S n => exists ls a, F_ E n ls /\ (extendP E ls n a) /\ l = if a is Some a then a :: ls else ls - end. + Inductive F_ (E: Extension) : nat -> list nat -> Prop := + | Base : F_ E O [] + | ExtendS n l a : F_ E n l -> extendP E l n a -> F_ E (S n) (a::l) + | ExtendN n l : F_ E n l -> (forall a, ~ extendP E l n a) -> F_ E (S n) l. - Definition F_with E x := exists l n, In x l /\ (F_ E n l). + Require Import Coq.Program.Equality. Lemma F_uni E : forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . Proof. - induction n; simpl; [intros ?? -> ->; easy |]. - intros l1 l2 (ls & a & H1 & H2 & ->) (ls' & a' & H1' & H2' & ->). - rewrite !(IHn _ _ H1 H1') in *. - rewrite !(extend_uni H2 H2') in *. - easy. + intros n l1 l2. + dependent induction n. + - intros H1 H2. inv H1. now inv H2. + - intros H1 H2. inv H1; inv H2. + assert(l = l0) as -> by now apply IHn. + f_equal. apply (extend_uni H3 H4). + assert (l = l2) as -> by now apply IHn. + exfalso. apply (H4 a H3). + assert (l = l1) as -> by now apply IHn. + exfalso. apply (H3 a H4). + now apply IHn. Qed. - Lemma F_computable E : exists f: nat -> list nat, forall n, F_ E n (f n). + Lemma F_mono E n m l1 l2: F_ E n l1 -> F_ E m l2 -> n <= m -> incl l1 l2. + Proof. + intros H1 H2 HE. + revert H1 H2; induction HE in l2 |-*; intros H1 H2. + - now assert (l1 = l2) as -> by (eapply F_uni; eauto). + - inv H2; last now apply IHHE. + specialize (IHHE l H1 H0). eauto. + Qed. + + Lemma F_pop E n x l: F_ E n (x::l) -> exists m, F_ E m l. + Proof. + intros H. dependent induction H. + - now exists n. + - eapply IHF_. eauto. easy. + Qed. + + Lemma F_pick E n x l: F_ E n (x::l) -> exists m, F_ E m l /\ extendP E l m x. + Proof. + intros H. dependent induction H. + - exists n; eauto. + - eapply IHF_; eauto. + Qed. + + Lemma F_computable E : Σ f: nat -> list nat, + forall n, F_ E n (f n) /\ length (f n) <= n. Proof. set (F := fix f x := match x with | O => [] - | S x => match projT1 ((extend_dec E) (f x) x) with - | None => f x - | Some a => a :: (f x) + | S x => match (extend_dec E) (f x) x with + | inr _ => f x + | inl aH => (projT1 aH) :: (f x) end end). exists F. induction n; simpl. - - easy. - - exists (F n). destruct (extend_dec E (F n) n). - exists x. cbn. split; [eauto|]. split; [eauto|]. - easy. + - split. constructor. easy. + - destruct (extend_dec E (F n) n); split. + + eapply ExtendS; first apply IHn. now destruct s. + + cbn. destruct IHn. lia. + + now eapply ExtendN. + + destruct IHn. lia. Qed. - Definition decider: nat -> list nat -> bool. + Definition F_func E := projT1 (F_computable E). + Lemma F_func_correctness {E}: forall n, F_ E n (F_func E n). Proof. - intros x l. - destruct (ListDec.In_dec EqDec.nat_eq_dec x l). - - exact true. - exact false. - Defined. + intros n; unfold F_func. + destruct (F_computable E) as [f H]. + now destruct (H n). + Qed. - Fact decider_spec: forall x l, (decider x l) = true <-> In x l. + Lemma F_func_correctness' {E}: forall n, length (F_func E n) <= n. Proof. - intros x l. unfold decider. - destruct (ListDec.In_dec EqDec.nat_eq_dec x l); firstorder. + intros n; unfold F_func. + destruct (F_computable E) as [f H]. + now destruct (H n). Qed. + + + Definition F_with E x := exists l n, In x l /\ (F_ E n l). Lemma F_with_semi_decidable E: semi_decidable (F_with E). Proof. unfold semi_decidable, semi_decider. - destruct (F_computable E) as [f Hf]. - exists (fun x n => decider x (f n)). + destruct (F_computable E) as [f Hf ]. + exists (fun x n => (Dec (In x (f n)))). intros x. split. - intros (l & n & Hxl & Hl). - exists n. rewrite decider_spec. - now rewrite (F_uni (Hf n) Hl). - - intros (n & Hn%decider_spec). + exists n. rewrite Dec_auto; first easy. + destruct (Hf n) as [Hf' _]. + now rewrite (F_uni Hf' Hl). + - intros (n & Hn%Dec_true). exists (f n), n; split; eauto. + apply Hf. Qed. - Section SimpleSet. +Section SimpleSet. + + Notation unique p := (forall x x', p x -> p x' -> x = x'). + Definition safe p n := forall k, k < n -> ~ p k. + Definition least p n := p n /\ safe p n. + + Definition W e x := exists n, W_ e n x. + Notation " P ↾ s" := (fun x => exists n, n <= s /\ P n x) (at level 20). + - Definition extend_simple l__n n x := - exists e, mu e (fun alpha => e < n /\ - (l__n l# W_ alpha n) /\ - mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). + Definition ext_intersect_W L n e := L # (W_ e) ↾ n. + Definition ext_has_wit L n e x := (W_ e ↾ n) x /\ 2 * e < x /\ ~ In x L. + Definition ext_pick L n x e := ext_intersect_W L n e /\ least (ext_has_wit L n e) x. + Definition ext_least_choice L n x := exists e, least (ext_pick L n x) e. + + (* Definition extend_simple l__n n x := + exists e, mu e (fun α => e < n /\ + (l__n # W_ α n) /\ + mu x (fun β => W_ α n β /\ 2 * α < β /\ ~ In β l__n)). *) + + (* Definition extend_simple' l__n n x := + exists e, search e (fun α => (l__n # W_ α n) /\ + search x (fun beta => W_ α n beta /\ 2 * α < beta) n) n. *) + + + Variable hy1: forall (l : list nat) (x : nat), (Σ y : nat, ext_least_choice l x y) + (forall y : nat, ~ ext_least_choice l x y). + Variable hy2: forall (l : list nat) (x y1 y2 : nat), ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. Definition simple_extendsion: Extension. Proof. - exists (fun l n a => forall x, a = Some x <-> extend_simple l n x). - - intros l x. admit. - - intros l x y1 y2 Hy1 Hy2. - admit. - Admitted. - - Definition E_simple P__n n x := - exists e, mu e (fun alpha => e < n /\ (W_ alpha n # P__n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). + unshelve econstructor. + - exact ext_least_choice. + - apply hy1. + - apply hy2. + Defined. + (* Definition E_simple P__n n x := + exists e, + mu e (fun alpha => e < n /\ (W_ alpha n # P__n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). *) (* Definition P_ := P_templete E_simple. *) Definition P_ := F_ simple_extendsion. - Definition W e x := exists n, W_ e n x. + Definition Pf_ := F_func simple_extendsion. Definition P := F_with simple_extendsion. - Definition R_simple P e := (~ exhaustible (W e)) -> ~ (W e) # P. + Definition R_simple_list L e := (~ exhaustible (W e)) -> ~ (L # (W e)). - Section StrongInduction. + Definition attention e n := exists x, least (ext_pick (Pf_ n) n x) e. + Definition active e n := ~ (Pf_ n) # ((W_ e) ↾ n). - Definition strong_induction (p: nat -> Type) : - (forall x, (forall y, y < x -> p y) -> p x) -> forall x, p x. - Proof. - intros H x; apply H. - induction x; [intros; lia| ]. - intros; apply H; intros; apply IHx; lia. - Defined. - End StrongInduction. + Lemma W_incl e n m: + n <= m -> forall x, (W_ e ↾ n) x -> (W_ e ↾ m) x. + Proof. + intros H x [y [H1 H2]]. + exists y. split; lia + easy. + Qed. - Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. + Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : + L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. + Proof. + intros H h1 h2 x Hx1 Hx2. + eapply (H x). now eapply h1. + now apply h2. + Qed. - Lemma P_meet_R_simple : forall e, R_simple P e. + Lemma active_incl e n: active e n -> forall m, n <= m -> active e m . + Proof. + intros H m Hm Hx. apply H. + eapply (intersect_mono Hx). + eapply F_mono; try eapply F_func_correctness; easy. + now eapply W_incl. + Qed. + + Lemma attention_active e k: attention e k -> active e (S k). + Proof. + intros [x H] Hcontr. + apply (Hcontr x). + assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x). + exists e. assert (Pf_ k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + apply H. assert (x = a) as ->. + eapply (@extend_uni simple_extendsion); cbn; eauto. eauto. + exfalso. eapply H3. cbn. + exists e. enough ((Pf_ k) = (Pf_ (S k))) as <- by apply H. + assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. + eapply F_uni; eauto. + firstorder. + Qed. + + Lemma active_not_attention e k: active e k -> ~ attention e k. + Proof. + now intros H [x [[Hx _] _]]. + Qed. + + Lemma active_hold e k: attention e k -> forall m, k < m -> active e m. + Proof. + intros. eapply active_incl. + apply attention_active. apply H. lia. + Qed. + + Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. + Proof. + intros H1 H2. + specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H1 a b)) as H1'. + specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H2 a b)) as H2'. + enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. + intro Hk. eapply H1'. apply Hk. easy. + intro Hk. eapply H2'. apply Hk. easy. + Qed. + + Definition spec (e x: nat) := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. + + Lemma spec_uni e x1 x2: spec e x1 -> spec e x2 -> x1 = x2 . + Proof. + intros [k [Ht Hk]] [k' [Ht' Hk']]. + assert (k=k') as <-. eapply attention_uni; eauto. + eapply (@extend_uni simple_extendsion); cbn; eauto. + Qed. + + Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, spec e x /\ e < n. + Proof. + intros [[L [k [Hin Hk]]] Hn]. + dependent induction L. inv Hin. + destruct (Dec (a = x)) as [->|]. + - clear IHL Hin. + destruct (F_pick Hk) as [k' [Hk' [e He]]]. + exists e. split. unfold spec. + exists k'. split; unfold attention. + + exists x. enough (Pf_ k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + exists e. enough (Pf_ k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + destruct He. destruct H. destruct H1. destruct H1. + lia. + - destruct (F_pop Hk) as [m' Hm']. + eapply (IHL m'); eauto. firstorder. + Qed. + + Lemma P_extract_spec n L: + (forall x, In x L -> P x /\ x <= 2 * n) -> + forall x, In x L -> exists c, spec c x /\ c < n. + Proof. + intros. induction L. inv H0. + destruct H0 as [-> | Hln]; last apply IHL; eauto. + apply P_meet_spec. now apply H. + Qed. + + Lemma DomC_pullback_list n L: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> + exists (LC: list nat), NoDup LC /\ length LC = length L /\ + forall c, In c LC -> exists x, spec c x /\ In x L /\ c < n. + Proof. + intros HL H1. + induction L. + - exists []; intuition. + - remember (@P_extract_spec n (a::L) H1 a). + assert (In a (a::L)) as H0 by intuition. + apply e in H0 as [c0 [H0 E1]]. + assert (NoDup L) by (inversion HL; intuition). + apply IHL in H as [LC H]. + exists (c0::LC). intuition. + + constructor; last now exact H2. + intros In. inv HL. + apply H4 in In as [y (Hs & h2 & h3)]. + now rewrite (spec_uni H0 Hs) in H6. + + cbn. rewrite H. trivial. + + destruct H3 as [->|]. + * exists a; intuition. + * destruct (H4 c H3) as [y Hy]. + exists y; intuition. + + intros y In1. assert (In y (a::L)) by intuition. + now apply H1 in H2. + Qed. + + Definition PredListTo p : list nat -> nat -> Prop + := fun L b => forall x, In x L <-> p x /\ x <= b. + + Lemma NoDupBoundH {L} b: + NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). + Proof. + intros ND H x E. + constructor. + - intros H1 % H. lia. + - exact ND. + Qed. + + Lemma PredNoDupListTo_NNExist p: + forall b, ~~ exists L, PredListTo p L b /\ NoDup L. Proof. - strong induction e. - intros Hex Hdisj. - unfold disj in Hdisj. + destruct (F_computable simple_extendsion) as [f Hf]. + induction b; intros H. + - ccase (p 0) as [H0 | H0]; apply H. + + exists [0]. split; try split. + * intros [E | E]; (try contradiction E). + rewrite <- E. intuition. + * intros E. assert (x = 0) by lia. + rewrite H1. intuition. + * constructor; intuition; constructor. + + exists nil. split; try split. + * contradiction. + * intros E. assert (x = 0) by lia. + rewrite H1 in E. firstorder. + * constructor. + - apply IHb. intros [L H1]. + ccase (p (1 + b)) as [H0 | H0]; apply H. + + exists ((1+ b) :: L). split; try split. + * intros [E | E]; try (rewrite <- E; intuition). + apply H1 in E. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** right. apply H1. intuition. + ** left. lia. + * apply (@NoDupBoundH _ b). + ** apply H1. + ** intros x H3 % H1. lia. + ** lia. + + exists L. split; try split. + * intros E % H1. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** apply H1. intuition. + ** rewrite E in E1. firstorder. + * apply H1. + Qed. + + Lemma P_bounded L n: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. + Proof. + intros ND [LC H] % DomC_pullback_list; intuition. + rewrite <- H. + assert (incl LC (seq 0 n)). unfold incl. + - intros c [e [x Hx]] % H2. apply in_seq. lia. + - apply pigeonhole_length in H1. + + now rewrite seq_length in H1. + + intros. decide (x1 = x2); tauto. + + exact H0. + Qed. + + Definition R_simple P e := (~ exhaustible (W e)) -> ~ (intersect (W e) P). + + Lemma R_simple_acc e L: + R_simple_list L e -> forall L', incl L L' -> R_simple_list L' e . + Proof. + intros H L' HL' h1 h2. + apply H; first easy. + firstorder. + Qed. + + Lemma list_accu e: + (forall k, k < e -> exists L, R_simple_list L k /\ exists n, F_ simple_extendsion n L) -> + exists m L, F_ simple_extendsion m L /\ forall n, n < e -> R_simple_list L n. + Proof. + intros. induction e. + { exists 0, []; split; first econstructor. intros n Hn. lia. } + destruct (H e) as [L [HL1 HL2]]; first lia. + destruct IHe as [m [HL' [Hm HL2']]]. + { intros n Hn. apply H. lia. } + destruct HL2 as [k Hk]. destruct (dec_le m k). + + exists k, L; split; eauto. + intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. + apply HL1. eapply R_simple_acc. apply HL2'; first easy. + eapply F_mono; eauto. + + exists m, HL'; split; eauto. + intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. + eapply R_simple_acc. apply HL1. eapply F_mono; eauto with lia. + now eapply HL2'. + Qed. + + + Lemma eventually_attention e m L: + ~ exhaustible (W e) -> + (forall n : nat, n < e -> ~ exhaustible (W n) -> ~ L # W n -> F_ simple_extendsion m L) -> + exists k, attention e k. + Proof. + intros H1 H2. + unfold attention. + Admitted. + Hypothesis O_O: LEM. + + Lemma try2: + forall e, exists L, R_simple_list L e /\ exists m, F_ simple_extendsion m L. + Proof. + strong induction e. apply list_accu in H. + destruct H as (m & L & HL & HL'). + destruct (O_O (exhaustible (W e))). + - exists []. split. intros H'; eauto. + exists 0; econstructor. + - unfold R_simple_list in HL'. +Admitted. + + Lemma P_meet_R_simple : forall e, R_simple P e. + Proof. + intros e. + destruct (try2 e) as [L [H1 [m H]]]. + unfold R_simple. intros H3. + intros Hin. destruct (H1 H3). + intros x Hx Hx'. + apply (Hin x Hx'). + unfold P. unfold F_with. + exists L, m. now split. + Qed. + Lemma P_semi_decidable : semi_decidable P. Proof. apply F_with_semi_decidable. Qed. + Lemma P_Listing: + forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). + Proof. + intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). + intros [L H1]. apply H. exists L; intuition. + apply P_bounded. + - exact H2. + - apply H0. + Qed. + + Lemma complToBound_compl p L b: + PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. + Proof. + intros H x. split. + - intros [H1 H1'] % in_filter_iff. + destruct Dec; cbn in H1'; try congruence. + enough (x <= b). + + intuition. + + apply in_seq in H1. lia. + - intros [H1 H2]. eapply in_filter_iff. split. + + apply in_seq; lia. + + destruct Dec; cbn; try tauto. exfalso. firstorder. + Qed. + + Lemma ComplP_Listing: + forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L + /\ forall x, In x L -> ~ P x. + Proof. + intros n H. + apply (@P_Listing n). intros [L H1]. + apply H. exists (complToBound L (2*n)). repeat split. + - remember (complToBound_length L (2*n)). lia. + - apply complToBound_NoDup. + - intros x I % (@complToBound_compl P); intuition. + Qed. + Lemma P_coinfinite : ~ exhaustible (compl P). Proof. - Admitted. + eapply weakly_unbounded_non_finite. + intros n H. eapply ComplP_Listing with (n := n). + intros (l & ? & ? & H2). + eapply H. + exists (firstn n l). + repeat split. + - rewrite firstn_length. lia. + - now eapply firstn_NoDup. + - intros ? ? % firstn_In. now eapply H2. + Qed. Lemma P_simple : simple P. Proof. From 6399aa48f3b8170679f73c26a40f71e43e897514 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sun, 17 Mar 2024 14:48:56 +0100 Subject: [PATCH 13/54] refactor --- theories/ReducibilityDegrees/lowsimple.v | 580 +---------------- .../ReducibilityDegrees/lowsimple_extension.v | 0 .../ReducibilityDegrees/priority_method.v | 136 ++++ .../ReducibilityDegrees/simple_extension.v | 586 ++++++++++++++++++ 4 files changed, 723 insertions(+), 579 deletions(-) create mode 100644 theories/ReducibilityDegrees/lowsimple_extension.v create mode 100644 theories/ReducibilityDegrees/priority_method.v create mode 100644 theories/ReducibilityDegrees/simple_extension.v diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index 690b1a6..00c638a 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -60,582 +60,4 @@ Proof. split; [destruct H2 as [H2 _]; eauto| now apply lowness]. Qed. -End LowSimplePredicate. - - -Require Import List. -Import ListNotations. -Notation "'Σ' x .. y , p" := - (sigT (fun x => .. (sigT (fun y => p)) ..)) - (at level 200, x binder, right associativity, - format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") - : type_scope. - -Section StrongInduction. - - Definition strong_induction (p: nat -> Type) : - (forall x, (forall y, y < x -> p y) -> p x) -> forall x, p x. - Proof. - intros H x; apply H. - induction x; [intros; lia| ]. - intros; apply H; intros; apply IHx; lia. - Defined. - -End StrongInduction. - -Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. - -Require Import SyntheticComputability.Shared.FinitenessFacts. - - -Section ComplToBound. -Definition complToBound L b : list nat - := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). - -Lemma complToBound_Bound L b : - forall x, In x (complToBound L b) -> x <= b. -Proof. - intros x [H % in_seq ?] % in_filter_iff. lia. -Qed. -Lemma filter_length {X} f (l : list X) : - length l = length (filter f l) + length (filter (fun x => (negb (f x))) l). -Proof. - induction l; cbn. - - reflexivity. - - destruct f; cbn; lia. -Qed. -Lemma filter_NoDup {X} f (l : list X) : - NoDup l -> NoDup (filter f l). -Proof. - induction 1; cbn. - - econstructor. - - destruct f; eauto. econstructor; auto. - intros ? % in_filter_iff. firstorder. -Qed. -Lemma complToBound_length L b: - length (complToBound L b) + length L >= S b. -Proof. - rewrite <- (seq_length (S b) 0). - erewrite filter_length with (l := seq 0 (S b)). - unfold complToBound. - eapply Plus.plus_le_compat_l_stt. - generalize (seq_NoDup (S b) 0). - generalize (seq 0 (S b)). clear. - intros. erewrite filter_ext with (g := fun x => Dec (In x L)). - 2:{ intros a. destruct Dec; cbn; destruct Dec; firstorder congruence. } - eapply NoDup_incl_length. now eapply filter_NoDup. - clear. induction l; cbn. - - firstorder. - - destruct Dec; cbn. 2: eauto. - intros ? [-> | ]; eauto. -Qed. -Lemma complToBound_NoDup L b: - NoDup (complToBound L b). -Proof. - eapply filter_NoDup, seq_NoDup. -Qed. -Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) -> In x l. -Proof. - induction n in x, l |- *; destruct l; cbn; firstorder. -Qed. - -Lemma firstn_NoDup {X} (l : list X) n : NoDup l -> NoDup (firstn n l). -Proof. - induction 1 in n |- *; destruct n; cbn; try now econstructor. - econstructor; eauto. - now intros ? % firstn_In. -Qed. - -End ComplToBound. - -Section Construction. - - Variable W_: nat -> nat -> nat -> Prop. - Hypothesis W_spec: forall P, semi_decidable P -> exists c, forall x, P x <-> exists n, W_ c n x. - - Definition disj {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. - Definition intersect {X} (A B: X -> Prop) := forall x, A x -> B x -> False. - Notation "A # B" := (disj A B) (at level 30). - - Record Extension := - { - extendP: list nat -> nat -> nat -> Prop; - extend_dec: forall l x, (Σ y, extendP l x y) + (forall y, ~ extendP l x y); - extend_uni: forall l x y1 y2, extendP l x y1 -> extendP l x y2 -> y1 = y2 - }. - - Inductive F_ (E: Extension) : nat -> list nat -> Prop := - | Base : F_ E O [] - | ExtendS n l a : F_ E n l -> extendP E l n a -> F_ E (S n) (a::l) - | ExtendN n l : F_ E n l -> (forall a, ~ extendP E l n a) -> F_ E (S n) l. - - Require Import Coq.Program.Equality. - - Lemma F_uni E : forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . - Proof. - intros n l1 l2. - dependent induction n. - - intros H1 H2. inv H1. now inv H2. - - intros H1 H2. inv H1; inv H2. - assert(l = l0) as -> by now apply IHn. - f_equal. apply (extend_uni H3 H4). - assert (l = l2) as -> by now apply IHn. - exfalso. apply (H4 a H3). - assert (l = l1) as -> by now apply IHn. - exfalso. apply (H3 a H4). - now apply IHn. - Qed. - - Lemma F_mono E n m l1 l2: F_ E n l1 -> F_ E m l2 -> n <= m -> incl l1 l2. - Proof. - intros H1 H2 HE. - revert H1 H2; induction HE in l2 |-*; intros H1 H2. - - now assert (l1 = l2) as -> by (eapply F_uni; eauto). - - inv H2; last now apply IHHE. - specialize (IHHE l H1 H0). eauto. - Qed. - - Lemma F_pop E n x l: F_ E n (x::l) -> exists m, F_ E m l. - Proof. - intros H. dependent induction H. - - now exists n. - - eapply IHF_. eauto. easy. - Qed. - - Lemma F_pick E n x l: F_ E n (x::l) -> exists m, F_ E m l /\ extendP E l m x. - Proof. - intros H. dependent induction H. - - exists n; eauto. - - eapply IHF_; eauto. - Qed. - - Lemma F_computable E : Σ f: nat -> list nat, - forall n, F_ E n (f n) /\ length (f n) <= n. - Proof. - set (F := fix f x := - match x with - | O => [] - | S x => match (extend_dec E) (f x) x with - | inr _ => f x - | inl aH => (projT1 aH) :: (f x) - end - end). - exists F. induction n; simpl. - - split. constructor. easy. - - destruct (extend_dec E (F n) n); split. - + eapply ExtendS; first apply IHn. now destruct s. - + cbn. destruct IHn. lia. - + now eapply ExtendN. - + destruct IHn. lia. - Qed. - - Definition F_func E := projT1 (F_computable E). - Lemma F_func_correctness {E}: forall n, F_ E n (F_func E n). - Proof. - intros n; unfold F_func. - destruct (F_computable E) as [f H]. - now destruct (H n). - Qed. - - Lemma F_func_correctness' {E}: forall n, length (F_func E n) <= n. - Proof. - intros n; unfold F_func. - destruct (F_computable E) as [f H]. - now destruct (H n). - Qed. - - - Definition F_with E x := exists l n, In x l /\ (F_ E n l). - - Lemma F_with_semi_decidable E: semi_decidable (F_with E). - Proof. - unfold semi_decidable, semi_decider. - destruct (F_computable E) as [f Hf ]. - exists (fun x n => (Dec (In x (f n)))). - intros x. split. - - intros (l & n & Hxl & Hl). - exists n. rewrite Dec_auto; first easy. - destruct (Hf n) as [Hf' _]. - now rewrite (F_uni Hf' Hl). - - intros (n & Hn%Dec_true). - exists (f n), n; split; eauto. - apply Hf. - Qed. - -Section SimpleSet. - - Notation unique p := (forall x x', p x -> p x' -> x = x'). - Definition safe p n := forall k, k < n -> ~ p k. - Definition least p n := p n /\ safe p n. - - Definition W e x := exists n, W_ e n x. - Notation " P ↾ s" := (fun x => exists n, n <= s /\ P n x) (at level 20). - - - Definition ext_intersect_W L n e := L # (W_ e) ↾ n. - Definition ext_has_wit L n e x := (W_ e ↾ n) x /\ 2 * e < x /\ ~ In x L. - - Definition ext_pick L n x e := ext_intersect_W L n e /\ least (ext_has_wit L n e) x. - Definition ext_least_choice L n x := exists e, least (ext_pick L n x) e. - - (* Definition extend_simple l__n n x := - exists e, mu e (fun α => e < n /\ - (l__n # W_ α n) /\ - mu x (fun β => W_ α n β /\ 2 * α < β /\ ~ In β l__n)). *) - - (* Definition extend_simple' l__n n x := - exists e, search e (fun α => (l__n # W_ α n) /\ - search x (fun beta => W_ α n beta /\ 2 * α < beta) n) n. *) - - - Variable hy1: forall (l : list nat) (x : nat), (Σ y : nat, ext_least_choice l x y) + (forall y : nat, ~ ext_least_choice l x y). - Variable hy2: forall (l : list nat) (x y1 y2 : nat), ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. - Definition simple_extendsion: Extension. - Proof. - unshelve econstructor. - - exact ext_least_choice. - - apply hy1. - - apply hy2. - Defined. - - (* Definition E_simple P__n n x := - exists e, - mu e (fun alpha => e < n /\ (W_ alpha n # P__n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). *) - -(* Definition P_ := P_templete E_simple. *) - Definition P_ := F_ simple_extendsion. - Definition Pf_ := F_func simple_extendsion. - Definition P := F_with simple_extendsion. - - Definition R_simple_list L e := (~ exhaustible (W e)) -> ~ (L # (W e)). - - Definition attention e n := exists x, least (ext_pick (Pf_ n) n x) e. - Definition active e n := ~ (Pf_ n) # ((W_ e) ↾ n). - - - Lemma W_incl e n m: - n <= m -> forall x, (W_ e ↾ n) x -> (W_ e ↾ m) x. - Proof. - intros H x [y [H1 H2]]. - exists y. split; lia + easy. - Qed. - - Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : - L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. - Proof. - intros H h1 h2 x Hx1 Hx2. - eapply (H x). now eapply h1. - now apply h2. - Qed. - - Lemma active_incl e n: active e n -> forall m, n <= m -> active e m . - Proof. - intros H m Hm Hx. apply H. - eapply (intersect_mono Hx). - eapply F_mono; try eapply F_func_correctness; easy. - now eapply W_incl. - Qed. - - Lemma attention_active e k: attention e k -> active e (S k). - Proof. - intros [x H] Hcontr. - apply (Hcontr x). - assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. - inv H0. cbn in H4. assert (ext_least_choice l k x). - exists e. assert (Pf_ k = l) as <-. - eapply F_uni. apply F_func_correctness. apply H3. - apply H. assert (x = a) as ->. - eapply (@extend_uni simple_extendsion); cbn; eauto. eauto. - exfalso. eapply H3. cbn. - exists e. enough ((Pf_ k) = (Pf_ (S k))) as <- by apply H. - assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. - eapply F_uni; eauto. - firstorder. - Qed. - - Lemma active_not_attention e k: active e k -> ~ attention e k. - Proof. - now intros H [x [[Hx _] _]]. - Qed. - - Lemma active_hold e k: attention e k -> forall m, k < m -> active e m. - Proof. - intros. eapply active_incl. - apply attention_active. apply H. lia. - Qed. - - Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. - Proof. - intros H1 H2. - specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H1 a b)) as H1'. - specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H2 a b)) as H2'. - enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. - intro Hk. eapply H1'. apply Hk. easy. - intro Hk. eapply H2'. apply Hk. easy. - Qed. - - Definition spec (e x: nat) := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. - - Lemma spec_uni e x1 x2: spec e x1 -> spec e x2 -> x1 = x2 . - Proof. - intros [k [Ht Hk]] [k' [Ht' Hk']]. - assert (k=k') as <-. eapply attention_uni; eauto. - eapply (@extend_uni simple_extendsion); cbn; eauto. - Qed. - - Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, spec e x /\ e < n. - Proof. - intros [[L [k [Hin Hk]]] Hn]. - dependent induction L. inv Hin. - destruct (Dec (a = x)) as [->|]. - - clear IHL Hin. - destruct (F_pick Hk) as [k' [Hk' [e He]]]. - exists e. split. unfold spec. - exists k'. split; unfold attention. - + exists x. enough (Pf_ k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + exists e. enough (Pf_ k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + destruct He. destruct H. destruct H1. destruct H1. - lia. - - destruct (F_pop Hk) as [m' Hm']. - eapply (IHL m'); eauto. firstorder. - Qed. - - Lemma P_extract_spec n L: - (forall x, In x L -> P x /\ x <= 2 * n) -> - forall x, In x L -> exists c, spec c x /\ c < n. - Proof. - intros. induction L. inv H0. - destruct H0 as [-> | Hln]; last apply IHL; eauto. - apply P_meet_spec. now apply H. - Qed. - - Lemma DomC_pullback_list n L: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> - exists (LC: list nat), NoDup LC /\ length LC = length L /\ - forall c, In c LC -> exists x, spec c x /\ In x L /\ c < n. - Proof. - intros HL H1. - induction L. - - exists []; intuition. - - remember (@P_extract_spec n (a::L) H1 a). - assert (In a (a::L)) as H0 by intuition. - apply e in H0 as [c0 [H0 E1]]. - assert (NoDup L) by (inversion HL; intuition). - apply IHL in H as [LC H]. - exists (c0::LC). intuition. - + constructor; last now exact H2. - intros In. inv HL. - apply H4 in In as [y (Hs & h2 & h3)]. - now rewrite (spec_uni H0 Hs) in H6. - + cbn. rewrite H. trivial. - + destruct H3 as [->|]. - * exists a; intuition. - * destruct (H4 c H3) as [y Hy]. - exists y; intuition. - + intros y In1. assert (In y (a::L)) by intuition. - now apply H1 in H2. - Qed. - - Definition PredListTo p : list nat -> nat -> Prop - := fun L b => forall x, In x L <-> p x /\ x <= b. - - Lemma NoDupBoundH {L} b: - NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). - Proof. - intros ND H x E. - constructor. - - intros H1 % H. lia. - - exact ND. - Qed. - - Lemma PredNoDupListTo_NNExist p: - forall b, ~~ exists L, PredListTo p L b /\ NoDup L. - Proof. - destruct (F_computable simple_extendsion) as [f Hf]. - induction b; intros H. - - ccase (p 0) as [H0 | H0]; apply H. - + exists [0]. split; try split. - * intros [E | E]; (try contradiction E). - rewrite <- E. intuition. - * intros E. assert (x = 0) by lia. - rewrite H1. intuition. - * constructor; intuition; constructor. - + exists nil. split; try split. - * contradiction. - * intros E. assert (x = 0) by lia. - rewrite H1 in E. firstorder. - * constructor. - - apply IHb. intros [L H1]. - ccase (p (1 + b)) as [H0 | H0]; apply H. - + exists ((1+ b) :: L). split; try split. - * intros [E | E]; try (rewrite <- E; intuition). - apply H1 in E. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** right. apply H1. intuition. - ** left. lia. - * apply (@NoDupBoundH _ b). - ** apply H1. - ** intros x H3 % H1. lia. - ** lia. - + exists L. split; try split. - * intros E % H1. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** apply H1. intuition. - ** rewrite E in E1. firstorder. - * apply H1. - Qed. - - Lemma P_bounded L n: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. - Proof. - intros ND [LC H] % DomC_pullback_list; intuition. - rewrite <- H. - assert (incl LC (seq 0 n)). unfold incl. - - intros c [e [x Hx]] % H2. apply in_seq. lia. - - apply pigeonhole_length in H1. - + now rewrite seq_length in H1. - + intros. decide (x1 = x2); tauto. - + exact H0. - Qed. - - Definition R_simple P e := (~ exhaustible (W e)) -> ~ (intersect (W e) P). - - Lemma R_simple_acc e L: - R_simple_list L e -> forall L', incl L L' -> R_simple_list L' e . - Proof. - intros H L' HL' h1 h2. - apply H; first easy. - firstorder. - Qed. - - Lemma list_accu e: - (forall k, k < e -> exists L, R_simple_list L k /\ exists n, F_ simple_extendsion n L) -> - exists m L, F_ simple_extendsion m L /\ forall n, n < e -> R_simple_list L n. - Proof. - intros. induction e. - { exists 0, []; split; first econstructor. intros n Hn. lia. } - destruct (H e) as [L [HL1 HL2]]; first lia. - destruct IHe as [m [HL' [Hm HL2']]]. - { intros n Hn. apply H. lia. } - destruct HL2 as [k Hk]. destruct (dec_le m k). - + exists k, L; split; eauto. - intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. - apply HL1. eapply R_simple_acc. apply HL2'; first easy. - eapply F_mono; eauto. - + exists m, HL'; split; eauto. - intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. - eapply R_simple_acc. apply HL1. eapply F_mono; eauto with lia. - now eapply HL2'. - Qed. - - - Lemma eventually_attention e m L: - ~ exhaustible (W e) -> - (forall n : nat, n < e -> ~ exhaustible (W n) -> ~ L # W n -> F_ simple_extendsion m L) -> - exists k, attention e k. - Proof. - intros H1 H2. - unfold attention. - - Admitted. - - Hypothesis O_O: LEM. - - Lemma try2: - forall e, exists L, R_simple_list L e /\ exists m, F_ simple_extendsion m L. - Proof. - strong induction e. apply list_accu in H. - destruct H as (m & L & HL & HL'). - destruct (O_O (exhaustible (W e))). - - exists []. split. intros H'; eauto. - exists 0; econstructor. - - unfold R_simple_list in HL'. -Admitted. - - Lemma P_meet_R_simple : forall e, R_simple P e. - Proof. - intros e. - destruct (try2 e) as [L [H1 [m H]]]. - unfold R_simple. intros H3. - intros Hin. destruct (H1 H3). - intros x Hx Hx'. - apply (Hin x Hx'). - unfold P. unfold F_with. - exists L, m. now split. - Qed. - - Lemma P_semi_decidable : semi_decidable P. - Proof. - apply F_with_semi_decidable. - Qed. - - Lemma P_Listing: - forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). - Proof. - intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). - intros [L H1]. apply H. exists L; intuition. - apply P_bounded. - - exact H2. - - apply H0. - Qed. - - Lemma complToBound_compl p L b: - PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. - Proof. - intros H x. split. - - intros [H1 H1'] % in_filter_iff. - destruct Dec; cbn in H1'; try congruence. - enough (x <= b). - + intuition. - + apply in_seq in H1. lia. - - intros [H1 H2]. eapply in_filter_iff. split. - + apply in_seq; lia. - + destruct Dec; cbn; try tauto. exfalso. firstorder. - Qed. - - Lemma ComplP_Listing: - forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L - /\ forall x, In x L -> ~ P x. - Proof. - intros n H. - apply (@P_Listing n). intros [L H1]. - apply H. exists (complToBound L (2*n)). repeat split. - - remember (complToBound_length L (2*n)). lia. - - apply complToBound_NoDup. - - intros x I % (@complToBound_compl P); intuition. - Qed. - - Lemma P_coinfinite : ~ exhaustible (compl P). - Proof. - eapply weakly_unbounded_non_finite. - intros n H. eapply ComplP_Listing with (n := n). - intros (l & ? & ? & H2). - eapply H. - exists (firstn n l). - repeat split. - - rewrite firstn_length. lia. - - now eapply firstn_NoDup. - - intros ? ? % firstn_In. now eapply H2. - Qed. - - Lemma P_simple : simple P. - Proof. - unfold simple; repeat split. - - rewrite EA.enum_iff. now apply P_semi_decidable. - - apply P_coinfinite. - - intros [q (Hqenum & Hqinf & Hq)]. - rewrite EA.enum_iff in Hqenum. - destruct (W_spec Hqenum) as [c HWq]. - apply (@P_meet_R_simple c). - intros [l Hqe]; apply Hqinf. - exists l. intros x Hqx. apply (Hqe x). - now rewrite HWq in Hqx. - intros x HWcx HPx. unfold W in HWcx. - rewrite <- HWq in HWcx. apply (Hq x HWcx HPx). - Qed. - - End SimpleSet. - -End Construction. +End LowSimplePredicate. \ No newline at end of file diff --git a/theories/ReducibilityDegrees/lowsimple_extension.v b/theories/ReducibilityDegrees/lowsimple_extension.v new file mode 100644 index 0000000..e69de29 diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v new file mode 100644 index 0000000..cd0fe45 --- /dev/null +++ b/theories/ReducibilityDegrees/priority_method.v @@ -0,0 +1,136 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +Require Import SyntheticComputability.Synthetic.DecidabilityFacts. +Require Export SyntheticComputability.Shared.FinitenessFacts. +Require Export SyntheticComputability.Shared.Pigeonhole. +Require Export SyntheticComputability.Shared.ListAutomation. +Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. +Import ListNotations. + +Notation "'Σ' x .. y , p" := + (sigT (fun x => .. (sigT (fun y => p)) ..)) + (at level 200, x binder, right associativity, + format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") + : type_scope. + +Definition disj {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. +Definition intersect {X} (A B: X -> Prop) := forall x, A x -> B x -> False. +Notation "A # B" := (disj A B) (at level 30). + +Section Construction. + Record Extension := + { + extendP: list nat -> nat -> nat -> Prop; + extend_dec: forall l x, (Σ y, extendP l x y) + (forall y, ~ extendP l x y); + extend_uni: forall l x y1 y2, extendP l x y1 -> extendP l x y2 -> y1 = y2 + }. + + Inductive F_ (E: Extension) : nat -> list nat -> Prop := + | Base : F_ E O [] + | ExtendS n l a : F_ E n l -> extendP E l n a -> F_ E (S n) (a::l) + | ExtendN n l : F_ E n l -> (forall a, ~ extendP E l n a) -> F_ E (S n) l. + + Lemma F_uni E : forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . + Proof. + intros n l1 l2. + dependent induction n. + - intros H1 H2. inv H1. now inv H2. + - intros H1 H2. inv H1; inv H2. + assert(l = l0) as -> by now apply IHn. + f_equal. apply (extend_uni H3 H4). + assert (l = l2) as -> by now apply IHn. + exfalso. apply (H4 a H3). + assert (l = l1) as -> by now apply IHn. + exfalso. apply (H3 a H4). + now apply IHn. + Qed. + + Lemma F_mono E n m l1 l2: F_ E n l1 -> F_ E m l2 -> n <= m -> incl l1 l2. + Proof. + intros H1 H2 HE. + revert H1 H2; induction HE in l2 |-*; intros H1 H2. + - now assert (l1 = l2) as -> by (eapply F_uni; eauto). + - inv H2; last now apply IHHE. + specialize (IHHE l H1 H0). eauto. + Qed. + + Lemma F_pop E n x l: F_ E n (x::l) -> exists m, F_ E m l. + Proof. + intros H. dependent induction H. + - now exists n. + - eapply IHF_. eauto. + Qed. + + Lemma F_pick E n x l: F_ E n (x::l) -> exists m, F_ E m l /\ extendP E l m x. + Proof. + intros H. dependent induction H. + - exists n; eauto. + - eapply IHF_; eauto. + Qed. + + Lemma F_computable E : Σ f: nat -> list nat, + forall n, F_ E n (f n) /\ length (f n) <= n. + Proof. + set (F := fix f x := + match x with + | O => [] + | S x => match (extend_dec E) (f x) x with + | inr _ => f x + | inl aH => (projT1 aH) :: (f x) + end + end). + exists F. induction n; simpl. + - split. constructor. easy. + - destruct (extend_dec E (F n) n); split. + + eapply ExtendS; first apply IHn. now destruct s. + + cbn. destruct IHn. lia. + + now eapply ExtendN. + + destruct IHn. lia. + Qed. + + Definition F_func E := projT1 (F_computable E). + Lemma F_func_correctness {E}: forall n, F_ E n (F_func E n). + Proof. + intros n; unfold F_func. + destruct (F_computable E) as [f H]. + now destruct (H n). + Qed. + + Lemma F_func_correctness' {E}: forall n, length (F_func E n) <= n. + Proof. + intros n; unfold F_func. + destruct (F_computable E) as [f H]. + now destruct (H n). + Qed. + + Definition F_with E x := exists l n, In x l /\ (F_ E n l). + + Lemma F_with_semi_decidable E: semi_decidable (F_with E). + Proof. + unfold semi_decidable, semi_decider. + destruct (F_computable E) as [f Hf ]. + exists (fun x n => (Dec (In x (f n)))). + intros x. split. + - intros (l & n & Hxl & Hl). + exists n. rewrite Dec_auto; first easy. + destruct (Hf n) as [Hf' _]. + now rewrite (F_uni Hf' Hl). + - intros (n & Hn%Dec_true). + exists (f n), n; split; eauto. + apply Hf. + Qed. + +End Construction. + +Section StrongInduction. + + Definition strong_induction (p: nat -> Type) : + (forall x, (forall y, y < x -> p y) -> p x) -> forall x, p x. + Proof. + intros H x; apply H. + induction x; [intros; lia| ]. + intros; apply H; intros; apply IHx; lia. + Defined. + +End StrongInduction. + +Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. \ No newline at end of file diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v new file mode 100644 index 0000000..c08f0cc --- /dev/null +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -0,0 +1,586 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +Require Import SyntheticComputability.Synthetic.DecidabilityFacts. +Require Export SyntheticComputability.Shared.FinitenessFacts. +Require Export SyntheticComputability.Shared.Pigeonhole. +Require Export SyntheticComputability.Shared.ListAutomation. +Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. +Require Import SyntheticComputability.ReducibilityDegrees.priority_method. +Import ListNotations. + +Section ComplToBound. + Definition complToBound L b : list nat + := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). + + Lemma complToBound_Bound L b : + forall x, In x (complToBound L b) -> x <= b. + Proof. + intros x [H % in_seq ?] % in_filter_iff. lia. + Qed. + Lemma filter_length {X} f (l : list X) : + length l = length (filter f l) + length (filter (fun x => (negb (f x))) l). + Proof. + induction l; cbn. + - reflexivity. + - destruct f; cbn; lia. + Qed. + Lemma filter_NoDup {X} f (l : list X) : + NoDup l -> NoDup (filter f l). + Proof. + induction 1; cbn. + - econstructor. + - destruct f; eauto. econstructor; auto. + intros ? % in_filter_iff. firstorder. + Qed. + Lemma complToBound_length L b: + length (complToBound L b) + length L >= S b. + Proof. + rewrite <- (seq_length (S b) 0). + erewrite filter_length with (l := seq 0 (S b)). + unfold complToBound. + eapply Plus.plus_le_compat_l_stt. + generalize (seq_NoDup (S b) 0). + generalize (seq 0 (S b)). clear. + intros. erewrite filter_ext with (g := fun x => Dec (In x L)). + 2:{ intros a. destruct Dec; cbn; destruct Dec; firstorder congruence. } + eapply NoDup_incl_length. now eapply filter_NoDup. + clear. induction l; cbn. + - firstorder. + - destruct Dec; cbn. 2: eauto. + intros ? [-> | ]; eauto. + Qed. + Lemma complToBound_NoDup L b: + NoDup (complToBound L b). + Proof. + eapply filter_NoDup, seq_NoDup. + Qed. + Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) -> In x l. + Proof. + induction n in x, l |- *; destruct l; cbn; firstorder. + Qed. + + Lemma firstn_NoDup {X} (l : list X) n : NoDup l -> NoDup (firstn n l). + Proof. + induction 1 in n |- *; destruct n; cbn; try now econstructor. + econstructor; eauto. + now intros ? % firstn_In. + Qed. +End ComplToBound. + +Notation unique p := (forall x x', p x -> p x' -> x = x'). +Section LeastWitness. + Definition safe p n := forall k, k < n -> ~ p k. + Definition least p n := p n /\ safe p n. +End LeastWitness. + + +Section Assume_EA. + +Section Requirements. + + Definition W e x := exists n, W_ e n x. + Notation " P ↾ s" := (fun x => exists n, n <= s /\ P n x) (at level 20). + + Definition ext_intersect_W L n e := L # (W_ e) ↾ n. + +End Requirements. + + Lemma exists_bounded_dec' P: + (forall x, dec (P x)) -> forall k, dec (exists n, n < k /\ P n). + Proof. + intros Hp k. + induction k. right; intros [? [? _]]; lia. + destruct IHk as [Hw|Hw]. + - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. + - destruct (Hp k). left. exists k; split; eauto; lia. + right. intros [x [Hx1 Hx2]]. + assert (x = k \/ x < k) as [->|Hk] by lia; firstorder. + Qed. + + Lemma exists_bounded_dec P: + (forall x, dec (P x)) -> forall k, dec (exists n, n <= k /\ P n). + Proof. + intros Hp k. + induction k. destruct (Hp 0). left. exists 0. eauto. + right. intros [x [Hx Hx']]. now assert (x=0) as -> by lia. + destruct IHk as [Hw|Hw]. + - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. + - destruct (Hp (S k)). left. exists (S k); split; eauto; lia. + right. intros [x [Hx1 Hx2]]. + assert (x = S k \/ x <= k) as [->|Hk] by lia; firstorder. + Qed. + + Lemma dec_neg_dec P: dec P -> dec (~ P). + Proof. intros [H|H]. right; eauto. now left. Qed. + + Lemma W_dec e x n : dec (W_ e x n). + Proof. + destruct (W_decider e x n) eqn: E. + left. now rewrite W_decidable. + right. intros Hw%W_decidable. congruence. + Qed. + + Fact W_dec_index n e x : dec ((W_ e ↾ n) x). + Proof. + cbn. eapply exists_bounded_dec. intro. eapply W_dec. + Qed. + + Fact ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). + Proof. + apply BaseLists.list_forall_dec. + intro x. eapply dec_neg_dec, exists_bounded_dec. + intros y. apply W_dec. + Qed. + + Definition ext_has_wit n e x := 2 * e < x /\ (W_ e ↾ n) x. + + Fact ext_has_wit_dec n e x : dec (ext_has_wit n e x). + Proof. + unfold ext_has_wit. apply and_dec. + apply Pigeonhole.dec_lt. + apply exists_bounded_dec. intro; apply W_dec. + Qed. + + Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit n e x. + + Fact dec_least_dec p n: (forall n, dec (p n)) -> dec (least p n). + Proof. + Print semi_decider. + Admitted. + + Fact least_unique p : unique (least p). + Proof. + intros n n' [H1 H2] [H1' H2']. + enough (~(n < n') /\ ~(n' < n)) by lia. + split; intros H. + - eapply H2'; eassumption. + - eapply H2; eassumption. + Qed. + + Fact magic_dec (P: nat -> nat -> Prop) (Q: nat -> Prop): + (forall y, dec (exists x, P x y)) -> + dec (exists y, Q y) -> + dec (exists y, Q y /\ exists x, P x y). + Proof. + intros H1 H2. + destruct H2 as [Hy| Hy]. + - + Admitted. + Fact ext_pick_dec L n e : dec (ext_pick L n e). + Proof. + eapply and_dec; first apply ext_intersect_W_dec. + unfold ext_has_wit. + eapply magic_dec. intro. eapply W_dec_index. + left. exists (2*e + 1). lia. + + Admitted. + + Definition ext_least_choice L n x := + exists e, e < n /\ + least (ext_pick L n) e /\ + least (ext_has_wit n e) x. + + Fact ext_least_choice_dec L n: + (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). + Proof. + unfold ext_least_choice. cbn. + Admitted. + + Fact ext_least_choice_uni l x y1 y2: + ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. + Proof. + intros [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. + assert (x0 = x1) as ->. eapply least_unique; eauto. + eapply least_unique; eauto. + Qed. + + + (* Definition extend_simple l__n n x := + exists e, mu e (fun α => e < n /\ + (l__n # W_ α n) /\ + mu x (fun β => W_ α n β /\ 2 * α < β /\ ~ In β l__n)). *) + + (* Definition extend_simple' l__n n x := + exists e, search e (fun α => (l__n # W_ α n) /\ + search x (fun beta => W_ α n beta /\ 2 * α < beta) n) n. *) + + + Variable hy1: forall (l : list nat) (x : nat), (Σ y : nat, ext_least_choice l x y) + (forall y : nat, ~ ext_least_choice l x y). + Variable hy2: forall (l : list nat) (x y1 y2 : nat), ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. + Definition simple_extendsion: Extension. + Proof. + unshelve econstructor. + - exact ext_least_choice. + - apply hy1. + - apply hy2. + Defined. + + (* Definition E_simple P__n n x := + exists e, + mu e (fun alpha => e < n /\ (W_ alpha n # P__n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). *) + +(* Definition P_ := P_templete E_simple. *) + Definition P_ := F_ simple_extendsion. + Definition Pf_ := F_func simple_extendsion. + Definition P := F_with simple_extendsion. + + Definition inifite e := ~ exhaustible (W e). + Definition incl_e L e := ~ (L # (W e)). + + Definition R_simple_list L e := inifite e -> incl_e L e. + + Definition attention e n := exists x, least (ext_pick (Pf_ n) n x) e. + Definition active e n := ~ (Pf_ n) # ((W_ e) ↾ n). + + + Lemma W_incl e n m: + n <= m -> forall x, (W_ e ↾ n) x -> (W_ e ↾ m) x. + Proof. + intros H x [y [H1 H2]]. + exists y. split; lia + easy. + Qed. + + Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : + L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. + Proof. + intros H h1 h2 x Hx1 Hx2. + eapply (H x). now eapply h1. + now apply h2. + Qed. + + Lemma active_incl e n: active e n -> forall m, n <= m -> active e m . + Proof. + intros H m Hm Hx. apply H. + eapply (intersect_mono Hx). + eapply F_mono; try eapply F_func_correctness; easy. + now eapply W_incl. + Qed. + + Lemma attention_active e k: attention e k -> active e (S k). + Proof. + intros [x H] Hcontr. + apply (Hcontr x). + assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x). + exists e. assert (Pf_ k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + apply H. assert (x = a) as ->. + eapply (@extend_uni simple_extendsion); cbn; eauto. eauto. + exfalso. eapply H3. cbn. + exists e. enough ((Pf_ k) = (Pf_ (S k))) as <- by apply H. + assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. + eapply F_uni; eauto. + firstorder. + Qed. + + Lemma active_not_attention e k: active e k -> ~ attention e k. + Proof. + now intros H [x [[Hx _] _]]. + Qed. + + Lemma active_hold e k: attention e k -> forall m, k < m -> active e m. + Proof. + intros. eapply active_incl. + apply attention_active. apply H. lia. + Qed. + + Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. + Proof. + intros H1 H2. + specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H1 a b)) as H1'. + specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H2 a b)) as H2'. + enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. + intro Hk. eapply H1'. apply Hk. easy. + intro Hk. eapply H2'. apply Hk. easy. + Qed. + + Definition spec (e x: nat) := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. + + Lemma spec_uni e x1 x2: spec e x1 -> spec e x2 -> x1 = x2 . + Proof. + intros [k [Ht Hk]] [k' [Ht' Hk']]. + assert (k=k') as <-. eapply attention_uni; eauto. + eapply (@extend_uni simple_extendsion); cbn; eauto. + Qed. + + Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, spec e x /\ e < n. + Proof. + intros [[L [k [Hin Hk]]] Hn]. + dependent induction L. inv Hin. + destruct (Dec (a = x)) as [->|]. + - clear IHL Hin. + destruct (F_pick Hk) as [k' [Hk' [e He]]]. + exists e. split. unfold spec. + exists k'. split; unfold attention. + + exists x. enough (Pf_ k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + exists e. enough (Pf_ k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + destruct He. destruct H. destruct H1. destruct H1. + lia. + - destruct (F_pop Hk) as [m' Hm']. + eapply (IHL m'); eauto. firstorder. + Qed. + + Lemma P_extract_spec n L: + (forall x, In x L -> P x /\ x <= 2 * n) -> + forall x, In x L -> exists c, spec c x /\ c < n. + Proof. + intros. induction L. inv H0. + destruct H0 as [-> | Hln]; last apply IHL; eauto. + apply P_meet_spec. now apply H. + Qed. + + Lemma DomC_pullback_list n L: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> + exists (LC: list nat), NoDup LC /\ length LC = length L /\ + forall c, In c LC -> exists x, spec c x /\ In x L /\ c < n. + Proof. + intros HL H1. + induction L. + - exists []; intuition. + - remember (@P_extract_spec n (a::L) H1 a). + assert (In a (a::L)) as H0 by intuition. + apply e in H0 as [c0 [H0 E1]]. + assert (NoDup L) by (inversion HL; intuition). + apply IHL in H as [LC H]. + exists (c0::LC). intuition. + + constructor; last now exact H2. + intros In. inv HL. + apply H4 in In as [y (Hs & h2 & h3)]. + now rewrite (spec_uni H0 Hs) in H6. + + cbn. rewrite H. trivial. + + destruct H3 as [->|]. + * exists a; intuition. + * destruct (H4 c H3) as [y Hy]. + exists y; intuition. + + intros y In1. assert (In y (a::L)) by intuition. + now apply H1 in H2. + Qed. + + Definition PredListTo p : list nat -> nat -> Prop + := fun L b => forall x, In x L <-> p x /\ x <= b. + + Lemma NoDupBoundH {L} b: + NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). + Proof. + intros ND H x E. + constructor. + - intros H1 % H. lia. + - exact ND. + Qed. + + Lemma PredNoDupListTo_NNExist p: + forall b, ~~ exists L, PredListTo p L b /\ NoDup L. + Proof. + destruct (F_computable simple_extendsion) as [f Hf]. + induction b; intros H. + - ccase (p 0) as [H0 | H0]; apply H. + + exists [0]. split; try split. + * intros [E | E]; (try contradiction E). + rewrite <- E. intuition. + * intros E. assert (x = 0) by lia. + rewrite H1. intuition. + * constructor; intuition; constructor. + + exists nil. split; try split. + * contradiction. + * intros E. assert (x = 0) by lia. + rewrite H1 in E. firstorder. + * constructor. + - apply IHb. intros [L H1]. + ccase (p (1 + b)) as [H0 | H0]; apply H. + + exists ((1+ b) :: L). split; try split. + * intros [E | E]; try (rewrite <- E; intuition). + apply H1 in E. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** right. apply H1. intuition. + ** left. lia. + * apply (@NoDupBoundH _ b). + ** apply H1. + ** intros x H3 % H1. lia. + ** lia. + + exists L. split; try split. + * intros E % H1. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** apply H1. intuition. + ** rewrite E in E1. firstorder. + * apply H1. + Qed. + + Lemma P_bounded L n: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. + Proof. + intros ND [LC H] % DomC_pullback_list; intuition. + rewrite <- H. + assert (incl LC (seq 0 n)). unfold incl. + - intros c [e [x Hx]] % H2. apply in_seq. lia. + - apply pigeonhole_length in H1. + + now rewrite seq_length in H1. + + intros. decide (x1 = x2); tauto. + + exact H0. + Qed. + + Definition R_simple P e := inifite e -> ~ (intersect (W e) P). + + Lemma R_simple_acc e L: + R_simple_list L e -> forall L', incl L L' -> R_simple_list L' e . + Proof. + intros H L' HL' h1 h2. + apply H; first easy. + firstorder. + Qed. + + (* Lemma list_accu e: + (forall k, k < e -> inifite k -> exists L, incl_e L k /\ exists n, F_ simple_extendsion n L) -> + exists m L, F_ simple_extendsion m L /\ forall n, n < e -> R_simple_list L n. + Proof. + intros. induction e. + { exists 0, []; split; first econstructor. intros n Hn. lia. } + destruct IHe as [m [HL' [Hm HL2']]]. + { intros n Hn He. apply H. lia. easy. } + pose (H e). + (* destruct (H e) as [L [HL1 HL2]]; first lia. *) + (* destruct HL2 as [k Hk]. *) + + (* destruct (dec_le m k). + + exists k, L; split; eauto. + intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. + apply HL1. eapply R_simple_acc. apply HL2'; first easy. + eapply F_mono; eauto. + + exists m, HL'; split; eauto. + intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. + eapply R_simple_acc. apply HL1. eapply F_mono; eauto with lia. + now eapply HL2'. + Qed. *) + Admitted. *) + + + (* Lemma eventually_attention e m L: + inifite e -> + (forall n : nat, n < e -> inifite n -> incl_e L n -> F_ simple_extendsion m L) -> + exists k, attention e k. + Proof. + intros H1 H2. + unfold attention. + + Admitted. *) + + + Lemma try1 e: + inifite e -> exists s, + ( ~~(Pf_ s) # (W_ e ↾ s)) \/ + (ext_intersect_W (Pf_ s) e s /\ exists x, 2 * e < x /\ (W_ e ↾ s) x). + Proof. + (* intros H HI. unfold inifite in H. *) + (* rewrite non_finite_spec in H. *) + Admitted. + + Lemma try3 e: + (forall n, n < e -> inifite n -> exists k, incl_e (Pf_ k) n) -> + inifite e -> + exists w, attention e w \/ incl_e (Pf_ w) e. + Proof. + intros He Hinfe. + destruct (try1 Hinfe) as [w [ Hx| Hw]]. + - exists w. right. unfold incl_e. + intro H. apply Hx. intro Hx'. + admit. + - + Admitted. + + Lemma try2: + forall e, inifite e -> exists m, incl_e (Pf_ m) e. + Proof. + strong induction e. intros He. + destruct (@try3 e H He) as [w [Hw| Hw]]. + specialize (attention_active Hw) as Hw'. + exists (S w). + intro Hx. unfold active in Hw'. + apply Hw'. intros y Hy1 [i [Hie Hi]]. + apply (Hx y Hy1). now exists i. + now exists w. + Qed. + + Lemma P_meet_R_simple : forall e, R_simple P e. + Proof. + intros e H3. + destruct (try2 H3) as [m Hm]. + intros Hin. apply Hm. + intros x Hx Hx'. + apply (Hin x Hx'). + unfold P. unfold F_with. + exists (Pf_ m), m. + split; last apply F_func_correctness. + easy. + Qed. + + Lemma P_semi_decidable : semi_decidable P. + Proof. + apply F_with_semi_decidable. + Qed. + + Lemma P_Listing: + forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). + Proof. + intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). + intros [L H1]. apply H. exists L; intuition. + apply P_bounded. + - exact H2. + - apply H0. + Qed. + + Lemma complToBound_compl p L b: + PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. + Proof. + intros H x. split. + - intros [H1 H1'] % in_filter_iff. + destruct Dec; cbn in H1'; try congruence. + enough (x <= b). + + intuition. + + apply in_seq in H1. lia. + - intros [H1 H2]. eapply in_filter_iff. split. + + apply in_seq; lia. + + destruct Dec; cbn; try tauto. exfalso. firstorder. + Qed. + + Lemma Compl_P_Listing: + forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L + /\ forall x, In x L -> ~ P x. + Proof. + intros n H. + apply (@P_Listing n). intros [L H1]. + apply H. exists (complToBound L (2*n)). repeat split. + - remember (complToBound_length L (2*n)). lia. + - apply complToBound_NoDup. + - intros x I % (@complToBound_compl P); intuition. + Qed. + + Lemma P_coinfinite : ~ exhaustible (compl P). + Proof. + eapply weakly_unbounded_non_finite. + intros n H. eapply Compl_P_Listing with (n := n). + intros (l & ? & ? & H2). + eapply H. + exists (firstn n l). + repeat split. + - rewrite firstn_length. lia. + - now eapply firstn_NoDup. + - intros ? ? % firstn_In. now eapply H2. + Qed. + + Lemma P_simple : simple P. + Proof. + unfold simple; repeat split. + - rewrite EA.enum_iff. now apply P_semi_decidable. + - apply P_coinfinite. + - intros [q (Hqenum & Hqinf & Hq)]. + rewrite EA.enum_iff in Hqenum. + destruct (W_spec Hqenum) as [c HWq]. + apply (@P_meet_R_simple c). + intros [l Hqe]; apply Hqinf. + exists l. intros x Hqx. apply (Hqe x). + now rewrite HWq in Hqx. + intros x HWcx HPx. unfold W in HWcx. + rewrite <- HWq in HWcx. apply (Hq x HWcx HPx). + Qed. + + End Assume_EA. \ No newline at end of file From bfa7cf6e98f71fab690ca411e225600e437582d7 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sun, 17 Mar 2024 20:35:51 +0100 Subject: [PATCH 14/54] decidability of extension --- theories/ReducibilityDegrees/lowsimple.v | 4 +- .../ReducibilityDegrees/priority_method.v | 4 - .../ReducibilityDegrees/simple_extension.v | 1126 +++++++++++------ 3 files changed, 709 insertions(+), 425 deletions(-) diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index 00c638a..a6ecd79 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -60,4 +60,6 @@ Proof. split; [destruct H2 as [H2 _]; eauto| now apply lowness]. Qed. -End LowSimplePredicate. \ No newline at end of file +End LowSimplePredicate. + +(*** Instance of low simple predicate ***) \ No newline at end of file diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index cd0fe45..aaf5699 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -12,10 +12,6 @@ Notation "'Σ' x .. y , p" := format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") : type_scope. -Definition disj {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. -Definition intersect {X} (A B: X -> Prop) := forall x, A x -> B x -> False. -Notation "A # B" := (disj A B) (at level 30). - Section Construction. Record Extension := { diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index c08f0cc..08fadc9 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -66,368 +66,733 @@ Section ComplToBound. Qed. End ComplToBound. -Notation unique p := (forall x x', p x -> p x' -> x = x'). -Section LeastWitness. - Definition safe p n := forall k, k < n -> ~ p k. - Definition least p n := p n /\ safe p n. -End LeastWitness. - +Section EWO. + Variable p: nat -> Prop. + Inductive T | (n: nat) : Prop := C (phi: ~p n -> T (S n)). -Section Assume_EA. + Definition T_elim (q: nat -> Type) + : (forall n, (~p n -> q (S n)) -> q n) -> + forall n, T n -> q n + := fun e => fix f n a := + let (phi) := a in + e n (fun h => f (S n) (phi h)). -Section Requirements. + (*** EWO for Numbers *) - Definition W e x := exists n, W_ e n x. - Notation " P ↾ s" := (fun x => exists n, n <= s /\ P n x) (at level 20). - - Definition ext_intersect_W L n e := L # (W_ e) ↾ n. - -End Requirements. - - Lemma exists_bounded_dec' P: - (forall x, dec (P x)) -> forall k, dec (exists n, n < k /\ P n). + Lemma TI n : + p n -> T n. Proof. - intros Hp k. - induction k. right; intros [? [? _]]; lia. - destruct IHk as [Hw|Hw]. - - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. - - destruct (Hp k). left. exists k; split; eauto; lia. - right. intros [x [Hx1 Hx2]]. - assert (x = k \/ x < k) as [->|Hk] by lia; firstorder. + intros H. constructor. intros H1. destruct (H1 H). Qed. - Lemma exists_bounded_dec P: - (forall x, dec (P x)) -> forall k, dec (exists n, n <= k /\ P n). + Lemma TD n : + T (S n) -> T n. Proof. - intros Hp k. - induction k. destruct (Hp 0). left. exists 0. eauto. - right. intros [x [Hx Hx']]. now assert (x=0) as -> by lia. - destruct IHk as [Hw|Hw]. - - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. - - destruct (Hp (S k)). left. exists (S k); split; eauto; lia. - right. intros [x [Hx1 Hx2]]. - assert (x = S k \/ x <= k) as [->|Hk] by lia; firstorder. + intros H. constructor. intros _. exact H. Qed. - Lemma dec_neg_dec P: dec P -> dec (~ P). - Proof. intros [H|H]. right; eauto. now left. Qed. + Variable p_dec: forall n, dec (p n). - Lemma W_dec e x n : dec (W_ e x n). + Definition TE (q: nat -> Type) + : (forall n, p n -> q n) -> + (forall n, ~p n -> q (S n) -> q n) -> + forall n, T n -> q n. Proof. - destruct (W_decider e x n) eqn: E. - left. now rewrite W_decidable. - right. intros Hw%W_decidable. congruence. + intros e1 e2. + apply T_elim. intros n IH. + destruct (p_dec n); auto. Qed. - Fact W_dec_index n e x : dec ((W_ e ↾ n) x). + (** From now on T will only be used through TI, TD, and TE *) + + + Lemma T_zero n : + T n -> T 0. Proof. - cbn. eapply exists_bounded_dec. intro. eapply W_dec. + induction n as [|n IH]. + - auto. + - intros H. apply IH. apply TD, H. Qed. - Fact ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). + Definition ewo_nat : + ex p -> Σ x, p x. Proof. - apply BaseLists.list_forall_dec. - intro x. eapply dec_neg_dec, exists_bounded_dec. - intros y. apply W_dec. - Qed. + intros H. + refine (@TE (fun _ => Σ x, p x) _ _ 0 _). + - eauto. + - easy. + - destruct H as [n H]. apply (@T_zero n), TI, H. + Qed. - Definition ext_has_wit n e x := 2 * e < x /\ (W_ e ↾ n) x. +End EWO. - Fact ext_has_wit_dec n e x : dec (ext_has_wit n e x). - Proof. - unfold ext_has_wit. apply and_dec. - apply Pigeonhole.dec_lt. - apply exists_bounded_dec. intro; apply W_dec. - Qed. +Notation unique p := (forall x x', p x -> p x' -> x = x'). +Section LeastWitness. + Definition safe p n := forall k, k < n -> ~ p k. + Definition least p n := p n /\ safe p n. - Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit n e x. + Fact least_unique p : unique (least p). + Proof. + intros n n' [H1 H2] [H1' H2']. + enough (~(n < n') /\ ~(n' < n)) by lia. + split; intros H. + - eapply H2'; eassumption. + - eapply H2; eassumption. + Qed. - Fact dec_least_dec p n: (forall n, dec (p n)) -> dec (least p n). - Proof. - Print semi_decider. - Admitted. + Fact safe_O p : + safe p 0. + Proof. + hnf. lia. + Qed. - Fact least_unique p : unique (least p). - Proof. - intros n n' [H1 H2] [H1' H2']. - enough (~(n < n') /\ ~(n' < n)) by lia. - split; intros H. - - eapply H2'; eassumption. - - eapply H2; eassumption. - Qed. + Fact safe_S p n : + safe p n -> ~p n -> safe p (S n). + Proof. + intros H1 H2 k H3. unfold safe in *. + assert (k < n \/ k = n) as [H|H] by lia. + - auto. + - congruence. + Qed. - Fact magic_dec (P: nat -> nat -> Prop) (Q: nat -> Prop): - (forall y, dec (exists x, P x y)) -> - dec (exists y, Q y) -> - dec (exists y, Q y /\ exists x, P x y). - Proof. - intros H1 H2. - destruct H2 as [Hy| Hy]. - - - Admitted. - Fact ext_pick_dec L n e : dec (ext_pick L n e). - Proof. - eapply and_dec; first apply ext_intersect_W_dec. - unfold ext_has_wit. - eapply magic_dec. intro. eapply W_dec_index. - left. exists (2*e + 1). lia. - - Admitted. + Fact safe_char p n : + safe p n <-> forall k, p k -> k >= n. + Proof. + split. + - intros H k H1. + enough (k < n -> False) by lia. + intros H2. apply H in H2. auto. + - intros H k H1 H2. apply H in H2. lia. + Qed. - Definition ext_least_choice L n x := - exists e, e < n /\ - least (ext_pick L n) e /\ - least (ext_has_wit n e) x. + Fact safe_char_S p n : + safe p (S n) <-> safe p n /\ ~p n. + Proof. + split. + - intros H. split. + + intros k H1. apply H. lia. + + apply H. lia. + - intros [H1 H2]. apply safe_S; assumption. + Qed. - Fact ext_least_choice_dec L n: - (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). - Proof. - unfold ext_least_choice. cbn. - Admitted. + Fact safe_eq p n k : + safe p n -> k <= n -> p k -> k = n. + Proof. + intros H1 H2 H3. unfold safe in *. + enough (~(k < n)) by lia. + specialize (H1 k). tauto. + Qed. - Fact ext_least_choice_uni l x y1 y2: - ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. - Proof. - intros [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. - assert (x0 = x1) as ->. eapply least_unique; eauto. - eapply least_unique; eauto. - Qed. - + Fact E14 x y z : + x - y = z <-> least (fun k => x <= y + k) z. + Proof. + assert (H: least (fun k => x <= y + k) (x - y)). + { split; unfold safe; lia. } + split. congruence. + eapply least_unique, H. + Qed. - (* Definition extend_simple l__n n x := - exists e, mu e (fun α => e < n /\ - (l__n # W_ α n) /\ - mu x (fun β => W_ α n β /\ 2 * α < β /\ ~ In β l__n)). *) + (*** Certifying LWOs *) - (* Definition extend_simple' l__n n x := - exists e, search e (fun α => (l__n # W_ α n) /\ - search x (fun beta => W_ α n beta /\ 2 * α < beta) n) n. *) + Section LWO. + Variable p : nat -> Prop. + Variable p_dec : forall x, dec (p x). + Definition lwo : + forall n, (Σ k, k < n /\ least p k) + safe p n. + Proof. + induction n as [|n IH]. + - right. apply safe_O. + - destruct IH as [[k H1]|H1]. + + left. exists k. destruct H1 as [H1 H2]. split. lia. exact H2. + + destruct (p_dec n). + * left. exists n. split. lia. easy. + * right. apply safe_S; assumption. + Defined. + + Definition lwo' : + forall n, (Σ k, k <= n /\ least p k) + safe p (S n). + Proof. + intros n. + destruct (lwo (S n)) as [(k&H1&H2)|H]. + - left. exists k. split. lia. exact H2. + - right. exact H. + Qed. - Variable hy1: forall (l : list nat) (x : nat), (Σ y : nat, ext_least_choice l x y) + (forall y : nat, ~ ext_least_choice l x y). - Variable hy2: forall (l : list nat) (x y1 y2 : nat), ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. - Definition simple_extendsion: Extension. - Proof. - unshelve econstructor. - - exact ext_least_choice. - - apply hy1. - - apply hy2. - Defined. + Definition least_sig : + (Σ x, p x) -> Σ x, (least p) x. + Proof. + intros [n H]. + destruct (lwo (S n)) as [(k&H1&H2)|H1]. + - exists k. exact H2. + - exfalso. apply (H1 n). lia. exact H. + Qed. - (* Definition E_simple P__n n x := - exists e, - mu e (fun alpha => e < n /\ (W_ alpha n # P__n) /\ mu x (fun beta => W_ alpha n beta /\ 2 * alpha < beta)). *) + Definition least_ex : + ex p -> ex (least p). + Proof. + intros [n H]. + destruct (lwo (S n)) as [(k&H1&H2)|H1]. + - exists k. exact H2. + - exfalso. apply (H1 n). lia. exact H. + Qed. -(* Definition P_ := P_templete E_simple. *) - Definition P_ := F_ simple_extendsion. - Definition Pf_ := F_func simple_extendsion. - Definition P := F_with simple_extendsion. + Definition safe_dec n : + dec (safe p n). + Proof. + destruct (lwo n) as [[k H1]|H1]. + - right. intros H. exfalso. + destruct H1 as [H1 H2]. apply (H k). exact H1. apply H2. + - left. exact H1. + Defined. + + Definition least_dec n : + dec (least p n). + Proof. + unfold least. + destruct (p_dec n) as [H|H]. + 2:{ right. tauto. } + destruct (safe_dec n) as [H1|H1]. + - left. easy. + - right. tauto. + Qed. + End LWO. - Definition inifite e := ~ exhaustible (W e). - Definition incl_e L e := ~ (L # (W e)). + Lemma exists_bounded_dec' P: + (forall x, dec (P x)) -> forall k, dec (exists n, n < k /\ P n). + Proof. + intros Hp k. + induction k. right; intros [? [? _]]; lia. + destruct IHk as [Hw|Hw]. + - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. + - destruct (Hp k). left. exists k; split; eauto; lia. + right. intros [x [Hx1 Hx2]]. + assert (x = k \/ x < k) as [->|Hk] by lia; firstorder. + Qed. - Definition R_simple_list L e := inifite e -> incl_e L e. + Lemma exists_bounded_dec P: + (forall x, dec (P x)) -> forall k, dec (exists n, n <= k /\ P n). + Proof. + intros Hp k. + induction k. destruct (Hp 0). left. exists 0. eauto. + right. intros [x [Hx Hx']]. now assert (x=0) as -> by lia. + destruct IHk as [Hw|Hw]. + - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. + - destruct (Hp (S k)). left. exists (S k); split; eauto; lia. + right. intros [x [Hx1 Hx2]]. + assert (x = S k \/ x <= k) as [->|Hk] by lia; firstorder. + Qed. - Definition attention e n := exists x, least (ext_pick (Pf_ n) n x) e. - Definition active e n := ~ (Pf_ n) # ((W_ e) ↾ n). + Definition bounded (P: nat -> Prop) := Σ N, forall x, P x -> x < N. + Fact bouned_le (P Q: nat -> Prop) N : + (forall x, P x -> x < N) -> + (exists x, P x /\ Q x) <-> exists x, x < N /\ P x /\ Q x. + Proof. + intros Hn; split. + - intros [x Hx]. exists x; intuition eauto. + - intros (x&c&Hc). exists x; intuition eauto. + Qed. - Lemma W_incl e n m: - n <= m -> forall x, (W_ e ↾ n) x -> (W_ e ↾ m) x. - Proof. - intros H x [y [H1 H2]]. - exists y. split; lia + easy. - Qed. + Lemma bounded_dec (P Q: nat -> Prop): + (forall x, dec (P x)) -> (forall x, dec (Q x)) -> + bounded P -> dec (exists n, P n /\ Q n). + Proof. + intros H1 H2 [N HN]. + assert (dec (exists n, n < N /\ P n /\ Q n)) as [H|H]. + - eapply exists_bounded_dec'. intro; eapply and_dec; eauto. + - left. rewrite bouned_le; eauto. + - right. intros H'%(@bouned_le _ _ N); easy. + Qed. + Lemma dec_neg_dec P: dec P -> dec (~ P). + Proof. intros [H|H]. right; eauto. now left. Qed. - Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : - L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. - Proof. - intros H h1 h2 x Hx1 Hx2. - eapply (H x). now eapply h1. - now apply h2. - Qed. +End LeastWitness. - Lemma active_incl e n: active e n -> forall m, n <= m -> active e m . - Proof. - intros H m Hm Hx. apply H. - eapply (intersect_mono Hx). - eapply F_mono; try eapply F_func_correctness; easy. - now eapply W_incl. - Qed. +Section Assume_EA. - Lemma attention_active e k: attention e k -> active e (S k). - Proof. - intros [x H] Hcontr. - apply (Hcontr x). - assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. - inv H0. cbn in H4. assert (ext_least_choice l k x). - exists e. assert (Pf_ k = l) as <-. - eapply F_uni. apply F_func_correctness. apply H3. - apply H. assert (x = a) as ->. - eapply (@extend_uni simple_extendsion); cbn; eauto. eauto. - exfalso. eapply H3. cbn. - exists e. enough ((Pf_ k) = (Pf_ (S k))) as <- by apply H. - assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. - eapply F_uni; eauto. - firstorder. - Qed. + Variable φ: nat -> nat -> option nat. + Definition EA := forall P, + semi_decidable P -> exists e, forall x, P x <-> exists n, φ e n = Some x. + Hypothesis EA: EA. - Lemma active_not_attention e k: active e k -> ~ attention e k. - Proof. - now intros H [x [[Hx _] _]]. - Qed. + Definition W_ n e x := φ n e = Some x. + Definition W e x := exists n, W_ e n x. - Lemma active_hold e k: attention e k -> forall m, k < m -> active e m. - Proof. - intros. eapply active_incl. - apply attention_active. apply H. lia. - Qed. + Lemma W_spec: forall P, semi_decidable P -> exists e, forall x, P x <-> W e x. + Proof. intros P [e He]%EA. exists e; intros x; now rewrite He. Qed. - Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. - Proof. - intros H1 H2. - specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H1 a b)) as H1'. - specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H2 a b)) as H2'. - enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. - intro Hk. eapply H1'. apply Hk. easy. - intro Hk. eapply H2'. apply Hk. easy. - Qed. + Notation "'W[' s ']' e" := (fun x => exists n, n <= s /\ W_ e n x) (at level 30). - Definition spec (e x: nat) := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. + Section EA_dec. - Lemma spec_uni e x1 x2: spec e x1 -> spec e x2 -> x1 = x2 . - Proof. - intros [k [Ht Hk]] [k' [Ht' Hk']]. - assert (k=k') as <-. eapply attention_uni; eauto. - eapply (@extend_uni simple_extendsion); cbn; eauto. - Qed. + Lemma W_dec e: forall x n, dec (W_ e n x). + Proof. + intros x n. + destruct (φ e n) eqn: E. + - destruct (Dec (x = n0)) as [E'|E']. + left. now subst. right. intros H. congruence. + - right. intros H. congruence. + Qed. - Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, spec e x /\ e < n. - Proof. - intros [[L [k [Hin Hk]]] Hn]. - dependent induction L. inv Hin. - destruct (Dec (a = x)) as [->|]. - - clear IHL Hin. - destruct (F_pick Hk) as [k' [Hk' [e He]]]. - exists e. split. unfold spec. - exists k'. split; unfold attention. - + exists x. enough (Pf_ k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + exists e. enough (Pf_ k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + destruct He. destruct H. destruct H1. destruct H1. - lia. - - destruct (F_pop Hk) as [m' Hm']. - eapply (IHL m'); eauto. firstorder. - Qed. + Lemma W_bounded_dec e : forall x s, dec ((W[s] e) x). + Proof. + intros x s. cbn. eapply exists_bounded_dec. + intro; apply W_dec. + Qed. - Lemma P_extract_spec n L: - (forall x, In x L -> P x /\ x <= 2 * n) -> - forall x, In x L -> exists c, spec c x /\ c < n. - Proof. - intros. induction L. inv H0. - destruct H0 as [-> | Hln]; last apply IHL; eauto. - apply P_meet_spec. now apply H. - Qed. + Lemma W_bounded_bounded e s: bounded (W[s] e). + Proof. + unfold bounded. + induction s. + - destruct (φ e 0) as [w|] eqn: E. + exists (S w). intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. rewrite E in Hn2. + injection Hn2. lia. + exists 42. intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. + rewrite E in Hn2. congruence. + - destruct IHs as [N HN]. + destruct (φ e (S s)) as [w|] eqn: E. + exists (max (S w) (S N)). intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. + rewrite E in Hn2. injection Hn2. lia. + eapply Nat.lt_trans. eapply HN. exists n; split; easy. lia. + exists N. intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. + rewrite E in Hn2. congruence. + eapply HN. exists n. split; eauto. + Qed. - Lemma DomC_pullback_list n L: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> - exists (LC: list nat), NoDup LC /\ length LC = length L /\ - forall c, In c LC -> exists x, spec c x /\ In x L /\ c < n. - Proof. - intros HL H1. - induction L. - - exists []; intuition. - - remember (@P_extract_spec n (a::L) H1 a). - assert (In a (a::L)) as H0 by intuition. - apply e in H0 as [c0 [H0 E1]]. - assert (NoDup L) by (inversion HL; intuition). - apply IHL in H as [LC H]. - exists (c0::LC). intuition. - + constructor; last now exact H2. - intros In. inv HL. - apply H4 in In as [y (Hs & h2 & h3)]. - now rewrite (spec_uni H0 Hs) in H6. - + cbn. rewrite H. trivial. - + destruct H3 as [->|]. - * exists a; intuition. - * destruct (H4 c H3) as [y Hy]. - exists y; intuition. - + intros y In1. assert (In y (a::L)) by intuition. - now apply H1 in H2. - Qed. + End EA_dec. - Definition PredListTo p : list nat -> nat -> Prop - := fun L b => forall x, In x L <-> p x /\ x <= b. + Definition disj {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. + Definition intersect {X} (A B: X -> Prop) := forall x, A x -> B x -> False. + Notation "A # B" := (disj A B) (at level 30). - Lemma NoDupBoundH {L} b: - NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). - Proof. - intros ND H x E. - constructor. - - intros H1 % H. lia. - - exact ND. - Qed. + Section Extension. - Lemma PredNoDupListTo_NNExist p: - forall b, ~~ exists L, PredListTo p L b /\ NoDup L. - Proof. - destruct (F_computable simple_extendsion) as [f Hf]. - induction b; intros H. - - ccase (p 0) as [H0 | H0]; apply H. - + exists [0]. split; try split. - * intros [E | E]; (try contradiction E). - rewrite <- E. intuition. - * intros E. assert (x = 0) by lia. - rewrite H1. intuition. - * constructor; intuition; constructor. - + exists nil. split; try split. - * contradiction. - * intros E. assert (x = 0) by lia. - rewrite H1 in E. firstorder. - * constructor. - - apply IHb. intros [L H1]. - ccase (p (1 + b)) as [H0 | H0]; apply H. - + exists ((1+ b) :: L). split; try split. - * intros [E | E]; try (rewrite <- E; intuition). - apply H1 in E. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** right. apply H1. intuition. - ** left. lia. - * apply (@NoDupBoundH _ b). - ** apply H1. - ** intros x H3 % H1. lia. - ** lia. - + exists L. split; try split. - * intros E % H1. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** apply H1. intuition. - ** rewrite E in E1. firstorder. - * apply H1. - Qed. + Definition ext_intersect_W L n e := L # W[n] e. + Definition ext_has_wit n e x := (W[n] e) x /\ 2 * e < x. + Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit n e x. + Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit n e) x. + Definition ext_least_choice L n x := exists e, ext_choice L n e x. - Lemma P_bounded L n: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. - Proof. - intros ND [LC H] % DomC_pullback_list; intuition. - rewrite <- H. - assert (incl LC (seq 0 n)). unfold incl. - - intros c [e [x Hx]] % H2. apply in_seq. lia. - - apply pigeonhole_length in H1. - + now rewrite seq_length in H1. - + intros. decide (x1 = x2); tauto. - + exact H0. - Qed. + End Extension. - Definition R_simple P e := inifite e -> ~ (intersect (W e) P). + Section Extension_Facts. - Lemma R_simple_acc e L: - R_simple_list L e -> forall L', incl L L' -> R_simple_list L' e . - Proof. - intros H L' HL' h1 h2. - apply H; first easy. - firstorder. - Qed. + #[export]Instance ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). + Proof. + apply BaseLists.list_forall_dec. + intro x. eapply dec_neg_dec, exists_bounded_dec. + intros y. apply W_dec. + Qed. + + #[export]Instance ext_has_wit_dec n e x : dec (ext_has_wit n e x). + Proof. + unfold ext_has_wit. apply and_dec. + apply exists_bounded_dec. intro; apply W_dec. + apply Pigeonhole.dec_lt. + Qed. + + #[export]Instance ext_has_wit_exists_dec n e : dec (exists x, ext_has_wit n e x). + Proof. + unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. + intro x; eapply W_bounded_dec. + intro x; eapply lt_dec. + Qed. + + #[export]Instance ext_pick_dec L n e : dec (ext_pick L n e). + Proof. + eapply and_dec; first apply ext_intersect_W_dec. + unfold ext_has_wit. + eapply bounded_dec; last apply W_bounded_bounded. + intros x. eapply exists_bounded_dec. + intro; apply W_dec. + apply lt_dec. + Qed. + + #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). + Proof. + eapply exists_bounded_dec'. intro x. + eapply least_dec. intros y. + eapply ext_pick_dec. + Qed. + + Fact ext_least_choice_dec L n: + (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). + Proof. + unfold ext_least_choice. + destruct (ext_pick_exists_dec L n) as [H|H]. + - left. apply ewo_nat; first last. + destruct H as [e (h1 & [(h4 & h2) h3])]. + eapply least_ex in h2. destruct h2 as [k h6]. + exists k, e. split; first easy; split. + 2: { eapply h6. } + split; last apply h3. split; first apply h4. + destruct h6. now exists k. + eapply ext_has_wit_dec. + intro x. unfold ext_choice. eapply exists_bounded_dec'. + intros x'. eapply and_dec. + apply least_dec. eapply ext_pick_dec. + apply least_dec. eapply ext_has_wit_dec. + - right. intros x [e (h1 & h2 & _)]. apply H. + exists e. split; eauto. + Qed. + + Fact ext_least_choice_uni l x y1 y2: + ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. + Proof. + intros [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. + assert (x0 = x1) as ->. eapply least_unique; eauto. + eapply least_unique; eauto. + Qed. + + End Extension_Facts. + + Section Simple_Extension. + + Definition simple_extendsion: Extension. + Proof. + unshelve econstructor. + - exact ext_least_choice. + - apply ext_least_choice_dec. + - apply ext_least_choice_uni. + Defined. + + Definition P_ := F_ simple_extendsion. + Definition Pf_ := F_func simple_extendsion. + Definition P := F_with simple_extendsion. + + Definition non_finite e := ~ exhaustible (W e). + Definition incl_e L e := ~ (L # (W e)). + + End Simple_Extension. + + Section Requirements. + Definition R_simple_list L e := non_finite e -> incl_e L e. + + Definition attention e n := e < n /\ least (ext_pick (Pf_ n) n) e. + Definition active e n := ~ (Pf_ n) # W[n] e. + Definition pick_el (e x: nat) := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. + + End Requirements. + + Section Requirements_Facts. + + Lemma ext_pick_witness L n e: + ext_pick L n e -> exists x, least (ext_has_wit n e) x. + Proof. + intros [H1 H2]. + eapply least_ex. intro; eapply ext_has_wit_dec. + trivial. + Qed. + + Lemma W_incl e n m: + n <= m -> forall x, (W[n] e) x -> (W[m] e) x. + Proof. + intros H x [y [H1 H2]]. + exists y. split; lia + easy. + Qed. + + Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : + L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. + Proof. + intros H h1 h2 x Hx1 Hx2. + eapply (H x). now eapply h1. + now apply h2. + Qed. + + Lemma active_incl e n: active e n -> forall m, n <= m -> active e m . + Proof. + intros H m Hm Hx. apply H. + eapply (intersect_mono Hx). + eapply F_mono; try eapply F_func_correctness; easy. + now eapply W_incl. + Qed. + + Lemma attention_active e k: attention e k -> active e (S k). + Proof. + intros [He H] Hcontr. + edestruct (ext_pick_witness) as [x Hx]. + { destruct H. eapply e0. } + apply (Hcontr x). + assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. + exists e. assert (Pf_ k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + split; first easy. split; first easy. easy. + assert (x = a) as ->. eapply (@extend_uni simple_extendsion); cbn; eauto. + eauto. exfalso. eapply (H3 x). cbn. + exists e. enough ((Pf_ k) = (Pf_ (S k))) as <-. + split; first easy. easy. + assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. + eapply F_uni; eauto. + firstorder. + Qed. + + Lemma active_not_attention e k: active e k -> ~ attention e k. + Proof. now intros h2 [_ [[h _] _]]. Qed. + + Lemma active_hold e k: attention e k -> forall m, k < m -> active e m. + Proof. + intros. eapply active_incl. + apply attention_active. apply H. lia. + Qed. + + Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. + Proof. + intros H1 H2. + specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H1 a b)) as H1'. + specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H2 a b)) as H2'. + enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. + intro Hk. eapply H1'. apply Hk. easy. + intro Hk. eapply H2'. apply Hk. easy. + Qed. + + Lemma spec_uni e x1 x2: pick_el e x1 -> pick_el e x2 -> x1 = x2 . + Proof. + intros [k [Ht Hk]] [k' [Ht' Hk']]. + assert (k=k') as <-. eapply attention_uni; eauto. + eapply (@extend_uni simple_extendsion); cbn; eauto. + Qed. + + End Requirements_Facts. + + Section Compl_P_non_finite. + + Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, pick_el e x /\ e < n. + Proof. + intros [[L [k [Hin Hk]]] Hn]. + dependent induction L. inv Hin. + destruct (Dec (a = x)) as [->|]. + - clear IHL Hin. + destruct (F_pick Hk) as [k' [Hk' [e He]]]. + exists e. split. + exists k'. split; unfold attention. + + destruct He; intuition eauto. enough (Pf_ k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + exists e. unfold ext_choice. destruct He; intuition eauto. + enough (Pf_ k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + destruct He. destruct H0. destruct H1. destruct H1. + lia. + - destruct (F_pop Hk) as [m' Hm']. + eapply (IHL m'); eauto. firstorder. + Qed. + + Lemma P_extract_spec n L: + (forall x, In x L -> P x /\ x <= 2 * n) -> + forall x, In x L -> exists c, pick_el c x /\ c < n. + Proof. + intros. induction L. inv H0. + destruct H0 as [-> | Hln]; last apply IHL; eauto. + apply P_meet_spec. now apply H. + Qed. + + Lemma P_pullback_list n L: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> + exists (LC: list nat), NoDup LC /\ length LC = length L /\ + forall c, In c LC -> exists x, pick_el c x /\ In x L /\ c < n. + Proof. + intros HL H1. + induction L. + - exists []; intuition. + - remember (@P_extract_spec n (a::L) H1 a). + assert (In a (a::L)) as H0 by intuition. + apply e in H0 as [c0 [H0 E1]]. + assert (NoDup L) by (inversion HL; intuition). + apply IHL in H as [LC H]. + exists (c0::LC). intuition. + + constructor; last now exact H2. + intros In. inv HL. + apply H4 in In as [y (Hs & h2 & h3)]. + now rewrite (spec_uni H0 Hs) in H6. + + cbn. rewrite H. trivial. + + destruct H3 as [->|]. + * exists a; intuition. + * destruct (H4 c H3) as [y Hy]. + exists y; intuition. + + intros y In1. assert (In y (a::L)) by intuition. + now apply H1 in H2. + Qed. + + Definition PredListTo p : list nat -> nat -> Prop + := fun L b => forall x, In x L <-> p x /\ x <= b. + + Lemma NoDupBoundH {L} b: + NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). + Proof. + intros ND H x E. + constructor. + - intros H1 % H. lia. + - exact ND. + Qed. + + Lemma PredNoDupListTo_NNExist p: + forall b, ~~ exists L, PredListTo p L b /\ NoDup L. + Proof. + destruct (F_computable simple_extendsion) as [f Hf]. + induction b; intros H. + - ccase (p 0) as [H0 | H0]; apply H. + + exists [0]. split; try split. + * intros [E | E]; (try contradiction E). + rewrite <- E. intuition. + * intros E. assert (x = 0) by lia. + rewrite H1. intuition. + * constructor; intuition; constructor. + + exists nil. split; try split. + * contradiction. + * intros E. assert (x = 0) by lia. + rewrite H1 in E. firstorder. + * constructor. + - apply IHb. intros [L H1]. + ccase (p (1 + b)) as [H0 | H0]; apply H. + + exists ((1+ b) :: L). split; try split. + * intros [E | E]; try (rewrite <- E; intuition). + apply H1 in E. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** right. apply H1. intuition. + ** left. lia. + * apply (@NoDupBoundH _ b). + ** apply H1. + ** intros x H3 % H1. lia. + ** lia. + + exists L. split; try split. + * intros E % H1. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** apply H1. intuition. + ** rewrite E in E1. firstorder. + * apply H1. + Qed. + + Lemma P_bounded L n: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. + Proof. + intros ND [LC H] % P_pullback_list; intuition. + rewrite <- H. + assert (incl LC (seq 0 n)). unfold incl. + - intros c [e [x Hx]] % H2. apply in_seq. lia. + - apply pigeonhole_length in H1. + + now rewrite seq_length in H1. + + intros. decide (x1 = x2); tauto. + + exact H0. + Qed. + + Lemma P_Listing: + forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). + Proof. + intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). + intros [L H1]. apply H. exists L; intuition. + apply P_bounded. + - exact H2. + - apply H0. + Qed. + + Lemma complToBound_compl p L b: + PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. + Proof. + intros H x. split. + - intros [H1 H1'] % in_filter_iff. + destruct Dec; cbn in H1'; try congruence. + enough (x <= b). + + intuition. + + apply in_seq in H1. lia. + - intros [H1 H2]. eapply in_filter_iff. split. + + apply in_seq; lia. + + destruct Dec; cbn; try tauto. exfalso. firstorder. + Qed. + + Lemma Compl_P_Listing: + forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L + /\ forall x, In x L -> ~ P x. + Proof. + intros n H. + apply (@P_Listing n). intros [L H1]. + apply H. exists (complToBound L (2*n)). repeat split. + - remember (complToBound_length L (2*n)). lia. + - apply complToBound_NoDup. + - intros x I % (@complToBound_compl P); intuition. + Qed. + + Lemma P_coinfinite : ~ exhaustible (compl P). + Proof. + eapply weakly_unbounded_non_finite. + intros n H. eapply Compl_P_Listing with (n := n). + intros (l & ? & ? & H2). + eapply H. + exists (firstn n l). + repeat split. + - rewrite firstn_length. lia. + - now eapply firstn_NoDup. + - intros ? ? % firstn_In. now eapply H2. + Qed. + + End Compl_P_non_finite. + + Section Meet_Requirement. + + Definition R_simple P e := non_finite e -> ~ (intersect (W e) P). + + Lemma R_simple_acc e L: + R_simple_list L e -> forall L', incl L L' -> R_simple_list L' e . + Proof. + intros H L' HL' h1 h2. + apply H; first easy. + firstorder. + Qed. + + + Lemma try1 e: + non_finite e -> exists s, + ( ~~(Pf_ s) # (W[s] e)) \/ + (ext_intersect_W (Pf_ s) e s /\ exists x, 2 * e < x /\ (W[s] e) x). + Proof. + (* intros H HI. unfold inifite in H. *) + (* rewrite non_finite_spec in H. *) + Admitted. + + Lemma try3 e: + (forall n, n < e -> non_finite n -> exists k, incl_e (Pf_ k) n) -> + non_finite e -> + exists w, attention e w \/ incl_e (Pf_ w) e. + Proof. + intros He Hinfe. + destruct (try1 Hinfe) as [w [ Hx| Hw]]. + - exists w. right. unfold incl_e. + intro H. apply Hx. intro Hx'. + admit. + - + Admitted. + + Lemma try2: + forall e, non_finite e -> exists m, incl_e (Pf_ m) e. + Proof. + strong induction e. intros He. + destruct (@try3 e H He) as [w [Hw| Hw]]. + specialize (attention_active Hw) as Hw'. + exists (S w). + intro Hx. unfold active in Hw'. + apply Hw'. intros y Hy1 [i [Hie Hi]]. + apply (Hx y Hy1). now exists i. + now exists w. + Qed. + + Lemma P_meet_R_simple : forall e, R_simple P e. + Proof. + intros e H3. + destruct (try2 H3) as [m Hm]. + intros Hin. apply Hm. + intros x Hx Hx'. + apply (Hin x Hx'). + unfold P. unfold F_with. + exists (Pf_ m), m. + split; last apply F_func_correctness. + easy. + Qed. + + End Meet_Requirement. (* Lemma list_accu e: (forall k, k < e -> inifite k -> exists L, incl_e L k /\ exists n, F_ simple_extendsion n L) -> @@ -465,122 +830,43 @@ End Requirements. Admitted. *) - Lemma try1 e: - inifite e -> exists s, - ( ~~(Pf_ s) # (W_ e ↾ s)) \/ - (ext_intersect_W (Pf_ s) e s /\ exists x, 2 * e < x /\ (W_ e ↾ s) x). - Proof. - (* intros H HI. unfold inifite in H. *) - (* rewrite non_finite_spec in H. *) - Admitted. - - Lemma try3 e: - (forall n, n < e -> inifite n -> exists k, incl_e (Pf_ k) n) -> - inifite e -> - exists w, attention e w \/ incl_e (Pf_ w) e. - Proof. - intros He Hinfe. - destruct (try1 Hinfe) as [w [ Hx| Hw]]. - - exists w. right. unfold incl_e. - intro H. apply Hx. intro Hx'. - admit. - - - Admitted. - - Lemma try2: - forall e, inifite e -> exists m, incl_e (Pf_ m) e. - Proof. - strong induction e. intros He. - destruct (@try3 e H He) as [w [Hw| Hw]]. - specialize (attention_active Hw) as Hw'. - exists (S w). - intro Hx. unfold active in Hw'. - apply Hw'. intros y Hy1 [i [Hie Hi]]. - apply (Hx y Hy1). now exists i. - now exists w. - Qed. + Section Result. - Lemma P_meet_R_simple : forall e, R_simple P e. - Proof. - intros e H3. - destruct (try2 H3) as [m Hm]. - intros Hin. apply Hm. - intros x Hx Hx'. - apply (Hin x Hx'). - unfold P. unfold F_with. - exists (Pf_ m), m. - split; last apply F_func_correctness. - easy. - Qed. + Lemma P_semi_decidable : semi_decidable P. + Proof. + apply F_with_semi_decidable. + Qed. - Lemma P_semi_decidable : semi_decidable P. - Proof. - apply F_with_semi_decidable. - Qed. - Lemma P_Listing: - forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). - Proof. - intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). - intros [L H1]. apply H. exists L; intuition. - apply P_bounded. - - exact H2. - - apply H0. - Qed. + Lemma P_simple : simple P. + Proof. + unfold simple; repeat split. + - rewrite EA.enum_iff. now apply P_semi_decidable. + - apply P_coinfinite. + - intros [q (Hqenum & Hqinf & Hq)]. + rewrite EA.enum_iff in Hqenum. + destruct (W_spec Hqenum) as [c HWq]. + apply (@P_meet_R_simple c). + intros [l Hqe]; apply Hqinf. + exists l. intros x Hqx. apply (Hqe x). + now rewrite HWq in Hqx. + intros x HWcx HPx. unfold W in HWcx. + rewrite <- (HWq x) in HWcx. apply (Hq x HWcx HPx). + Qed. - Lemma complToBound_compl p L b: - PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. - Proof. - intros H x. split. - - intros [H1 H1'] % in_filter_iff. - destruct Dec; cbn in H1'; try congruence. - enough (x <= b). - + intuition. - + apply in_seq in H1. lia. - - intros [H1 H2]. eapply in_filter_iff. split. - + apply in_seq; lia. - + destruct Dec; cbn; try tauto. exfalso. firstorder. - Qed. + End Result. - Lemma Compl_P_Listing: - forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L - /\ forall x, In x L -> ~ P x. - Proof. - intros n H. - apply (@P_Listing n). intros [L H1]. - apply H. exists (complToBound L (2*n)). repeat split. - - remember (complToBound_length L (2*n)). lia. - - apply complToBound_NoDup. - - intros x I % (@complToBound_compl P); intuition. - Qed. +End Assume_EA. - Lemma P_coinfinite : ~ exhaustible (compl P). - Proof. - eapply weakly_unbounded_non_finite. - intros n H. eapply Compl_P_Listing with (n := n). - intros (l & ? & ? & H2). - eapply H. - exists (firstn n l). - repeat split. - - rewrite firstn_length. lia. - - now eapply firstn_NoDup. - - intros ? ? % firstn_In. now eapply H2. - Qed. - - Lemma P_simple : simple P. - Proof. - unfold simple; repeat split. - - rewrite EA.enum_iff. now apply P_semi_decidable. - - apply P_coinfinite. - - intros [q (Hqenum & Hqinf & Hq)]. - rewrite EA.enum_iff in Hqenum. - destruct (W_spec Hqenum) as [c HWq]. - apply (@P_meet_R_simple c). - intros [l Hqe]; apply Hqinf. - exists l. intros x Hqx. apply (Hqe x). - now rewrite HWq in Hqx. - intros x HWcx HPx. unfold W in HWcx. - rewrite <- HWq in HWcx. apply (Hq x HWcx HPx). - Qed. - End Assume_EA. \ No newline at end of file +Require SyntheticComputability.Axioms.EA. + +Lemma EA_correctness: Σ φ, EA φ. +Proof. + Import SyntheticComputability.Axioms.EA.Assume_EA. + exists φ. + intros P HP%SyntheticComputability.Axioms.EA.enum_iff. + rewrite W_spec in HP. destruct HP as [c Hc]. + exists c. intros x. unfold W in Hc. + apply Hc. +Qed. \ No newline at end of file From f5c55134e8eb595aa7560047f49322a45f884f16 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 18 Mar 2024 18:39:51 +0100 Subject: [PATCH 15/54] dn of requirements --- .../ReducibilityDegrees/simple_extension.v | 98 ++++++++++++++++--- 1 file changed, 87 insertions(+), 11 deletions(-) diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 08fadc9..e21871a 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -742,8 +742,7 @@ Section Assume_EA. apply H; first easy. firstorder. Qed. - - +(* Lemma try1 e: non_finite e -> exists s, ( ~~(Pf_ s) # (W[s] e)) \/ @@ -764,22 +763,27 @@ Section Assume_EA. intro H. apply Hx. intro Hx'. admit. - - Admitted. + Admitted. *) + Lemma impl_dn (P Q: Prop): + (~~ (P -> Q)) <-> (P -> ~~Q) . + Proof. split; firstorder. Qed. + +(* Lemma try2: - forall e, non_finite e -> exists m, incl_e (Pf_ m) e. + forall e, non_finite e -> ~~ exists m, incl_e (Pf_ m) e. Proof. - strong induction e. intros He. + (* strong induction e. intros He. destruct (@try3 e H He) as [w [Hw| Hw]]. specialize (attention_active Hw) as Hw'. exists (S w). intro Hx. unfold active in Hw'. apply Hw'. intros y Hy1 [i [Hie Hi]]. apply (Hx y Hy1). now exists i. - now exists w. - Qed. + now exists w. *) + Admitted. *) - Lemma P_meet_R_simple : forall e, R_simple P e. + (* Lemma P_meet_R_simple : forall e, R_simple P e. Proof. intros e H3. destruct (try2 H3) as [m Hm]. @@ -790,8 +794,80 @@ Section Assume_EA. exists (Pf_ m), m. split; last apply F_func_correctness. easy. - Qed. - + Abort. *) + + Lemma which_one_is_the_best {X} (P Q: X -> Prop): + (forall x, P x -> Q x -> False) <-> (~ exists x, P x /\ Q x) . + Proof. + split. + - intros Hf [x [h1 h2]]. eapply Hf; eauto. + - intros Hf x px qx. eapply Hf; eauto. + Qed. + + Lemma quant_swap {X} (P: X -> Prop): + (~ forall x, ~ P x) -> ~~ exists x, P x. + Proof. eauto. Qed. + + + Lemma non_finite_not_bounded e: + non_finite e -> + ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x. + Proof. + intro H. unfold non_finite in H. + intros He. rewrite non_finite_nat in H. + specialize (H (2 * e + 1)). + apply H. intros [m [Hm1 [k Hmk]]]. + apply He. exists k, m; split; last lia. + now exists k. + Qed. +(* + Lemma : + forall y, y < e -> ~~ R_simple P y + (exists k, exists x, W[k] e x /\ 2 * e < x). + Proof. + + Qed. + *) + Lemma P_meet_R_simple' : forall e, ~ ~ R_simple P e. + Proof. + intro e. strong induction e. + unfold R_simple. rewrite impl_dn. + intros B%non_finite_not_bounded. + intros Hi. eapply B. intros (k&x&Hk&Hx). + eapply Hi. intros Hin. + pose (N := S (max e k)). + destruct (Dec (ext_intersect_W (Pf_ N) N e)) as [He|He]. + - assert (ext_pick (Pf_ N) N e). + { unfold ext_pick. split. easy. + exists x. split. destruct Hk as [w [Hw1 Hw2]]. + exists w. split; eauto. lia. eauto. } + assert (attention e N). + unfold attention. split. lia. + split. easy. intros j Hj%H. + intros Hj2. apply Hj. + intros Hj3. + unfold ext_pick in Hj2. + admit. + eapply attention_active in H1. + apply H1. intros y Hy1 [w [_ Hw]]. + eapply Hin. exists w. eapply Hw. + exists (Pf_ (S N)), (S N). + split. easy. eapply F_func_correctness. + - eapply He. intros y Hy1 [w [_ Hw]]. + eapply Hin. exists w. eapply Hw. + exists (Pf_ N), N; split. + easy. eapply F_func_correctness. + Admitted. + + (* Lemma Good_req Q e: ~~ R_simple Q e -> R_simple Q e. + Proof. + intros H. unfold R_simple, intersect. + rewrite which_one_is_the_best. + rewrite <- impl_dn. + Qed. *) + + + End Meet_Requirement. (* Lemma list_accu e: @@ -846,7 +922,7 @@ Section Assume_EA. - intros [q (Hqenum & Hqinf & Hq)]. rewrite EA.enum_iff in Hqenum. destruct (W_spec Hqenum) as [c HWq]. - apply (@P_meet_R_simple c). + apply (@P_meet_R_simple' c). intros H. apply H. intros [l Hqe]; apply Hqinf. exists l. intros x Hqx. apply (Hqe x). now rewrite HWq in Hqx. From 65c55d8f70fe3782fee9eeaf77fe85f2d021b02b Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 19 Mar 2024 15:01:20 +0100 Subject: [PATCH 16/54] P is simple --- .../ReducibilityDegrees/simple_extension.v | 140 ++++++++++++++---- 1 file changed, 113 insertions(+), 27 deletions(-) diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index e21871a..fc064c1 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -541,6 +541,12 @@ Section Assume_EA. apply attention_active. apply H. lia. Qed. + Lemma attention_not_hold e k: attention e k -> forall m, k < m -> ~ attention e m. + Proof. + intros H1 m Hm. eapply active_not_attention. + apply (active_hold H1 Hm). + Qed. + Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. Proof. intros H1 H2. @@ -742,6 +748,108 @@ Section Assume_EA. apply H; first easy. firstorder. Qed. + + Lemma attention_at_most_one e: ~ ~ (exists s, forall s', s < s' -> ~ attention e s'). + Proof. + ccase (exists k, attention e k) as [[w Hw]|H]. + - intros H. eapply H. exists w. + now eapply attention_not_hold. + - intros Hk. apply Hk. exists 0. + intros k' Hk' Ha. + apply H. now exists k'. + Qed. + + Lemma attention_at_most_one' k: + ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attention e s'). + Proof. + induction k. + - intros H. apply H. exists 42. intros ??. lia. + - intros H. apply IHk. intros [s Hs]; clear IHk. + specialize (@attention_at_most_one k) as Hk. + apply Hk. intros [sk Hsk]; clear Hk. + set (max sk s) as N. + eapply H. exists N. intros e He. + assert (e = k \/ e < k) as [->|Hek] by lia. + intros s' Hs'. eapply Hsk. lia. + intros s' Hs'. eapply Hs; lia. + Qed. + + Lemma non_finite_not_bounded e: + non_finite e -> + ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x. + Proof. + intro H. unfold non_finite in H. + intros He. rewrite non_finite_nat in H. + specialize (H (2 * e + 1)). + apply H. intros [m [Hm1 [k Hmk]]]. + apply He. exists k, m; split; last lia. + now exists k. + Qed. + + Lemma Pick_leas N e : + e < N -> + (exists w, w < e /\ ext_pick (Pf_ N) N w) -> (exists w, w < e /\ attention w N). + Proof. + intros HeN [w (Hw1 & Hw2)]. + assert (exists w, ext_pick (Pf_ N) N w) by now exists w. + eapply least_ex in H; last eauto. + destruct H as [k Hk]. assert (k <= w). + { enough (~ k > w) by lia. intro Hkw. + destruct Hk. rewrite safe_char in H0. + specialize (H0 w Hw2). lia. } + exists k. do 2 (split; first lia). eapply Hk. + Qed. + + Lemma non_finite_e_attention e: + non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (Pf_ k) k e \/ attention e k) . + Proof. + intros H He. + eapply (non_finite_not_bounded H); intros (b & x & Hxb1 & Hxb2). + eapply (@attention_at_most_one' e); intros [s Hs]. + eapply He. + pose (N := S (max (max b s) e)). + destruct (Dec (ext_intersect_W (Pf_ N) N e)) as [He'|He']. + - exists N. right. + assert (ext_pick (Pf_ N) N e). + { split; first easy. exists x. split; last easy. + eapply W_incl; last apply Hxb1. lia. } + split. lia. split. easy. + intros w HEw Hw. + assert (exists w, w < e /\ ext_pick (Pf_ N) N w). + now exists w. eapply Pick_leas in H1. + destruct H1 as [g [HEg Hg]]. + eapply (Hs g HEg); last apply Hg. lia. lia. + - exists N. now left. + Qed. + + Lemma impl_dn (P Q: Prop): (~~ (P -> Q)) <-> (P -> ~~Q) . + Proof. split; firstorder. Qed. + + Lemma P_meet_R_simple' : forall e, ~ ~ R_simple P e. + Proof. + intros e. + unfold R_simple. rewrite impl_dn. + intros He He'. + eapply (non_finite_e_attention He). + intros [k [H|H]]. + - apply He'. intros x. + apply H. unfold ext_intersect_W. + intros y Hy1 [w Hy2]. + apply (x y). now exists w. + exists (Pf_ k), k; split; eauto. + eapply F_func_correctness. + - eapply attention_active in H. + unfold active in H. + apply He'. intros x. + apply H. intros y Hy1 [w Hy2]. + apply (x y). now exists w. + exists (Pf_ (S k)), (S k); split; eauto. + eapply F_func_correctness. + Qed. + + + + (* Lemma try1 e: non_finite e -> exists s, @@ -765,10 +873,7 @@ Section Assume_EA. - Admitted. *) - Lemma impl_dn (P Q: Prop): - (~~ (P -> Q)) <-> (P -> ~~Q) . - Proof. split; firstorder. Qed. - + (* Lemma try2: forall e, non_finite e -> ~~ exists m, incl_e (Pf_ m) e. @@ -795,7 +900,7 @@ Section Assume_EA. split; last apply F_func_correctness. easy. Abort. *) - +(* Lemma which_one_is_the_best {X} (P Q: X -> Prop): (forall x, P x -> Q x -> False) <-> (~ exists x, P x /\ Q x) . Proof. @@ -807,28 +912,9 @@ Section Assume_EA. Lemma quant_swap {X} (P: X -> Prop): (~ forall x, ~ P x) -> ~~ exists x, P x. Proof. eauto. Qed. + *) - - Lemma non_finite_not_bounded e: - non_finite e -> - ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x. - Proof. - intro H. unfold non_finite in H. - intros He. rewrite non_finite_nat in H. - specialize (H (2 * e + 1)). - apply H. intros [m [Hm1 [k Hmk]]]. - apply He. exists k, m; split; last lia. - now exists k. - Qed. -(* - Lemma : - forall y, y < e -> ~~ R_simple P y - (exists k, exists x, W[k] e x /\ 2 * e < x). - Proof. - - Qed. - *) - Lemma P_meet_R_simple' : forall e, ~ ~ R_simple P e. + (* Lemma P_meet_R_simple' : forall e, ~ ~ R_simple P e. Proof. intro e. strong induction e. unfold R_simple. rewrite impl_dn. @@ -857,7 +943,7 @@ Section Assume_EA. eapply Hin. exists w. eapply Hw. exists (Pf_ N), N; split. easy. eapply F_func_correctness. - Admitted. + Admitted. *) (* Lemma Good_req Q e: ~~ R_simple Q e -> R_simple Q e. Proof. From 161b808a28a43db4db1087c52fdc67e5d3c0c54b Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 19 Mar 2024 15:49:58 +0100 Subject: [PATCH 17/54] refactor --- .../ReducibilityDegrees/priority_method.v | 248 ++++++++- .../ReducibilityDegrees/simple_extension.v | 525 +++--------------- 2 files changed, 311 insertions(+), 462 deletions(-) diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index aaf5699..a237ebf 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -129,4 +129,250 @@ Section StrongInduction. End StrongInduction. -Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. \ No newline at end of file +Tactic Notation "strong" "induction" ident(n) := induction n using strong_induction. + +Section EWO. + Variable p: nat -> Prop. + Inductive T | (n: nat) : Prop := C (phi: ~p n -> T (S n)). + + Definition T_elim (q: nat -> Type) + : (forall n, (~p n -> q (S n)) -> q n) -> + forall n, T n -> q n + := fun e => fix f n a := + let (phi) := a in + e n (fun h => f (S n) (phi h)). + + (*** EWO for Numbers *) + + Lemma TI n : + p n -> T n. + Proof. + intros H. constructor. intros H1. destruct (H1 H). + Qed. + + Lemma TD n : + T (S n) -> T n. + Proof. + intros H. constructor. intros _. exact H. + Qed. + + Variable p_dec: forall n, dec (p n). + + Definition TE (q: nat -> Type) + : (forall n, p n -> q n) -> + (forall n, ~p n -> q (S n) -> q n) -> + forall n, T n -> q n. + Proof. + intros e1 e2. + apply T_elim. intros n IH. + destruct (p_dec n); auto. + Qed. + + (** From now on T will only be used through TI, TD, and TE *) + + + Lemma T_zero n : + T n -> T 0. + Proof. + induction n as [|n IH]. + - auto. + - intros H. apply IH. apply TD, H. + Qed. + + Definition ewo_nat : + ex p -> Σ x, p x. + Proof. + intros H. + refine (@TE (fun _ => Σ x, p x) _ _ 0 _). + - eauto. + - easy. + - destruct H as [n H]. apply (@T_zero n), TI, H. + Qed. + +End EWO. + +Notation unique p := (forall x x', p x -> p x' -> x = x'). + +Section LeastWitness. + + Definition safe p n := forall k, k < n -> ~ p k. + Definition least p n := p n /\ safe p n. + + Fact least_unique p : unique (least p). + Proof. + intros n n' [H1 H2] [H1' H2']. + enough (~(n < n') /\ ~(n' < n)) by lia. + split; intros H. + - eapply H2'; eassumption. + - eapply H2; eassumption. + Qed. + + Fact safe_O p : + safe p 0. + Proof. + hnf. lia. + Qed. + + Fact safe_S p n : + safe p n -> ~p n -> safe p (S n). + Proof. + intros H1 H2 k H3. unfold safe in *. + assert (k < n \/ k = n) as [H|H] by lia. + - auto. + - congruence. + Qed. + + Fact safe_char p n : + safe p n <-> forall k, p k -> k >= n. + Proof. + split. + - intros H k H1. + enough (k < n -> False) by lia. + intros H2. apply H in H2. auto. + - intros H k H1 H2. apply H in H2. lia. + Qed. + + Fact safe_char_S p n : + safe p (S n) <-> safe p n /\ ~p n. + Proof. + split. + - intros H. split. + + intros k H1. apply H. lia. + + apply H. lia. + - intros [H1 H2]. apply safe_S; assumption. + Qed. + + Fact safe_eq p n k : + safe p n -> k <= n -> p k -> k = n. + Proof. + intros H1 H2 H3. unfold safe in *. + enough (~(k < n)) by lia. + specialize (H1 k). tauto. + Qed. + + Fact E14 x y z : + x - y = z <-> least (fun k => x <= y + k) z. + Proof. + assert (H: least (fun k => x <= y + k) (x - y)). + { split; unfold safe; lia. } + split. congruence. + eapply least_unique, H. + Qed. + + (*** Certifying LWOs *) + + Section LWO. + Variable p : nat -> Prop. + Variable p_dec : forall x, dec (p x). + + Definition lwo : + forall n, (Σ k, k < n /\ least p k) + safe p n. + Proof. + induction n as [|n IH]. + - right. apply safe_O. + - destruct IH as [[k H1]|H1]. + + left. exists k. destruct H1 as [H1 H2]. split. lia. exact H2. + + destruct (p_dec n). + * left. exists n. split. lia. easy. + * right. apply safe_S; assumption. + Defined. + + Definition lwo' : + forall n, (Σ k, k <= n /\ least p k) + safe p (S n). + Proof. + intros n. + destruct (lwo (S n)) as [(k&H1&H2)|H]. + - left. exists k. split. lia. exact H2. + - right. exact H. + Qed. + + Definition least_sig : + (Σ x, p x) -> Σ x, (least p) x. + Proof. + intros [n H]. + destruct (lwo (S n)) as [(k&H1&H2)|H1]. + - exists k. exact H2. + - exfalso. apply (H1 n). lia. exact H. + Qed. + + Definition least_ex : + ex p -> ex (least p). + Proof. + intros [n H]. + destruct (lwo (S n)) as [(k&H1&H2)|H1]. + - exists k. exact H2. + - exfalso. apply (H1 n). lia. exact H. + Qed. + + Definition safe_dec n : + dec (safe p n). + Proof. + destruct (lwo n) as [[k H1]|H1]. + - right. intros H. exfalso. + destruct H1 as [H1 H2]. apply (H k). exact H1. apply H2. + - left. exact H1. + Defined. + + Definition least_dec n : + dec (least p n). + Proof. + unfold least. + destruct (p_dec n) as [H|H]. + 2:{ right. tauto. } + destruct (safe_dec n) as [H1|H1]. + - left. easy. + - right. tauto. + Qed. + End LWO. + + Lemma exists_bounded_dec' P: + (forall x, dec (P x)) -> forall k, dec (exists n, n < k /\ P n). + Proof. + intros Hp k. + induction k. right; intros [? [? _]]; lia. + destruct IHk as [Hw|Hw]. + - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. + - destruct (Hp k). left. exists k; split; eauto; lia. + right. intros [x [Hx1 Hx2]]. + assert (x = k \/ x < k) as [->|Hk] by lia; firstorder. + Qed. + + Lemma exists_bounded_dec P: + (forall x, dec (P x)) -> forall k, dec (exists n, n <= k /\ P n). + Proof. + intros Hp k. + induction k. destruct (Hp 0). left. exists 0. eauto. + right. intros [x [Hx Hx']]. now assert (x=0) as -> by lia. + destruct IHk as [Hw|Hw]. + - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. + - destruct (Hp (S k)). left. exists (S k); split; eauto; lia. + right. intros [x [Hx1 Hx2]]. + assert (x = S k \/ x <= k) as [->|Hk] by lia; firstorder. + Qed. + + Definition bounded (P: nat -> Prop) := Σ N, forall x, P x -> x < N. + + Fact bouned_le (P Q: nat -> Prop) N : + (forall x, P x -> x < N) -> + (exists x, P x /\ Q x) <-> exists x, x < N /\ P x /\ Q x. + Proof. + intros Hn; split. + - intros [x Hx]. exists x; intuition eauto. + - intros (x&c&Hc). exists x; intuition eauto. + Qed. + + Lemma bounded_dec (P Q: nat -> Prop): + (forall x, dec (P x)) -> (forall x, dec (Q x)) -> + bounded P -> dec (exists n, P n /\ Q n). + Proof. + intros H1 H2 [N HN]. + assert (dec (exists n, n < N /\ P n /\ Q n)) as [H|H]. + - eapply exists_bounded_dec'. intro; eapply and_dec; eauto. + - left. rewrite bouned_le; eauto. + - right. intros H'%(@bouned_le _ _ N); easy. + Qed. + Lemma dec_neg_dec P: dec P -> dec (~ P). + Proof. intros [H|H]. right; eauto. now left. Qed. + +End LeastWitness. + diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index fc064c1..56ed7a7 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -66,249 +66,6 @@ Section ComplToBound. Qed. End ComplToBound. -Section EWO. - Variable p: nat -> Prop. - Inductive T | (n: nat) : Prop := C (phi: ~p n -> T (S n)). - - Definition T_elim (q: nat -> Type) - : (forall n, (~p n -> q (S n)) -> q n) -> - forall n, T n -> q n - := fun e => fix f n a := - let (phi) := a in - e n (fun h => f (S n) (phi h)). - - (*** EWO for Numbers *) - - Lemma TI n : - p n -> T n. - Proof. - intros H. constructor. intros H1. destruct (H1 H). - Qed. - - Lemma TD n : - T (S n) -> T n. - Proof. - intros H. constructor. intros _. exact H. - Qed. - - Variable p_dec: forall n, dec (p n). - - Definition TE (q: nat -> Type) - : (forall n, p n -> q n) -> - (forall n, ~p n -> q (S n) -> q n) -> - forall n, T n -> q n. - Proof. - intros e1 e2. - apply T_elim. intros n IH. - destruct (p_dec n); auto. - Qed. - - (** From now on T will only be used through TI, TD, and TE *) - - - Lemma T_zero n : - T n -> T 0. - Proof. - induction n as [|n IH]. - - auto. - - intros H. apply IH. apply TD, H. - Qed. - - Definition ewo_nat : - ex p -> Σ x, p x. - Proof. - intros H. - refine (@TE (fun _ => Σ x, p x) _ _ 0 _). - - eauto. - - easy. - - destruct H as [n H]. apply (@T_zero n), TI, H. - Qed. - -End EWO. - -Notation unique p := (forall x x', p x -> p x' -> x = x'). -Section LeastWitness. - Definition safe p n := forall k, k < n -> ~ p k. - Definition least p n := p n /\ safe p n. - - Fact least_unique p : unique (least p). - Proof. - intros n n' [H1 H2] [H1' H2']. - enough (~(n < n') /\ ~(n' < n)) by lia. - split; intros H. - - eapply H2'; eassumption. - - eapply H2; eassumption. - Qed. - - Fact safe_O p : - safe p 0. - Proof. - hnf. lia. - Qed. - - Fact safe_S p n : - safe p n -> ~p n -> safe p (S n). - Proof. - intros H1 H2 k H3. unfold safe in *. - assert (k < n \/ k = n) as [H|H] by lia. - - auto. - - congruence. - Qed. - - Fact safe_char p n : - safe p n <-> forall k, p k -> k >= n. - Proof. - split. - - intros H k H1. - enough (k < n -> False) by lia. - intros H2. apply H in H2. auto. - - intros H k H1 H2. apply H in H2. lia. - Qed. - - Fact safe_char_S p n : - safe p (S n) <-> safe p n /\ ~p n. - Proof. - split. - - intros H. split. - + intros k H1. apply H. lia. - + apply H. lia. - - intros [H1 H2]. apply safe_S; assumption. - Qed. - - Fact safe_eq p n k : - safe p n -> k <= n -> p k -> k = n. - Proof. - intros H1 H2 H3. unfold safe in *. - enough (~(k < n)) by lia. - specialize (H1 k). tauto. - Qed. - - Fact E14 x y z : - x - y = z <-> least (fun k => x <= y + k) z. - Proof. - assert (H: least (fun k => x <= y + k) (x - y)). - { split; unfold safe; lia. } - split. congruence. - eapply least_unique, H. - Qed. - - (*** Certifying LWOs *) - - Section LWO. - Variable p : nat -> Prop. - Variable p_dec : forall x, dec (p x). - - Definition lwo : - forall n, (Σ k, k < n /\ least p k) + safe p n. - Proof. - induction n as [|n IH]. - - right. apply safe_O. - - destruct IH as [[k H1]|H1]. - + left. exists k. destruct H1 as [H1 H2]. split. lia. exact H2. - + destruct (p_dec n). - * left. exists n. split. lia. easy. - * right. apply safe_S; assumption. - Defined. - - Definition lwo' : - forall n, (Σ k, k <= n /\ least p k) + safe p (S n). - Proof. - intros n. - destruct (lwo (S n)) as [(k&H1&H2)|H]. - - left. exists k. split. lia. exact H2. - - right. exact H. - Qed. - - Definition least_sig : - (Σ x, p x) -> Σ x, (least p) x. - Proof. - intros [n H]. - destruct (lwo (S n)) as [(k&H1&H2)|H1]. - - exists k. exact H2. - - exfalso. apply (H1 n). lia. exact H. - Qed. - - Definition least_ex : - ex p -> ex (least p). - Proof. - intros [n H]. - destruct (lwo (S n)) as [(k&H1&H2)|H1]. - - exists k. exact H2. - - exfalso. apply (H1 n). lia. exact H. - Qed. - - Definition safe_dec n : - dec (safe p n). - Proof. - destruct (lwo n) as [[k H1]|H1]. - - right. intros H. exfalso. - destruct H1 as [H1 H2]. apply (H k). exact H1. apply H2. - - left. exact H1. - Defined. - - Definition least_dec n : - dec (least p n). - Proof. - unfold least. - destruct (p_dec n) as [H|H]. - 2:{ right. tauto. } - destruct (safe_dec n) as [H1|H1]. - - left. easy. - - right. tauto. - Qed. - End LWO. - - Lemma exists_bounded_dec' P: - (forall x, dec (P x)) -> forall k, dec (exists n, n < k /\ P n). - Proof. - intros Hp k. - induction k. right; intros [? [? _]]; lia. - destruct IHk as [Hw|Hw]. - - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. - - destruct (Hp k). left. exists k; split; eauto; lia. - right. intros [x [Hx1 Hx2]]. - assert (x = k \/ x < k) as [->|Hk] by lia; firstorder. - Qed. - - Lemma exists_bounded_dec P: - (forall x, dec (P x)) -> forall k, dec (exists n, n <= k /\ P n). - Proof. - intros Hp k. - induction k. destruct (Hp 0). left. exists 0. eauto. - right. intros [x [Hx Hx']]. now assert (x=0) as -> by lia. - destruct IHk as [Hw|Hw]. - - left. destruct Hw as [x [Hx1 Hx2]]. exists x; split; eauto; lia. - - destruct (Hp (S k)). left. exists (S k); split; eauto; lia. - right. intros [x [Hx1 Hx2]]. - assert (x = S k \/ x <= k) as [->|Hk] by lia; firstorder. - Qed. - - Definition bounded (P: nat -> Prop) := Σ N, forall x, P x -> x < N. - - Fact bouned_le (P Q: nat -> Prop) N : - (forall x, P x -> x < N) -> - (exists x, P x /\ Q x) <-> exists x, x < N /\ P x /\ Q x. - Proof. - intros Hn; split. - - intros [x Hx]. exists x; intuition eauto. - - intros (x&c&Hc). exists x; intuition eauto. - Qed. - - Lemma bounded_dec (P Q: nat -> Prop): - (forall x, dec (P x)) -> (forall x, dec (Q x)) -> - bounded P -> dec (exists n, P n /\ Q n). - Proof. - intros H1 H2 [N HN]. - assert (dec (exists n, n < N /\ P n /\ Q n)) as [H|H]. - - eapply exists_bounded_dec'. intro; eapply and_dec; eauto. - - left. rewrite bouned_le; eauto. - - right. intros H'%(@bouned_le _ _ N); easy. - Qed. - Lemma dec_neg_dec P: dec P -> dec (~ P). - Proof. intros [H|H]. right; eauto. now left. Qed. - -End LeastWitness. - Section Assume_EA. Variable φ: nat -> nat -> option nat. @@ -366,9 +123,11 @@ Section Assume_EA. End EA_dec. - Definition disj {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. - Definition intersect {X} (A B: X -> Prop) := forall x, A x -> B x -> False. - Notation "A # B" := (disj A B) (at level 30). + Definition disj_list_pred {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. + Definition disj_pred {X} (A B: X -> Prop) := forall x, A x -> B x -> False. + Notation "A # B" := (disj_list_pred A B) (at level 30). + Notation "A #ₚ B" := (disj_pred A B) (at level 30). + Section Extension. @@ -466,16 +225,21 @@ Section Assume_EA. Definition P := F_with simple_extendsion. Definition non_finite e := ~ exhaustible (W e). - Definition incl_e L e := ~ (L # (W e)). + + Fact In_Pf k y: In y (Pf_ k) -> P y . + Proof. + intros H. exists (Pf_ k), k. + split; [easy| apply F_func_correctness]. + Qed. End Simple_Extension. Section Requirements. - Definition R_simple_list L e := non_finite e -> incl_e L e. + Definition R_simple P e := non_finite e -> ~ (W e #ₚ P). Definition attention e n := e < n /\ least (ext_pick (Pf_ n) n) e. Definition active e n := ~ (Pf_ n) # W[n] e. - Definition pick_el (e x: nat) := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. + Definition pick_el e x := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. End Requirements. @@ -504,7 +268,7 @@ Section Assume_EA. now apply h2. Qed. - Lemma active_incl e n: active e n -> forall m, n <= m -> active e m . + Lemma active_always_active e n: active e n -> forall m, n <= m -> active e m . Proof. intros H m Hm Hx. apply H. eapply (intersect_mono Hx). @@ -535,29 +299,40 @@ Section Assume_EA. Lemma active_not_attention e k: active e k -> ~ attention e k. Proof. now intros h2 [_ [[h _] _]]. Qed. - Lemma active_hold e k: attention e k -> forall m, k < m -> active e m. + Lemma attention_always_active e k: attention e k -> forall m, k < m -> active e m. Proof. - intros. eapply active_incl. + intros. eapply active_always_active. apply attention_active. apply H. lia. Qed. - Lemma attention_not_hold e k: attention e k -> forall m, k < m -> ~ attention e m. + Lemma attention_always_not_attention e k: + attention e k -> forall m, k < m -> ~ attention e m. Proof. intros H1 m Hm. eapply active_not_attention. - apply (active_hold H1 Hm). + apply (attention_always_active H1 Hm). + Qed. + + Lemma attention_at_most_once e: ~ ~ (exists s, forall s', s < s' -> ~ attention e s'). + Proof. + ccase (exists k, attention e k) as [[w Hw]|H]. + - intros H. eapply H. exists w. + now eapply attention_always_not_attention. + - intros Hk. apply Hk. exists 0. + intros k' Hk' Ha. + apply H. now exists k'. Qed. Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. Proof. intros H1 H2. - specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H1 a b)) as H1'. - specialize (fun a b => @active_not_attention _ _ (@active_hold _ _ H2 a b)) as H2'. + specialize (fun a b => @active_not_attention _ _ (@attention_always_active _ _ H1 a b)) as H1'. + specialize (fun a b => @active_not_attention _ _ (@attention_always_active _ _ H2 a b)) as H2'. enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. intro Hk. eapply H1'. apply Hk. easy. intro Hk. eapply H2'. apply Hk. easy. Qed. - Lemma spec_uni e x1 x2: pick_el e x1 -> pick_el e x2 -> x1 = x2 . + Lemma pick_el_uni e x1 x2: pick_el e x1 -> pick_el e x2 -> x1 = x2 . Proof. intros [k [Ht Hk]] [k' [Ht' Hk']]. assert (k=k') as <-. eapply attention_uni; eauto. @@ -614,7 +389,7 @@ Section Assume_EA. + constructor; last now exact H2. intros In. inv HL. apply H4 in In as [y (Hs & h2 & h3)]. - now rewrite (spec_uni H0 Hs) in H6. + now rewrite (pick_el_uni H0 Hs) in H6. + cbn. rewrite H. trivial. + destruct H3 as [->|]. * exists a; intuition. @@ -710,7 +485,7 @@ Section Assume_EA. + destruct Dec; cbn; try tauto. exfalso. firstorder. Qed. - Lemma Compl_P_Listing: + Lemma compl_P_Listing: forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L /\ forall x, In x L -> ~ P x. Proof. @@ -725,7 +500,7 @@ Section Assume_EA. Lemma P_coinfinite : ~ exhaustible (compl P). Proof. eapply weakly_unbounded_non_finite. - intros n H. eapply Compl_P_Listing with (n := n). + intros n H. eapply compl_P_Listing with (n := n). intros (l & ? & ? & H2). eapply H. exists (firstn n l). @@ -739,33 +514,13 @@ Section Assume_EA. Section Meet_Requirement. - Definition R_simple P e := non_finite e -> ~ (intersect (W e) P). - - Lemma R_simple_acc e L: - R_simple_list L e -> forall L', incl L L' -> R_simple_list L' e . - Proof. - intros H L' HL' h1 h2. - apply H; first easy. - firstorder. - Qed. - - Lemma attention_at_most_one e: ~ ~ (exists s, forall s', s < s' -> ~ attention e s'). - Proof. - ccase (exists k, attention e k) as [[w Hw]|H]. - - intros H. eapply H. exists w. - now eapply attention_not_hold. - - intros Hk. apply Hk. exists 0. - intros k' Hk' Ha. - apply H. now exists k'. - Qed. - - Lemma attention_at_most_one' k: + Lemma attention_at_most_once_bound k: ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attention e s'). Proof. induction k. - intros H. apply H. exists 42. intros ??. lia. - intros H. apply IHk. intros [s Hs]; clear IHk. - specialize (@attention_at_most_one k) as Hk. + specialize (@attention_at_most_once k) as Hk. apply Hk. intros [sk Hsk]; clear Hk. set (max sk s) as N. eapply H. exists N. intros e He. @@ -775,8 +530,7 @@ Section Assume_EA. Qed. Lemma non_finite_not_bounded e: - non_finite e -> - ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x. + non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x. Proof. intro H. unfold non_finite in H. intros He. rewrite non_finite_nat in H. @@ -786,9 +540,9 @@ Section Assume_EA. now exists k. Qed. - Lemma Pick_leas N e : - e < N -> - (exists w, w < e /\ ext_pick (Pf_ N) N w) -> (exists w, w < e /\ attention w N). + Lemma ext_pick_attention N e: + e < N -> (exists w, w < e /\ ext_pick (Pf_ N) N w) -> + (exists w, w < e /\ attention w N). Proof. intros HeN [w (Hw1 & Hw2)]. assert (exists w, ext_pick (Pf_ N) N w) by now exists w. @@ -800,12 +554,12 @@ Section Assume_EA. exists k. do 2 (split; first lia). eapply Hk. Qed. - Lemma non_finite_e_attention e: + Lemma non_finite_attention e: non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (Pf_ k) k e \/ attention e k) . Proof. intros H He. eapply (non_finite_not_bounded H); intros (b & x & Hxb1 & Hxb2). - eapply (@attention_at_most_one' e); intros [s Hs]. + eapply (@attention_at_most_once_bound e); intros [s Hs]. eapply He. pose (N := S (max (max b s) e)). destruct (Dec (ext_intersect_W (Pf_ N) N e)) as [He'|He']. @@ -816,182 +570,34 @@ Section Assume_EA. split. lia. split. easy. intros w HEw Hw. assert (exists w, w < e /\ ext_pick (Pf_ N) N w). - now exists w. eapply Pick_leas in H1. + now exists w. eapply ext_pick_attention in H1. destruct H1 as [g [HEg Hg]]. eapply (Hs g HEg); last apply Hg. lia. lia. - exists N. now left. - Qed. - - Lemma impl_dn (P Q: Prop): (~~ (P -> Q)) <-> (P -> ~~Q) . - Proof. split; firstorder. Qed. - - Lemma P_meet_R_simple' : forall e, ~ ~ R_simple P e. - Proof. - intros e. - unfold R_simple. rewrite impl_dn. - intros He He'. - eapply (non_finite_e_attention He). - intros [k [H|H]]. - - apply He'. intros x. - apply H. unfold ext_intersect_W. - intros y Hy1 [w Hy2]. - apply (x y). now exists w. - exists (Pf_ k), k; split; eauto. - eapply F_func_correctness. - - eapply attention_active in H. - unfold active in H. - apply He'. intros x. - apply H. intros y Hy1 [w Hy2]. - apply (x y). now exists w. - exists (Pf_ (S k)), (S k); split; eauto. - eapply F_func_correctness. - Qed. - - - - -(* - Lemma try1 e: - non_finite e -> exists s, - ( ~~(Pf_ s) # (W[s] e)) \/ - (ext_intersect_W (Pf_ s) e s /\ exists x, 2 * e < x /\ (W[s] e) x). - Proof. - (* intros H HI. unfold inifite in H. *) - (* rewrite non_finite_spec in H. *) - Admitted. - - Lemma try3 e: - (forall n, n < e -> non_finite n -> exists k, incl_e (Pf_ k) n) -> - non_finite e -> - exists w, attention e w \/ incl_e (Pf_ w) e. - Proof. - intros He Hinfe. - destruct (try1 Hinfe) as [w [ Hx| Hw]]. - - exists w. right. unfold incl_e. - intro H. apply Hx. intro Hx'. - admit. - - - Admitted. *) - - -(* - Lemma try2: - forall e, non_finite e -> ~~ exists m, incl_e (Pf_ m) e. - Proof. - (* strong induction e. intros He. - destruct (@try3 e H He) as [w [Hw| Hw]]. - specialize (attention_active Hw) as Hw'. - exists (S w). - intro Hx. unfold active in Hw'. - apply Hw'. intros y Hy1 [i [Hie Hi]]. - apply (Hx y Hy1). now exists i. - now exists w. *) - Admitted. *) - - (* Lemma P_meet_R_simple : forall e, R_simple P e. - Proof. - intros e H3. - destruct (try2 H3) as [m Hm]. - intros Hin. apply Hm. - intros x Hx Hx'. - apply (Hin x Hx'). - unfold P. unfold F_with. - exists (Pf_ m), m. - split; last apply F_func_correctness. - easy. - Abort. *) -(* - Lemma which_one_is_the_best {X} (P Q: X -> Prop): - (forall x, P x -> Q x -> False) <-> (~ exists x, P x /\ Q x) . - Proof. - split. - - intros Hf [x [h1 h2]]. eapply Hf; eauto. - - intros Hf x px qx. eapply Hf; eauto. Qed. - Lemma quant_swap {X} (P: X -> Prop): - (~ forall x, ~ P x) -> ~~ exists x, P x. - Proof. eauto. Qed. - *) - - (* Lemma P_meet_R_simple' : forall e, ~ ~ R_simple P e. + Lemma ext_intersect_W_intersect k e: + ~ (Pf_ k # W[k] e) -> W e #ₚ P -> False. Proof. - intro e. strong induction e. - unfold R_simple. rewrite impl_dn. - intros B%non_finite_not_bounded. - intros Hi. eapply B. intros (k&x&Hk&Hx). - eapply Hi. intros Hin. - pose (N := S (max e k)). - destruct (Dec (ext_intersect_W (Pf_ N) N e)) as [He|He]. - - assert (ext_pick (Pf_ N) N e). - { unfold ext_pick. split. easy. - exists x. split. destruct Hk as [w [Hw1 Hw2]]. - exists w. split; eauto. lia. eauto. } - assert (attention e N). - unfold attention. split. lia. - split. easy. intros j Hj%H. - intros Hj2. apply Hj. - intros Hj3. - unfold ext_pick in Hj2. - admit. - eapply attention_active in H1. - apply H1. intros y Hy1 [w [_ Hw]]. - eapply Hin. exists w. eapply Hw. - exists (Pf_ (S N)), (S N). - split. easy. eapply F_func_correctness. - - eapply He. intros y Hy1 [w [_ Hw]]. - eapply Hin. exists w. eapply Hw. - exists (Pf_ N), N; split. - easy. eapply F_func_correctness. - Admitted. *) - - (* Lemma Good_req Q e: ~~ R_simple Q e -> R_simple Q e. - Proof. - intros H. unfold R_simple, intersect. - rewrite which_one_is_the_best. - rewrite <- impl_dn. - Qed. *) - + unfold ext_intersect_W. + intros H1 H2. apply H1. + intros y Hy1 [w Hy2]. + eapply (H2 y). now exists w. + eapply (In_Pf Hy1). + Qed. + Lemma P_meet_R_simple : forall e, R_simple P e. + Proof. + intros e He. intros He'. + eapply (non_finite_attention He). + intros [k [H|H]]. + - eapply ext_intersect_W_intersect; eauto. + - eapply attention_active in H. + eapply ext_intersect_W_intersect; eauto. + Qed. End Meet_Requirement. - (* Lemma list_accu e: - (forall k, k < e -> inifite k -> exists L, incl_e L k /\ exists n, F_ simple_extendsion n L) -> - exists m L, F_ simple_extendsion m L /\ forall n, n < e -> R_simple_list L n. - Proof. - intros. induction e. - { exists 0, []; split; first econstructor. intros n Hn. lia. } - destruct IHe as [m [HL' [Hm HL2']]]. - { intros n Hn He. apply H. lia. easy. } - pose (H e). - (* destruct (H e) as [L [HL1 HL2]]; first lia. *) - (* destruct HL2 as [k Hk]. *) - - (* destruct (dec_le m k). - + exists k, L; split; eauto. - intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. - apply HL1. eapply R_simple_acc. apply HL2'; first easy. - eapply F_mono; eauto. - + exists m, HL'; split; eauto. - intros n Hn. assert (n = e \/ n < e) as [->|H'] by lia. - eapply R_simple_acc. apply HL1. eapply F_mono; eauto with lia. - now eapply HL2'. - Qed. *) - Admitted. *) - - - (* Lemma eventually_attention e m L: - inifite e -> - (forall n : nat, n < e -> inifite n -> incl_e L n -> F_ simple_extendsion m L) -> - exists k, attention e k. - Proof. - intros H1 H2. - unfold attention. - - Admitted. *) - - Section Result. Lemma P_semi_decidable : semi_decidable P. @@ -999,7 +605,6 @@ Section Assume_EA. apply F_with_semi_decidable. Qed. - Lemma P_simple : simple P. Proof. unfold simple; repeat split. @@ -1008,7 +613,7 @@ Section Assume_EA. - intros [q (Hqenum & Hqinf & Hq)]. rewrite EA.enum_iff in Hqenum. destruct (W_spec Hqenum) as [c HWq]. - apply (@P_meet_R_simple' c). intros H. apply H. + apply (@P_meet_R_simple c). intros [l Hqe]; apply Hqinf. exists l. intros x Hqx. apply (Hqe x). now rewrite HWq in Hqx. @@ -1020,14 +625,12 @@ Section Assume_EA. End Assume_EA. - Require SyntheticComputability.Axioms.EA. Lemma EA_correctness: Σ φ, EA φ. Proof. - Import SyntheticComputability.Axioms.EA.Assume_EA. - exists φ. - intros P HP%SyntheticComputability.Axioms.EA.enum_iff. + Import SyntheticComputability.Axioms.EA.Assume_EA. + exists φ. intros P HP%SyntheticComputability.Axioms.EA.enum_iff. rewrite W_spec in HP. destruct HP as [c Hc]. exists c. intros x. unfold W in Hc. apply Hc. From e790d585323488018c18bf2c7f93117efc8dd22c Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 19 Mar 2024 20:38:29 +0100 Subject: [PATCH 18/54] adds detail --- .../ReducibilityDegrees/priority_method.v | 3 +- .../ReducibilityDegrees/simple_extension.v | 67 ++++++++++++++++++- 2 files changed, 66 insertions(+), 4 deletions(-) diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index a237ebf..09ba2d9 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -374,5 +374,4 @@ Section LeastWitness. Lemma dec_neg_dec P: dec P -> dec (~ P). Proof. intros [H|H]. right; eauto. now left. Qed. -End LeastWitness. - +End LeastWitness. \ No newline at end of file diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 56ed7a7..b97afa2 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -478,7 +478,7 @@ Section Assume_EA. - intros [H1 H1'] % in_filter_iff. destruct Dec; cbn in H1'; try congruence. enough (x <= b). - + intuition. + + firstorder. + apply in_seq in H1. lia. - intros [H1 H2]. eapply in_filter_iff. split. + apply in_seq; lia. @@ -605,7 +605,7 @@ Section Assume_EA. apply F_with_semi_decidable. Qed. - Lemma P_simple : simple P. + Theorem P_simple : simple P. Proof. unfold simple; repeat split. - rewrite EA.enum_iff. now apply P_semi_decidable. @@ -623,6 +623,69 @@ Section Assume_EA. End Result. + Section Effectively_Simple. + + Definition effectively_simple P := + simple P /\ exists f, + forall e, (forall x, W e x -> (compl P) x) -> forall x, W e x -> x < (f e). + + Lemma attention_pick e k: attention e k -> exists x, x > 2*e /\ P x /\ W e x. + Proof. + intros [He H]. + edestruct (ext_pick_witness) as [x Hx]. + { destruct H. eapply e0. } + assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. + exists e. assert (Pf_ k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + split; first easy. split; first easy. easy. + assert (x = a) as ->. eapply (@extend_uni simple_extendsion); cbn; eauto. + exists a. split. destruct H4, H0, H1, H4, H4. + assert (x=e) as HE. + { eapply least_unique; last eapply H. + enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } + lia. split. exists (Pf_ (S k)), (S k); split; eauto. now rewrite <- H2. + apply F_func_correctness. destruct H4, H0, H1, H4, H4, H4, H4. + assert (x=e) as HE. + { eapply least_unique; last eapply H. + enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } + exists x0; congruence. + exfalso. eapply (H3 x); exists e; do 2 (split; eauto). + enough ((Pf_ k) = (Pf_ (S k))) as <- by easy. + assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. + eapply F_uni; eauto. + Qed. + + Theorem P_effectively_simple: effectively_simple P. + Proof. + split; first apply P_simple. + exists (fun e => 2 * e + 1). + intros e He x Hex. enough (~ 2 * e < x) by lia. + intros Hex'. + assert (W e #ₚ P). + { intros y Hy1 Hy2. now apply (He y). } + assert (forall k, (Pf_ k) # W[k] e). + { intros k y Hy1 [w [_ Hy2]]. eapply (H y). exists w; eauto. + exists (Pf_ k), k; split; eauto. apply F_func_correctness. } + enough (exists k, attention e k) as [k Hk]. + (* apply H1. intros [k Hk]. *) + eapply attention_pick in Hk. + destruct Hk as [y (Hx1&Hx2&Hx3)]. + eapply (He y); eauto. + { unfold attention. + assert (exists k, least (ext_pick (Pf_ k) k) e /\ e < k). + { destruct Hex as [k Hk]. + pose (S (max e k)) as N. + unfold ext_pick. exists N. split. split. split. admit. + exists x. unfold ext_has_wit. split; last easy. + exists k; split; eauto. lia. 2: {lia. } admit. + } + admit. + } + Admitted. + + End Effectively_Simple. + End Assume_EA. Require SyntheticComputability.Axioms.EA. From 22242acbe64f46e85ef2f672da198da98ac42426 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sun, 24 Mar 2024 19:30:30 +0100 Subject: [PATCH 19/54] wall function --- .../low_simple_extension.v | 701 ++++++++++++++++++ .../ReducibilityDegrees/priority_method.v | 54 +- 2 files changed, 736 insertions(+), 19 deletions(-) create mode 100644 theories/ReducibilityDegrees/low_simple_extension.v diff --git a/theories/ReducibilityDegrees/low_simple_extension.v b/theories/ReducibilityDegrees/low_simple_extension.v new file mode 100644 index 0000000..472ccdf --- /dev/null +++ b/theories/ReducibilityDegrees/low_simple_extension.v @@ -0,0 +1,701 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +Require Import SyntheticComputability.Synthetic.DecidabilityFacts. +Require Export SyntheticComputability.Shared.FinitenessFacts. +Require Export SyntheticComputability.Shared.Pigeonhole. +Require Export SyntheticComputability.Shared.ListAutomation. +Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. +Require Import SyntheticComputability.ReducibilityDegrees.priority_method. +Import ListNotations. + +Section ComplToBound. + Definition complToBound L b : list nat + := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). + + Lemma complToBound_Bound L b : + forall x, In x (complToBound L b) -> x <= b. + Proof. + intros x [H % in_seq ?] % in_filter_iff. lia. + Qed. + Lemma filter_length {X} f (l : list X) : + length l = length (filter f l) + length (filter (fun x => (negb (f x))) l). + Proof. + induction l; cbn. + - reflexivity. + - destruct f; cbn; lia. + Qed. + Lemma filter_NoDup {X} f (l : list X) : + NoDup l -> NoDup (filter f l). + Proof. + induction 1; cbn. + - econstructor. + - destruct f; eauto. econstructor; auto. + intros ? % in_filter_iff. firstorder. + Qed. + Lemma complToBound_length L b: + length (complToBound L b) + length L >= S b. + Proof. + rewrite <- (seq_length (S b) 0). + erewrite filter_length with (l := seq 0 (S b)). + unfold complToBound. + eapply Plus.plus_le_compat_l_stt. + generalize (seq_NoDup (S b) 0). + generalize (seq 0 (S b)). clear. + intros. erewrite filter_ext with (g := fun x => Dec (In x L)). + 2:{ intros a. destruct Dec; cbn; destruct Dec; firstorder congruence. } + eapply NoDup_incl_length. now eapply filter_NoDup. + clear. induction l; cbn. + - firstorder. + - destruct Dec; cbn. 2: eauto. + intros ? [-> | ]; eauto. + Qed. + Lemma complToBound_NoDup L b: + NoDup (complToBound L b). + Proof. + eapply filter_NoDup, seq_NoDup. + Qed. + Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) -> In x l. + Proof. + induction n in x, l |- *; destruct l; cbn; firstorder. + Qed. + + Lemma firstn_NoDup {X} (l : list X) n : NoDup l -> NoDup (firstn n l). + Proof. + induction 1 in n |- *; destruct n; cbn; try now econstructor. + econstructor; eauto. + now intros ? % firstn_In. + Qed. +End ComplToBound. + +Section Assume_EA. + + Variable φ: nat -> nat -> option nat. + Definition EA := forall P, + semi_decidable P -> exists e, forall x, P x <-> exists n, φ e n = Some x. + Hypothesis EA: EA. + + Definition W_ n e x := φ n e = Some x. + Definition W e x := exists n, W_ e n x. + + Lemma W_spec: forall P, semi_decidable P -> exists e, forall x, P x <-> W e x. + Proof. intros P [e He]%EA. exists e; intros x; now rewrite He. Qed. + + Notation "'W[' s ']' e" := (fun x => exists n, n <= s /\ W_ e n x) (at level 30). + + Section EA_dec. + + Lemma W_dec e: forall x n, dec (W_ e n x). + Proof. + intros x n. + destruct (φ e n) eqn: E. + - destruct (Dec (x = n0)) as [E'|E']. + left. now subst. right. intros H. congruence. + - right. intros H. congruence. + Qed. + + Lemma W_bounded_dec e : forall x s, dec ((W[s] e) x). + Proof. + intros x s. cbn. eapply exists_bounded_dec. + intro; apply W_dec. + Qed. + + Lemma W_bounded_bounded e s: bounded (W[s] e). + Proof. + unfold bounded. + induction s. + - destruct (φ e 0) as [w|] eqn: E. + exists (S w). intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. rewrite E in Hn2. + injection Hn2. lia. + exists 42. intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. + rewrite E in Hn2. congruence. + - destruct IHs as [N HN]. + destruct (φ e (S s)) as [w|] eqn: E. + exists (max (S w) (S N)). intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. + rewrite E in Hn2. injection Hn2. lia. + eapply Nat.lt_trans. eapply HN. exists n; split; easy. lia. + exists N. intros x [n [Hn1 Hn2]]. + inv Hn1. unfold W_ in Hn2. + rewrite E in Hn2. congruence. + eapply HN. exists n. split; eauto. + Qed. + + End EA_dec. + + Definition disj_list_pred {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. + Definition disj_pred {X} (A B: X -> Prop) := forall x, A x -> B x -> False. + Notation "A # B" := (disj_list_pred A B) (at level 30). + Notation "A #ₚ B" := (disj_pred A B) (at level 30). + + Lemma extra_bounded f m: Σ b, forall n, n < m -> f n < b. + Proof. + induction m. + - exists 42. intros. inv H. + - destruct IHm as [b' Hb']. + exists (S (max b' (f m))). + intros ? H. inv H. lia. specialize (Hb' n H1). lia. + Qed. + + Section Assume_WALL. + Variable wall: nat -> list nat -> nat -> nat. + + Section Extension. + + Definition ext_intersect_W L n e := L # W[n] e. + Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. + Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. + Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. + Definition ext_least_choice L n x := exists e, ext_choice L n e x. + + End Extension. + + Section Extension_Facts. + + #[export]Instance ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). + Proof. + apply BaseLists.list_forall_dec. + intro x. eapply dec_neg_dec, exists_bounded_dec; eauto. + Qed. + + #[export] Instance ext_wall L n e x: dec (forall i, i <= e -> wall i L n < x). + Proof. eapply forall_bounded_dec; eauto. Qed. + + #[export]Instance ext_has_wit_dec L n e x : dec (ext_has_wit L n e x). + Proof. apply and_dec; first apply exists_bounded_dec; eauto. Qed. + + #[export]Instance ext_has_wit_exists_dec L n e : dec (exists x, ext_has_wit L n e x). + Proof. + unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. + intro x; eapply W_bounded_dec; eauto. eauto. + Qed. + + #[export]Instance ext_pick_dec L n e : dec (ext_pick L n e). + Proof. + eapply and_dec; first apply ext_intersect_W_dec. + unfold ext_has_wit. + eapply bounded_dec; last apply W_bounded_bounded. + intros x. eapply exists_bounded_dec. + intro; apply W_dec. eauto. + Qed. + + #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). + Proof. + eapply exists_bounded_dec'. intro x. + eapply least_dec. intros y. + eapply ext_pick_dec. + Qed. + + Fact ext_least_choice_dec L n: + (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). + Proof. + unfold ext_least_choice. + destruct (ext_pick_exists_dec L n) as [H|H]. + - left. apply ewo_nat; first last. + destruct H as [e (h1 & [(h4 & h2) h3])]. + eapply least_ex in h2. destruct h2 as [k h6]. + exists k, e. split; first easy; split. + 2: { eapply h6. } + split; last apply h3. split; first apply h4. + destruct h6. now exists k. + eapply ext_has_wit_dec. + intro x. unfold ext_choice. eapply exists_bounded_dec'. + intros x'. eapply and_dec. + apply least_dec. eapply ext_pick_dec. + apply least_dec. eapply ext_has_wit_dec. + - right. intros x [e (h1 & h2 & _)]. apply H. + exists e. split; eauto. + Qed. + + Fact ext_least_choice_uni l x y1 y2: + ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. + Proof. + intros [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. + assert (x0 = x1) as ->. eapply least_unique; eauto. + eapply least_unique; eauto. + Qed. + + End Extension_Facts. + + Section Simple_Extension. + + Instance simple_extension: Extension. + Proof. + unshelve econstructor. + - exact ext_least_choice. + - apply ext_least_choice_dec. + - apply ext_least_choice_uni. + Defined. + + Definition P_ := F_ simple_extension. + Definition P_func := F_func simple_extension. + Definition P := F_with simple_extension. + + Definition non_finite e := ~ exhaustible (W e). + + Fact In_Pf k y: In y (P_func k) -> P y. + Proof. + intros H. exists (P_func k), k. + split; [easy| apply F_func_correctness]. + Qed. + + Definition lim_to (f: list nat -> nat -> nat) b := (exists n, forall m, n <= m -> f (P_func m) m = b). + + End Simple_Extension. + + Hypothesis wall_spec: forall e, exists b, lim_to (wall e) b. + + Section Requirements. + + Definition R_simple P e := non_finite e -> ~ (W e #ₚ P). + Definition attend e n := e < n /\ least (ext_pick (P_func n) n) e. + Definition act e n := ~ (P_func n) # W[n] e. + Definition pick_el e x := exists k, attend e k /\ ext_least_choice (P_func k) k x. + + End Requirements. + + Section Requirements_Facts. + + Lemma ext_pick_witness L n e: + ext_pick L n e -> exists x, least (ext_has_wit L n e) x. + Proof. + intros [H1 H2]. + eapply least_ex. intro; eapply ext_has_wit_dec. + trivial. + Qed. + + Lemma W_incl e n m: + n <= m -> forall x, (W[n] e) x -> (W[m] e) x. + Proof. + intros H x [y [H1 H2]]. + exists y. split; lia + easy. + Qed. + + Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : + L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. + Proof. + intros H h1 h2 x Hx1 Hx2. + eapply (H x). now eapply h1. + now apply h2. + Qed. + + Lemma act_always_act e n: act e n -> forall m, n <= m -> act e m . + Proof. + intros H m Hm Hx. apply H. + eapply (intersect_mono Hx). + eapply F_mono; try eapply F_func_correctness; easy. + now eapply W_incl. + Qed. + + Lemma attend_act e k: attend e k -> act e (S k). + Proof. + intros [He H] Hcontr. + edestruct (ext_pick_witness) as [x Hx]. + { destruct H. eapply e0. } + apply (Hcontr x). + assert (P_ (S k) (P_func (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. + exists e. assert (P_func k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + split; first easy. split; first easy. easy. + assert (x = a) as ->. eapply (@extend_uni simple_extension); cbn; eauto. + eauto. exfalso. eapply (H3 x). cbn. + exists e. enough ((P_func k) = (P_func (S k))) as <-. + split; first easy. easy. + assert (F_ simple_extension k (P_func k)) by apply F_func_correctness. + eapply F_uni; eauto. + firstorder. + Qed. + + Lemma act_not_attend e k: act e k -> ~ attend e k. + Proof. now intros h2 [_ [[h _] _]]. Qed. + + Lemma attend_always_act e k: attend e k -> forall m, k < m -> act e m. + Proof. + intros. eapply act_always_act. + apply attend_act. apply H. lia. + Qed. + + Lemma attend_always_not_attend e k: + attend e k -> forall m, k < m -> ~ attend e m. + Proof. + intros H1 m Hm. eapply act_not_attend. + apply (attend_always_act H1 Hm). + Qed. + + Lemma attend_at_most_once e: ~ ~ (exists s, forall s', s < s' -> ~ attend e s'). + Proof. + ccase (exists k, attend e k) as [[w Hw]|H]. + - intros H. eapply H. exists w. + now eapply attend_always_not_attend. + - intros Hk. apply Hk. exists 0. + intros k' Hk' Ha. + apply H. now exists k'. + Qed. + + Lemma attend_uni e k1 k2 : attend e k1 -> attend e k2 -> k1 = k2. + Proof. + intros H1 H2. + specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H1 a b)) as H1'. + specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H2 a b)) as H2'. + enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. + intro Hk. eapply H1'. apply Hk. easy. + intro Hk. eapply H2'. apply Hk. easy. + Qed. + + Lemma pick_el_uni e x1 x2: pick_el e x1 -> pick_el e x2 -> x1 = x2 . + Proof. + intros [k [Ht Hk]] [k' [Ht' Hk']]. + assert (k=k') as <-. eapply attend_uni; eauto. + eapply (@extend_uni simple_extension); cbn; eauto. + Qed. + + End Requirements_Facts. + + Section Compl_P_non_finite. + + Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, pick_el e x /\ e < n. + Proof. + intros [[L [k [Hin Hk]]] Hn]. + dependent induction L. inv Hin. + destruct (Dec (a = x)) as [->|]. + - clear IHL Hin. + destruct (F_pick Hk) as [k' [Hk' [e He]]]. + exists e. split. + exists k'. split; unfold attend. + + destruct He; intuition eauto. enough (P_func k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + exists e. unfold ext_choice. + enough (P_func k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. + destruct He; intuition eauto. + + destruct He. destruct H0. destruct H1. destruct H1. + lia. + - destruct (F_pop Hk) as [m' Hm']. + eapply (IHL m'); eauto. firstorder. + Qed. + + Lemma P_extract_spec n L: + (forall x, In x L -> P x /\ x <= 2 * n) -> + forall x, In x L -> exists c, pick_el c x /\ c < n. + Proof. + intros. induction L. inv H0. + destruct H0 as [-> | Hln]; last apply IHL; eauto. + apply P_meet_spec. now apply H. + Qed. + + Lemma P_pullback_list n L: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> + exists (LC: list nat), NoDup LC /\ length LC = length L /\ + forall c, In c LC -> exists x, pick_el c x /\ In x L /\ c < n. + Proof. + intros HL H1. + induction L. + - exists []; intuition. + - remember (@P_extract_spec n (a::L) H1 a). + assert (In a (a::L)) as H0 by intuition. + apply e in H0 as [c0 [H0 E1]]. + assert (NoDup L) by (inversion HL; intuition). + apply IHL in H as [LC H]. + exists (c0::LC). intuition. + + constructor; last now exact H2. + intros In. inv HL. + apply H4 in In as [y (Hs & h2 & h3)]. + now rewrite (pick_el_uni H0 Hs) in H6. + + cbn. rewrite H. trivial. + + destruct H3 as [->|]. + * exists a; intuition. + * destruct (H4 c H3) as [y Hy]. + exists y; intuition. + + intros y In1. assert (In y (a::L)) by intuition. + now apply H1 in H2. + Qed. + + Definition PredListTo p : list nat -> nat -> Prop + := fun L b => forall x, In x L <-> p x /\ x <= b. + + Lemma NoDupBoundH {L} b: + NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). + Proof. + intros ND H x E. + constructor. + - intros H1 % H. lia. + - exact ND. + Qed. + + Lemma PredNoDupListTo_NNExist p: + forall b, ~~ exists L, PredListTo p L b /\ NoDup L. + Proof. + destruct (F_computable simple_extension) as [f Hf]. + induction b; intros H. + - ccase (p 0) as [H0 | H0]; apply H. + + exists [0]. split; try split. + * intros [E | E]; (try contradiction E). + rewrite <- E. intuition. + * intros E. assert (x = 0) by lia. + rewrite H1. intuition. + * constructor; intuition; constructor. + + exists nil. split; try split. + * contradiction. + * intros E. assert (x = 0) by lia. + rewrite H1 in E. firstorder. + * constructor. + - apply IHb. intros [L H1]. + ccase (p (1 + b)) as [H0 | H0]; apply H. + + exists ((1+ b) :: L). split; try split. + * intros [E | E]; try (rewrite <- E; intuition). + apply H1 in E. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** right. apply H1. intuition. + ** left. lia. + * apply (@NoDupBoundH _ b). + ** apply H1. + ** intros x H3 % H1. lia. + ** lia. + + exists L. split; try split. + * intros E % H1. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** apply H1. intuition. + ** rewrite E in E1. firstorder. + * apply H1. + Qed. + + Lemma P_bounded L n: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. + Proof. + intros ND [LC H] % P_pullback_list; intuition. + rewrite <- H. + assert (incl LC (seq 0 n)). unfold incl. + - intros c [e [x Hx]] % H2. apply in_seq. lia. + - apply pigeonhole_length in H1. + + now rewrite seq_length in H1. + + intros. decide (x1 = x2); tauto. + + exact H0. + Qed. + + Lemma P_Listing: + forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). + Proof. + intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). + intros [L H1]. apply H. exists L; intuition. + apply P_bounded. + - exact H2. + - apply H0. + Qed. + + Lemma complToBound_compl p L b: + PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. + Proof. + intros H x. split. + - intros [H1 H1'] % in_filter_iff. + destruct Dec; cbn in H1'; try congruence. + enough (x <= b). + + firstorder. + + apply in_seq in H1. lia. + - intros [H1 H2]. eapply in_filter_iff. split. + + apply in_seq; lia. + + destruct Dec; cbn; try tauto. exfalso. firstorder. + Qed. + + Lemma compl_P_Listing: + forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L + /\ forall x, In x L -> ~ P x. + Proof. + intros n H. + apply (@P_Listing n). intros [L H1]. + apply H. exists (complToBound L (2*n)). repeat split. + - remember (complToBound_length L (2*n)). lia. + - apply complToBound_NoDup. + - intros x I % (@complToBound_compl P); intuition. + Qed. + + Lemma P_coinfinite : ~ exhaustible (compl P). + Proof. + eapply weakly_unbounded_non_finite. + intros n H. eapply compl_P_Listing with (n := n). + intros (l & ? & ? & H2). + eapply H. + exists (firstn n l). + repeat split. + - rewrite firstn_length. lia. + - now eapply firstn_NoDup. + - intros ? ? % firstn_In. now eapply H2. + Qed. + + End Compl_P_non_finite. + + Section Meet_Requirement. + + Lemma wall_of_wall' e: exists w, forall x, wall e (P_func x) x < w. + Proof. + destruct (wall_spec e) as [v [k Hk]]. + destruct (@extra_bounded (fun k => wall e (P_func k) k) k) as [w Hw]. + exists (S (max v w)). intros x. + destruct (Dec (x < k)). apply Hw in l. lia. + assert (wall e (P_func x) x = v). apply Hk. lia. lia. + Qed. + + Lemma wall_of_wall e: exists w, forall i x, i <= e -> wall i (P_func x) x < w. + Proof. + induction e. + - destruct (wall_of_wall' 0) as [w Hw]. exists w. + intros i x ?. inv H. trivial. + - destruct IHe as [w IHe]. + destruct (wall_of_wall' (S e)) as [w' Hw]. + exists (S (max w w')). intros i x H. + inv H. specialize (Hw x). lia. + specialize (IHe i x H1). lia. + Qed. + + Lemma attend_at_most_once_bound k: + ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). + Proof. + induction k. + - intros H. apply H. exists 42. intros ??. lia. + - intros H. apply IHk. intros [s Hs]; clear IHk. + specialize (@attend_at_most_once k) as Hk. + apply Hk. intros [sk Hsk]; clear Hk. + set (max sk s) as N. + eapply H. exists N. intros e He. + assert (e = k \/ e < k) as [->|Hek] by lia. + intros s' Hs'. eapply Hsk. lia. + intros s' Hs'. eapply Hs; lia. + Qed. + + Lemma non_finite_not_bounded e: + non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ + (forall n, forall i, i <= e -> wall i (P_func n) n < x). + Proof. + intro H. unfold non_finite in H. + intros He. rewrite non_finite_nat in H. + destruct (wall_of_wall e) as [w Hw]. + pose (N := max (2*e + 1) w). specialize (H N). + apply H. intros [m [Hm1 [k Hmk]]]. + apply He. exists k, m. + repeat split. now exists k. + lia. intros n i Hie. specialize (Hw i n Hie). lia. + Qed. + + Lemma ext_pick_attend N e: + e < N -> (exists w, w < e /\ ext_pick (P_func N) N w) -> + (exists w, w < e /\ attend w N). + Proof. + intros HeN [w (Hw1 & Hw2)]. + assert (exists w, ext_pick (P_func N) N w) by now exists w. + eapply least_ex in H; last eauto. + destruct H as [k Hk]. assert (k <= w). + { enough (~ k > w) by lia. intro Hkw. + destruct Hk. rewrite safe_char in H0. + specialize (H0 w Hw2). lia. } + exists k. do 2 (split; first lia). eapply Hk. + Qed. + + Lemma non_finite_attend e: + non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ attend e k) . + Proof. + intros H He. + eapply (non_finite_not_bounded H); intros (b & x & (Hx1 & Hx2 & Hx3)). + eapply (@attend_at_most_once_bound e); intros [s Hs]. + pose (N := S (max (max b s) e)). apply He. + destruct (Dec (ext_intersect_W (P_func N) N e)) as [He'|He']. + - exists N. right. + assert (ext_pick (P_func N) N e). + { split; first easy. exists x. split. + eapply W_incl; last apply Hx1. lia. + split; first lia. eapply Hx3. + } + split. lia. split. easy. + intros w HEw Hw. + assert (exists w, w < e /\ ext_pick (P_func N) N w). + now exists w. eapply ext_pick_attend in H1. + destruct H1 as [g [HEg Hg]]. + eapply (Hs g HEg); last apply Hg. lia. lia. + - exists N. now left. + Qed. + + Lemma ext_intersect_W_intersect k e: + ~ (P_func k # W[k] e) -> W e #ₚ P -> False. + Proof. + unfold ext_intersect_W. + intros H1 H2. apply H1. + intros y Hy1 [w Hy2]. + eapply (H2 y). now exists w. + eapply (In_Pf Hy1). + Qed. + + Lemma P_meet_R_simple : forall e, R_simple P e. + Proof. + intros e He. intros He'. + eapply (non_finite_attend He). + intros [k [H|H]]. + - eapply ext_intersect_W_intersect; eauto. + - eapply attend_act in H. + eapply ext_intersect_W_intersect; eauto. + Qed. + + End Meet_Requirement. + + Section Result. + + Lemma P_semi_decidable : semi_decidable P. + Proof. + apply F_with_semi_decidable. + Qed. + + Theorem P_simple : simple P. + Proof. + unfold simple; repeat split. + - rewrite EA.enum_iff. now apply P_semi_decidable. + - apply P_coinfinite. + - intros [q (Hqenum & Hqinf & Hq)]. + rewrite EA.enum_iff in Hqenum. + destruct (W_spec Hqenum) as [c HWq]. + apply (@P_meet_R_simple c). + intros [l Hqe]; apply Hqinf. + exists l. intros x Hqx. apply (Hqe x). + now rewrite HWq in Hqx. + intros x HWcx HPx. unfold W in HWcx. + rewrite <- (HWq x) in HWcx. apply (Hq x HWcx HPx). + Qed. + + End Result. + + End Assume_WALL. + + Section Instance_of_Wall. + + Definition low_wall (e: nat) (l: list nat) (n: nat) := 42. + Lemma low_wall_spec: forall e, exists b, lim_to low_wall (low_wall e) b. + Proof. intro e. exists 42. exists 0; intuition. Qed. + + Theorem P_simple_low_wall: simple (P low_wall). + Proof. + eapply P_simple. apply low_wall_spec. + Qed. + + End Instance_of_Wall. + +End Assume_EA. + +Require SyntheticComputability.Axioms.EA. + +Lemma EA_correctness: Σ φ, EA φ. +Proof. + Import SyntheticComputability.Axioms.EA.Assume_EA. + exists φ. intros P HP%SyntheticComputability.Axioms.EA.enum_iff. + rewrite W_spec in HP. destruct HP as [c Hc]. + exists c. intros x. unfold W in Hc. + apply Hc. +Qed. + +Section Requirements_Lowness_Correctness. + + Variable P: nat -> Prop. + + Lemma Jump_limit :limit_computable (P´). + Proof. + unfold J. + unfold limit_computable. + admit. + Admitted. +End Requirements_Lowness_Correctness. diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index 09ba2d9..7c28898 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -12,8 +12,7 @@ Notation "'Σ' x .. y , p" := format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") : type_scope. -Section Construction. - Record Extension := + Class Extension := { extendP: list nat -> nat -> nat -> Prop; extend_dec: forall l x, (Σ y, extendP l x y) + (forall y, ~ extendP l x y); @@ -22,10 +21,14 @@ Section Construction. Inductive F_ (E: Extension) : nat -> list nat -> Prop := | Base : F_ E O [] - | ExtendS n l a : F_ E n l -> extendP E l n a -> F_ E (S n) (a::l) - | ExtendN n l : F_ E n l -> (forall a, ~ extendP E l n a) -> F_ E (S n) l. + | ExtendS n l a : F_ E n l -> extendP l n a -> F_ E (S n) (a::l) + | ExtendN n l : F_ E n l -> (forall a, ~ extendP l n a) -> F_ E (S n) l. + +Section Construction. + + Variable E: Extension. - Lemma F_uni E : forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . + Lemma F_uni: forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . Proof. intros n l1 l2. dependent induction n. @@ -40,7 +43,7 @@ Section Construction. now apply IHn. Qed. - Lemma F_mono E n m l1 l2: F_ E n l1 -> F_ E m l2 -> n <= m -> incl l1 l2. + Lemma F_mono n m l1 l2: F_ E n l1 -> F_ E m l2 -> n <= m -> incl l1 l2. Proof. intros H1 H2 HE. revert H1 H2; induction HE in l2 |-*; intros H1 H2. @@ -49,61 +52,61 @@ Section Construction. specialize (IHHE l H1 H0). eauto. Qed. - Lemma F_pop E n x l: F_ E n (x::l) -> exists m, F_ E m l. + Lemma F_pop n x l: F_ E n (x::l) -> exists m, F_ E m l. Proof. intros H. dependent induction H. - now exists n. - eapply IHF_. eauto. Qed. - Lemma F_pick E n x l: F_ E n (x::l) -> exists m, F_ E m l /\ extendP E l m x. + Lemma F_pick n x l: F_ E n (x::l) -> exists m, F_ E m l /\ extendP l m x. Proof. intros H. dependent induction H. - exists n; eauto. - eapply IHF_; eauto. Qed. - Lemma F_computable E : Σ f: nat -> list nat, + Lemma F_computable : Σ f: nat -> list nat, forall n, F_ E n (f n) /\ length (f n) <= n. Proof. set (F := fix f x := match x with | O => [] - | S x => match (extend_dec E) (f x) x with + | S x => match extend_dec (f x) x with | inr _ => f x | inl aH => (projT1 aH) :: (f x) end end). exists F. induction n; simpl. - split. constructor. easy. - - destruct (extend_dec E (F n) n); split. + - destruct (extend_dec (F n) n); split. + eapply ExtendS; first apply IHn. now destruct s. + cbn. destruct IHn. lia. + now eapply ExtendN. + destruct IHn. lia. Qed. - Definition F_func E := projT1 (F_computable E). - Lemma F_func_correctness {E}: forall n, F_ E n (F_func E n). + Definition F_func := projT1 F_computable. + Lemma F_func_correctness: forall n, F_ E n (F_func n). Proof. intros n; unfold F_func. - destruct (F_computable E) as [f H]. + destruct F_computable as [f H]. now destruct (H n). Qed. - Lemma F_func_correctness' {E}: forall n, length (F_func E n) <= n. + Lemma F_func_correctness': forall n, length (F_func n) <= n. Proof. intros n; unfold F_func. - destruct (F_computable E) as [f H]. + destruct F_computable as [f H]. now destruct (H n). Qed. - Definition F_with E x := exists l n, In x l /\ (F_ E n l). + Definition F_with x := exists l n, In x l /\ (F_ E n l). - Lemma F_with_semi_decidable E: semi_decidable (F_with E). + Lemma F_with_semi_decidable: semi_decidable F_with. Proof. unfold semi_decidable, semi_decider. - destruct (F_computable E) as [f Hf ]. + destruct F_computable as [f Hf ]. exists (fun x n => (Dec (In x (f n)))). intros x. split. - intros (l & n & Hxl & Hl). @@ -374,4 +377,17 @@ Section LeastWitness. Lemma dec_neg_dec P: dec P -> dec (~ P). Proof. intros [H|H]. right; eauto. now left. Qed. + Lemma forall_bounded_dec P e: + (forall x, dec (P x)) -> dec(forall i, i <= e -> P i). + Proof. + intros H. + induction e. destruct (H 0); [left|right]; intros. + now inv H0. intros H'. apply n, H'. lia. + destruct IHe as [H1|H1]. + destruct (H (S e)); [left|right]; intros. + inv H0; first easy. now apply H1. + intros H0. apply n. apply H0. easy. + right. intro H'. apply H1. intros. apply H'. lia. + Qed. + End LeastWitness. \ No newline at end of file From a04cc7ec0a784c2ae066b72df61212533e73838a Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 25 Mar 2024 14:36:14 +0100 Subject: [PATCH 20/54] refactor --- .../low_simple_extension.v | 701 ------------------ theories/ReducibilityDegrees/low_wall.v | 20 + .../ReducibilityDegrees/simple_extension.v | 347 +++++---- 3 files changed, 213 insertions(+), 855 deletions(-) delete mode 100644 theories/ReducibilityDegrees/low_simple_extension.v create mode 100644 theories/ReducibilityDegrees/low_wall.v diff --git a/theories/ReducibilityDegrees/low_simple_extension.v b/theories/ReducibilityDegrees/low_simple_extension.v deleted file mode 100644 index 472ccdf..0000000 --- a/theories/ReducibilityDegrees/low_simple_extension.v +++ /dev/null @@ -1,701 +0,0 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. -Require Import SyntheticComputability.Synthetic.DecidabilityFacts. -Require Export SyntheticComputability.Shared.FinitenessFacts. -Require Export SyntheticComputability.Shared.Pigeonhole. -Require Export SyntheticComputability.Shared.ListAutomation. -Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. -Require Import SyntheticComputability.ReducibilityDegrees.priority_method. -Import ListNotations. - -Section ComplToBound. - Definition complToBound L b : list nat - := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). - - Lemma complToBound_Bound L b : - forall x, In x (complToBound L b) -> x <= b. - Proof. - intros x [H % in_seq ?] % in_filter_iff. lia. - Qed. - Lemma filter_length {X} f (l : list X) : - length l = length (filter f l) + length (filter (fun x => (negb (f x))) l). - Proof. - induction l; cbn. - - reflexivity. - - destruct f; cbn; lia. - Qed. - Lemma filter_NoDup {X} f (l : list X) : - NoDup l -> NoDup (filter f l). - Proof. - induction 1; cbn. - - econstructor. - - destruct f; eauto. econstructor; auto. - intros ? % in_filter_iff. firstorder. - Qed. - Lemma complToBound_length L b: - length (complToBound L b) + length L >= S b. - Proof. - rewrite <- (seq_length (S b) 0). - erewrite filter_length with (l := seq 0 (S b)). - unfold complToBound. - eapply Plus.plus_le_compat_l_stt. - generalize (seq_NoDup (S b) 0). - generalize (seq 0 (S b)). clear. - intros. erewrite filter_ext with (g := fun x => Dec (In x L)). - 2:{ intros a. destruct Dec; cbn; destruct Dec; firstorder congruence. } - eapply NoDup_incl_length. now eapply filter_NoDup. - clear. induction l; cbn. - - firstorder. - - destruct Dec; cbn. 2: eauto. - intros ? [-> | ]; eauto. - Qed. - Lemma complToBound_NoDup L b: - NoDup (complToBound L b). - Proof. - eapply filter_NoDup, seq_NoDup. - Qed. - Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) -> In x l. - Proof. - induction n in x, l |- *; destruct l; cbn; firstorder. - Qed. - - Lemma firstn_NoDup {X} (l : list X) n : NoDup l -> NoDup (firstn n l). - Proof. - induction 1 in n |- *; destruct n; cbn; try now econstructor. - econstructor; eauto. - now intros ? % firstn_In. - Qed. -End ComplToBound. - -Section Assume_EA. - - Variable φ: nat -> nat -> option nat. - Definition EA := forall P, - semi_decidable P -> exists e, forall x, P x <-> exists n, φ e n = Some x. - Hypothesis EA: EA. - - Definition W_ n e x := φ n e = Some x. - Definition W e x := exists n, W_ e n x. - - Lemma W_spec: forall P, semi_decidable P -> exists e, forall x, P x <-> W e x. - Proof. intros P [e He]%EA. exists e; intros x; now rewrite He. Qed. - - Notation "'W[' s ']' e" := (fun x => exists n, n <= s /\ W_ e n x) (at level 30). - - Section EA_dec. - - Lemma W_dec e: forall x n, dec (W_ e n x). - Proof. - intros x n. - destruct (φ e n) eqn: E. - - destruct (Dec (x = n0)) as [E'|E']. - left. now subst. right. intros H. congruence. - - right. intros H. congruence. - Qed. - - Lemma W_bounded_dec e : forall x s, dec ((W[s] e) x). - Proof. - intros x s. cbn. eapply exists_bounded_dec. - intro; apply W_dec. - Qed. - - Lemma W_bounded_bounded e s: bounded (W[s] e). - Proof. - unfold bounded. - induction s. - - destruct (φ e 0) as [w|] eqn: E. - exists (S w). intros x [n [Hn1 Hn2]]. - inv Hn1. unfold W_ in Hn2. rewrite E in Hn2. - injection Hn2. lia. - exists 42. intros x [n [Hn1 Hn2]]. - inv Hn1. unfold W_ in Hn2. - rewrite E in Hn2. congruence. - - destruct IHs as [N HN]. - destruct (φ e (S s)) as [w|] eqn: E. - exists (max (S w) (S N)). intros x [n [Hn1 Hn2]]. - inv Hn1. unfold W_ in Hn2. - rewrite E in Hn2. injection Hn2. lia. - eapply Nat.lt_trans. eapply HN. exists n; split; easy. lia. - exists N. intros x [n [Hn1 Hn2]]. - inv Hn1. unfold W_ in Hn2. - rewrite E in Hn2. congruence. - eapply HN. exists n. split; eauto. - Qed. - - End EA_dec. - - Definition disj_list_pred {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. - Definition disj_pred {X} (A B: X -> Prop) := forall x, A x -> B x -> False. - Notation "A # B" := (disj_list_pred A B) (at level 30). - Notation "A #ₚ B" := (disj_pred A B) (at level 30). - - Lemma extra_bounded f m: Σ b, forall n, n < m -> f n < b. - Proof. - induction m. - - exists 42. intros. inv H. - - destruct IHm as [b' Hb']. - exists (S (max b' (f m))). - intros ? H. inv H. lia. specialize (Hb' n H1). lia. - Qed. - - Section Assume_WALL. - Variable wall: nat -> list nat -> nat -> nat. - - Section Extension. - - Definition ext_intersect_W L n e := L # W[n] e. - Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. - Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. - Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. - Definition ext_least_choice L n x := exists e, ext_choice L n e x. - - End Extension. - - Section Extension_Facts. - - #[export]Instance ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). - Proof. - apply BaseLists.list_forall_dec. - intro x. eapply dec_neg_dec, exists_bounded_dec; eauto. - Qed. - - #[export] Instance ext_wall L n e x: dec (forall i, i <= e -> wall i L n < x). - Proof. eapply forall_bounded_dec; eauto. Qed. - - #[export]Instance ext_has_wit_dec L n e x : dec (ext_has_wit L n e x). - Proof. apply and_dec; first apply exists_bounded_dec; eauto. Qed. - - #[export]Instance ext_has_wit_exists_dec L n e : dec (exists x, ext_has_wit L n e x). - Proof. - unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. - intro x; eapply W_bounded_dec; eauto. eauto. - Qed. - - #[export]Instance ext_pick_dec L n e : dec (ext_pick L n e). - Proof. - eapply and_dec; first apply ext_intersect_W_dec. - unfold ext_has_wit. - eapply bounded_dec; last apply W_bounded_bounded. - intros x. eapply exists_bounded_dec. - intro; apply W_dec. eauto. - Qed. - - #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). - Proof. - eapply exists_bounded_dec'. intro x. - eapply least_dec. intros y. - eapply ext_pick_dec. - Qed. - - Fact ext_least_choice_dec L n: - (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). - Proof. - unfold ext_least_choice. - destruct (ext_pick_exists_dec L n) as [H|H]. - - left. apply ewo_nat; first last. - destruct H as [e (h1 & [(h4 & h2) h3])]. - eapply least_ex in h2. destruct h2 as [k h6]. - exists k, e. split; first easy; split. - 2: { eapply h6. } - split; last apply h3. split; first apply h4. - destruct h6. now exists k. - eapply ext_has_wit_dec. - intro x. unfold ext_choice. eapply exists_bounded_dec'. - intros x'. eapply and_dec. - apply least_dec. eapply ext_pick_dec. - apply least_dec. eapply ext_has_wit_dec. - - right. intros x [e (h1 & h2 & _)]. apply H. - exists e. split; eauto. - Qed. - - Fact ext_least_choice_uni l x y1 y2: - ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. - Proof. - intros [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. - assert (x0 = x1) as ->. eapply least_unique; eauto. - eapply least_unique; eauto. - Qed. - - End Extension_Facts. - - Section Simple_Extension. - - Instance simple_extension: Extension. - Proof. - unshelve econstructor. - - exact ext_least_choice. - - apply ext_least_choice_dec. - - apply ext_least_choice_uni. - Defined. - - Definition P_ := F_ simple_extension. - Definition P_func := F_func simple_extension. - Definition P := F_with simple_extension. - - Definition non_finite e := ~ exhaustible (W e). - - Fact In_Pf k y: In y (P_func k) -> P y. - Proof. - intros H. exists (P_func k), k. - split; [easy| apply F_func_correctness]. - Qed. - - Definition lim_to (f: list nat -> nat -> nat) b := (exists n, forall m, n <= m -> f (P_func m) m = b). - - End Simple_Extension. - - Hypothesis wall_spec: forall e, exists b, lim_to (wall e) b. - - Section Requirements. - - Definition R_simple P e := non_finite e -> ~ (W e #ₚ P). - Definition attend e n := e < n /\ least (ext_pick (P_func n) n) e. - Definition act e n := ~ (P_func n) # W[n] e. - Definition pick_el e x := exists k, attend e k /\ ext_least_choice (P_func k) k x. - - End Requirements. - - Section Requirements_Facts. - - Lemma ext_pick_witness L n e: - ext_pick L n e -> exists x, least (ext_has_wit L n e) x. - Proof. - intros [H1 H2]. - eapply least_ex. intro; eapply ext_has_wit_dec. - trivial. - Qed. - - Lemma W_incl e n m: - n <= m -> forall x, (W[n] e) x -> (W[m] e) x. - Proof. - intros H x [y [H1 H2]]. - exists y. split; lia + easy. - Qed. - - Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : - L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. - Proof. - intros H h1 h2 x Hx1 Hx2. - eapply (H x). now eapply h1. - now apply h2. - Qed. - - Lemma act_always_act e n: act e n -> forall m, n <= m -> act e m . - Proof. - intros H m Hm Hx. apply H. - eapply (intersect_mono Hx). - eapply F_mono; try eapply F_func_correctness; easy. - now eapply W_incl. - Qed. - - Lemma attend_act e k: attend e k -> act e (S k). - Proof. - intros [He H] Hcontr. - edestruct (ext_pick_witness) as [x Hx]. - { destruct H. eapply e0. } - apply (Hcontr x). - assert (P_ (S k) (P_func (S k))) by apply F_func_correctness. - inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. - exists e. assert (P_func k = l) as <-. - eapply F_uni. apply F_func_correctness. apply H3. - split; first easy. split; first easy. easy. - assert (x = a) as ->. eapply (@extend_uni simple_extension); cbn; eauto. - eauto. exfalso. eapply (H3 x). cbn. - exists e. enough ((P_func k) = (P_func (S k))) as <-. - split; first easy. easy. - assert (F_ simple_extension k (P_func k)) by apply F_func_correctness. - eapply F_uni; eauto. - firstorder. - Qed. - - Lemma act_not_attend e k: act e k -> ~ attend e k. - Proof. now intros h2 [_ [[h _] _]]. Qed. - - Lemma attend_always_act e k: attend e k -> forall m, k < m -> act e m. - Proof. - intros. eapply act_always_act. - apply attend_act. apply H. lia. - Qed. - - Lemma attend_always_not_attend e k: - attend e k -> forall m, k < m -> ~ attend e m. - Proof. - intros H1 m Hm. eapply act_not_attend. - apply (attend_always_act H1 Hm). - Qed. - - Lemma attend_at_most_once e: ~ ~ (exists s, forall s', s < s' -> ~ attend e s'). - Proof. - ccase (exists k, attend e k) as [[w Hw]|H]. - - intros H. eapply H. exists w. - now eapply attend_always_not_attend. - - intros Hk. apply Hk. exists 0. - intros k' Hk' Ha. - apply H. now exists k'. - Qed. - - Lemma attend_uni e k1 k2 : attend e k1 -> attend e k2 -> k1 = k2. - Proof. - intros H1 H2. - specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H1 a b)) as H1'. - specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H2 a b)) as H2'. - enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. - intro Hk. eapply H1'. apply Hk. easy. - intro Hk. eapply H2'. apply Hk. easy. - Qed. - - Lemma pick_el_uni e x1 x2: pick_el e x1 -> pick_el e x2 -> x1 = x2 . - Proof. - intros [k [Ht Hk]] [k' [Ht' Hk']]. - assert (k=k') as <-. eapply attend_uni; eauto. - eapply (@extend_uni simple_extension); cbn; eauto. - Qed. - - End Requirements_Facts. - - Section Compl_P_non_finite. - - Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, pick_el e x /\ e < n. - Proof. - intros [[L [k [Hin Hk]]] Hn]. - dependent induction L. inv Hin. - destruct (Dec (a = x)) as [->|]. - - clear IHL Hin. - destruct (F_pick Hk) as [k' [Hk' [e He]]]. - exists e. split. - exists k'. split; unfold attend. - + destruct He; intuition eauto. enough (P_func k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + exists e. unfold ext_choice. - enough (P_func k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. - destruct He; intuition eauto. - + destruct He. destruct H0. destruct H1. destruct H1. - lia. - - destruct (F_pop Hk) as [m' Hm']. - eapply (IHL m'); eauto. firstorder. - Qed. - - Lemma P_extract_spec n L: - (forall x, In x L -> P x /\ x <= 2 * n) -> - forall x, In x L -> exists c, pick_el c x /\ c < n. - Proof. - intros. induction L. inv H0. - destruct H0 as [-> | Hln]; last apply IHL; eauto. - apply P_meet_spec. now apply H. - Qed. - - Lemma P_pullback_list n L: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> - exists (LC: list nat), NoDup LC /\ length LC = length L /\ - forall c, In c LC -> exists x, pick_el c x /\ In x L /\ c < n. - Proof. - intros HL H1. - induction L. - - exists []; intuition. - - remember (@P_extract_spec n (a::L) H1 a). - assert (In a (a::L)) as H0 by intuition. - apply e in H0 as [c0 [H0 E1]]. - assert (NoDup L) by (inversion HL; intuition). - apply IHL in H as [LC H]. - exists (c0::LC). intuition. - + constructor; last now exact H2. - intros In. inv HL. - apply H4 in In as [y (Hs & h2 & h3)]. - now rewrite (pick_el_uni H0 Hs) in H6. - + cbn. rewrite H. trivial. - + destruct H3 as [->|]. - * exists a; intuition. - * destruct (H4 c H3) as [y Hy]. - exists y; intuition. - + intros y In1. assert (In y (a::L)) by intuition. - now apply H1 in H2. - Qed. - - Definition PredListTo p : list nat -> nat -> Prop - := fun L b => forall x, In x L <-> p x /\ x <= b. - - Lemma NoDupBoundH {L} b: - NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). - Proof. - intros ND H x E. - constructor. - - intros H1 % H. lia. - - exact ND. - Qed. - - Lemma PredNoDupListTo_NNExist p: - forall b, ~~ exists L, PredListTo p L b /\ NoDup L. - Proof. - destruct (F_computable simple_extension) as [f Hf]. - induction b; intros H. - - ccase (p 0) as [H0 | H0]; apply H. - + exists [0]. split; try split. - * intros [E | E]; (try contradiction E). - rewrite <- E. intuition. - * intros E. assert (x = 0) by lia. - rewrite H1. intuition. - * constructor; intuition; constructor. - + exists nil. split; try split. - * contradiction. - * intros E. assert (x = 0) by lia. - rewrite H1 in E. firstorder. - * constructor. - - apply IHb. intros [L H1]. - ccase (p (1 + b)) as [H0 | H0]; apply H. - + exists ((1+ b) :: L). split; try split. - * intros [E | E]; try (rewrite <- E; intuition). - apply H1 in E. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** right. apply H1. intuition. - ** left. lia. - * apply (@NoDupBoundH _ b). - ** apply H1. - ** intros x H3 % H1. lia. - ** lia. - + exists L. split; try split. - * intros E % H1. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** apply H1. intuition. - ** rewrite E in E1. firstorder. - * apply H1. - Qed. - - Lemma P_bounded L n: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. - Proof. - intros ND [LC H] % P_pullback_list; intuition. - rewrite <- H. - assert (incl LC (seq 0 n)). unfold incl. - - intros c [e [x Hx]] % H2. apply in_seq. lia. - - apply pigeonhole_length in H1. - + now rewrite seq_length in H1. - + intros. decide (x1 = x2); tauto. - + exact H0. - Qed. - - Lemma P_Listing: - forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). - Proof. - intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). - intros [L H1]. apply H. exists L; intuition. - apply P_bounded. - - exact H2. - - apply H0. - Qed. - - Lemma complToBound_compl p L b: - PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. - Proof. - intros H x. split. - - intros [H1 H1'] % in_filter_iff. - destruct Dec; cbn in H1'; try congruence. - enough (x <= b). - + firstorder. - + apply in_seq in H1. lia. - - intros [H1 H2]. eapply in_filter_iff. split. - + apply in_seq; lia. - + destruct Dec; cbn; try tauto. exfalso. firstorder. - Qed. - - Lemma compl_P_Listing: - forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L - /\ forall x, In x L -> ~ P x. - Proof. - intros n H. - apply (@P_Listing n). intros [L H1]. - apply H. exists (complToBound L (2*n)). repeat split. - - remember (complToBound_length L (2*n)). lia. - - apply complToBound_NoDup. - - intros x I % (@complToBound_compl P); intuition. - Qed. - - Lemma P_coinfinite : ~ exhaustible (compl P). - Proof. - eapply weakly_unbounded_non_finite. - intros n H. eapply compl_P_Listing with (n := n). - intros (l & ? & ? & H2). - eapply H. - exists (firstn n l). - repeat split. - - rewrite firstn_length. lia. - - now eapply firstn_NoDup. - - intros ? ? % firstn_In. now eapply H2. - Qed. - - End Compl_P_non_finite. - - Section Meet_Requirement. - - Lemma wall_of_wall' e: exists w, forall x, wall e (P_func x) x < w. - Proof. - destruct (wall_spec e) as [v [k Hk]]. - destruct (@extra_bounded (fun k => wall e (P_func k) k) k) as [w Hw]. - exists (S (max v w)). intros x. - destruct (Dec (x < k)). apply Hw in l. lia. - assert (wall e (P_func x) x = v). apply Hk. lia. lia. - Qed. - - Lemma wall_of_wall e: exists w, forall i x, i <= e -> wall i (P_func x) x < w. - Proof. - induction e. - - destruct (wall_of_wall' 0) as [w Hw]. exists w. - intros i x ?. inv H. trivial. - - destruct IHe as [w IHe]. - destruct (wall_of_wall' (S e)) as [w' Hw]. - exists (S (max w w')). intros i x H. - inv H. specialize (Hw x). lia. - specialize (IHe i x H1). lia. - Qed. - - Lemma attend_at_most_once_bound k: - ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). - Proof. - induction k. - - intros H. apply H. exists 42. intros ??. lia. - - intros H. apply IHk. intros [s Hs]; clear IHk. - specialize (@attend_at_most_once k) as Hk. - apply Hk. intros [sk Hsk]; clear Hk. - set (max sk s) as N. - eapply H. exists N. intros e He. - assert (e = k \/ e < k) as [->|Hek] by lia. - intros s' Hs'. eapply Hsk. lia. - intros s' Hs'. eapply Hs; lia. - Qed. - - Lemma non_finite_not_bounded e: - non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ - (forall n, forall i, i <= e -> wall i (P_func n) n < x). - Proof. - intro H. unfold non_finite in H. - intros He. rewrite non_finite_nat in H. - destruct (wall_of_wall e) as [w Hw]. - pose (N := max (2*e + 1) w). specialize (H N). - apply H. intros [m [Hm1 [k Hmk]]]. - apply He. exists k, m. - repeat split. now exists k. - lia. intros n i Hie. specialize (Hw i n Hie). lia. - Qed. - - Lemma ext_pick_attend N e: - e < N -> (exists w, w < e /\ ext_pick (P_func N) N w) -> - (exists w, w < e /\ attend w N). - Proof. - intros HeN [w (Hw1 & Hw2)]. - assert (exists w, ext_pick (P_func N) N w) by now exists w. - eapply least_ex in H; last eauto. - destruct H as [k Hk]. assert (k <= w). - { enough (~ k > w) by lia. intro Hkw. - destruct Hk. rewrite safe_char in H0. - specialize (H0 w Hw2). lia. } - exists k. do 2 (split; first lia). eapply Hk. - Qed. - - Lemma non_finite_attend e: - non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ attend e k) . - Proof. - intros H He. - eapply (non_finite_not_bounded H); intros (b & x & (Hx1 & Hx2 & Hx3)). - eapply (@attend_at_most_once_bound e); intros [s Hs]. - pose (N := S (max (max b s) e)). apply He. - destruct (Dec (ext_intersect_W (P_func N) N e)) as [He'|He']. - - exists N. right. - assert (ext_pick (P_func N) N e). - { split; first easy. exists x. split. - eapply W_incl; last apply Hx1. lia. - split; first lia. eapply Hx3. - } - split. lia. split. easy. - intros w HEw Hw. - assert (exists w, w < e /\ ext_pick (P_func N) N w). - now exists w. eapply ext_pick_attend in H1. - destruct H1 as [g [HEg Hg]]. - eapply (Hs g HEg); last apply Hg. lia. lia. - - exists N. now left. - Qed. - - Lemma ext_intersect_W_intersect k e: - ~ (P_func k # W[k] e) -> W e #ₚ P -> False. - Proof. - unfold ext_intersect_W. - intros H1 H2. apply H1. - intros y Hy1 [w Hy2]. - eapply (H2 y). now exists w. - eapply (In_Pf Hy1). - Qed. - - Lemma P_meet_R_simple : forall e, R_simple P e. - Proof. - intros e He. intros He'. - eapply (non_finite_attend He). - intros [k [H|H]]. - - eapply ext_intersect_W_intersect; eauto. - - eapply attend_act in H. - eapply ext_intersect_W_intersect; eauto. - Qed. - - End Meet_Requirement. - - Section Result. - - Lemma P_semi_decidable : semi_decidable P. - Proof. - apply F_with_semi_decidable. - Qed. - - Theorem P_simple : simple P. - Proof. - unfold simple; repeat split. - - rewrite EA.enum_iff. now apply P_semi_decidable. - - apply P_coinfinite. - - intros [q (Hqenum & Hqinf & Hq)]. - rewrite EA.enum_iff in Hqenum. - destruct (W_spec Hqenum) as [c HWq]. - apply (@P_meet_R_simple c). - intros [l Hqe]; apply Hqinf. - exists l. intros x Hqx. apply (Hqe x). - now rewrite HWq in Hqx. - intros x HWcx HPx. unfold W in HWcx. - rewrite <- (HWq x) in HWcx. apply (Hq x HWcx HPx). - Qed. - - End Result. - - End Assume_WALL. - - Section Instance_of_Wall. - - Definition low_wall (e: nat) (l: list nat) (n: nat) := 42. - Lemma low_wall_spec: forall e, exists b, lim_to low_wall (low_wall e) b. - Proof. intro e. exists 42. exists 0; intuition. Qed. - - Theorem P_simple_low_wall: simple (P low_wall). - Proof. - eapply P_simple. apply low_wall_spec. - Qed. - - End Instance_of_Wall. - -End Assume_EA. - -Require SyntheticComputability.Axioms.EA. - -Lemma EA_correctness: Σ φ, EA φ. -Proof. - Import SyntheticComputability.Axioms.EA.Assume_EA. - exists φ. intros P HP%SyntheticComputability.Axioms.EA.enum_iff. - rewrite W_spec in HP. destruct HP as [c Hc]. - exists c. intros x. unfold W in Hc. - apply Hc. -Qed. - -Section Requirements_Lowness_Correctness. - - Variable P: nat -> Prop. - - Lemma Jump_limit :limit_computable (P´). - Proof. - unfold J. - unfold limit_computable. - admit. - Admitted. -End Requirements_Lowness_Correctness. diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v new file mode 100644 index 0000000..022846b --- /dev/null +++ b/theories/ReducibilityDegrees/low_wall.v @@ -0,0 +1,20 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +Require Import SyntheticComputability.Synthetic.DecidabilityFacts. +Require Export SyntheticComputability.Shared.FinitenessFacts. +Require Export SyntheticComputability.Shared.Pigeonhole. +Require Export SyntheticComputability.Shared.ListAutomation. +Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. +Require Import SyntheticComputability.ReducibilityDegrees.priority_method. +Require Import SyntheticComputability.ReducibilityDegrees.simple_extension. + +Section Requirements_Lowness_Correctness. + + Variable P: nat -> Prop. + + Lemma Jump_limit :limit_computable (P´). + Proof. + unfold J. + unfold limit_computable. + admit. + Admitted. +End Requirements_Lowness_Correctness. diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index b97afa2..3dcfa6b 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -128,13 +128,24 @@ Section Assume_EA. Notation "A # B" := (disj_list_pred A B) (at level 30). Notation "A #ₚ B" := (disj_pred A B) (at level 30). + Lemma extra_bounded f m: Σ b, forall n, n < m -> f n < b. + Proof. + induction m. + - exists 42. intros. inv H. + - destruct IHm as [b' Hb']. + exists (S (max b' (f m))). + intros ? H. inv H. lia. specialize (Hb' n H1). lia. + Qed. + + Section Assume_WALL. + Variable wall: nat -> list nat -> nat -> nat. Section Extension. Definition ext_intersect_W L n e := L # W[n] e. - Definition ext_has_wit n e x := (W[n] e) x /\ 2 * e < x. - Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit n e x. - Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit n e) x. + Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. + Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. + Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. Definition ext_least_choice L n x := exists e, ext_choice L n e x. End Extension. @@ -144,22 +155,19 @@ Section Assume_EA. #[export]Instance ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). Proof. apply BaseLists.list_forall_dec. - intro x. eapply dec_neg_dec, exists_bounded_dec. - intros y. apply W_dec. - Qed. + intro x. eapply dec_neg_dec, exists_bounded_dec; eauto. + Qed. - #[export]Instance ext_has_wit_dec n e x : dec (ext_has_wit n e x). - Proof. - unfold ext_has_wit. apply and_dec. - apply exists_bounded_dec. intro; apply W_dec. - apply Pigeonhole.dec_lt. - Qed. + #[export] Instance ext_wall L n e x: dec (forall i, i <= e -> wall i L n < x). + Proof. eapply forall_bounded_dec; eauto. Qed. + + #[export]Instance ext_has_wit_dec L n e x : dec (ext_has_wit L n e x). + Proof. apply and_dec; first apply exists_bounded_dec; eauto. Qed. - #[export]Instance ext_has_wit_exists_dec n e : dec (exists x, ext_has_wit n e x). + #[export]Instance ext_has_wit_exists_dec L n e : dec (exists x, ext_has_wit L n e x). Proof. unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. - intro x; eapply W_bounded_dec. - intro x; eapply lt_dec. + intro x; eapply W_bounded_dec; eauto. eauto. Qed. #[export]Instance ext_pick_dec L n e : dec (ext_pick L n e). @@ -168,12 +176,11 @@ Section Assume_EA. unfold ext_has_wit. eapply bounded_dec; last apply W_bounded_bounded. intros x. eapply exists_bounded_dec. - intro; apply W_dec. - apply lt_dec. + intro; apply W_dec. eauto. Qed. #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). - Proof. + Proof. eapply exists_bounded_dec'. intro x. eapply least_dec. intros y. eapply ext_pick_dec. @@ -200,10 +207,9 @@ Section Assume_EA. exists e. split; eauto. Qed. - Fact ext_least_choice_uni l x y1 y2: - ext_least_choice l x y1 -> ext_least_choice l x y2 -> y1 = y2. + Fact ext_least_choice_uni l x: unique (ext_least_choice l x). Proof. - intros [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. + intros y1 y2 [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. assert (x0 = x1) as ->. eapply least_unique; eauto. eapply least_unique; eauto. Qed. @@ -212,41 +218,45 @@ Section Assume_EA. Section Simple_Extension. - Definition simple_extendsion: Extension. - Proof. + Instance simple_extension: Extension. + Proof. unshelve econstructor. - exact ext_least_choice. - apply ext_least_choice_dec. - apply ext_least_choice_uni. - Defined. + Defined. - Definition P_ := F_ simple_extendsion. - Definition Pf_ := F_func simple_extendsion. - Definition P := F_with simple_extendsion. + Definition P_ := F_ simple_extension. + Definition P_func := F_func simple_extension. + Definition P := F_with simple_extension. Definition non_finite e := ~ exhaustible (W e). - Fact In_Pf k y: In y (Pf_ k) -> P y . + Fact In_Pf k y: In y (P_func k) -> P y. Proof. - intros H. exists (Pf_ k), k. + intros H. exists (P_func k), k. split; [easy| apply F_func_correctness]. Qed. + Definition lim_to (f: list nat -> nat -> nat) b := (exists n, forall m, n <= m -> f (P_func m) m = b). + End Simple_Extension. + Hypothesis wall_spec: forall e, exists b, lim_to (wall e) b. + Section Requirements. Definition R_simple P e := non_finite e -> ~ (W e #ₚ P). - Definition attention e n := e < n /\ least (ext_pick (Pf_ n) n) e. - Definition active e n := ~ (Pf_ n) # W[n] e. - Definition pick_el e x := exists k, attention e k /\ ext_least_choice (Pf_ k) k x. + Definition attend e n := e < n /\ least (ext_pick (P_func n) n) e. + Definition act e n := ~ (P_func n) # W[n] e. + Definition pick_el e x := exists k, attend e k /\ ext_least_choice (P_func k) k x. End Requirements. Section Requirements_Facts. Lemma ext_pick_witness L n e: - ext_pick L n e -> exists x, least (ext_has_wit n e) x. + ext_pick L n e -> exists x, least (ext_has_wit L n e) x. Proof. intros [H1 H2]. eapply least_ex. intro; eapply ext_has_wit_dec. @@ -268,7 +278,7 @@ Section Assume_EA. now apply h2. Qed. - Lemma active_always_active e n: active e n -> forall m, n <= m -> active e m . + Lemma act_always_act e n: act e n -> forall m, n <= m -> act e m . Proof. intros H m Hm Hx. apply H. eapply (intersect_mono Hx). @@ -276,67 +286,67 @@ Section Assume_EA. now eapply W_incl. Qed. - Lemma attention_active e k: attention e k -> active e (S k). + Lemma attend_act e k: attend e k -> act e (S k). Proof. intros [He H] Hcontr. edestruct (ext_pick_witness) as [x Hx]. { destruct H. eapply e0. } apply (Hcontr x). - assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. + assert (P_ (S k) (P_func (S k))) by apply F_func_correctness. inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. - exists e. assert (Pf_ k = l) as <-. + exists e. assert (P_func k = l) as <-. eapply F_uni. apply F_func_correctness. apply H3. split; first easy. split; first easy. easy. - assert (x = a) as ->. eapply (@extend_uni simple_extendsion); cbn; eauto. + assert (x = a) as ->. eapply (@extend_uni simple_extension); cbn; eauto. eauto. exfalso. eapply (H3 x). cbn. - exists e. enough ((Pf_ k) = (Pf_ (S k))) as <-. + exists e. enough ((P_func k) = (P_func (S k))) as <-. split; first easy. easy. - assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. + assert (F_ simple_extension k (P_func k)) by apply F_func_correctness. eapply F_uni; eauto. firstorder. Qed. - Lemma active_not_attention e k: active e k -> ~ attention e k. + Lemma act_not_attend e k: act e k -> ~ attend e k. Proof. now intros h2 [_ [[h _] _]]. Qed. - Lemma attention_always_active e k: attention e k -> forall m, k < m -> active e m. + Lemma attend_always_act e k: attend e k -> forall m, k < m -> act e m. Proof. - intros. eapply active_always_active. - apply attention_active. apply H. lia. + intros. eapply act_always_act. + apply attend_act. apply H. lia. Qed. - Lemma attention_always_not_attention e k: - attention e k -> forall m, k < m -> ~ attention e m. + Lemma attend_always_not_attend e k: + attend e k -> forall m, k < m -> ~ attend e m. Proof. - intros H1 m Hm. eapply active_not_attention. - apply (attention_always_active H1 Hm). + intros H1 m Hm. eapply act_not_attend. + apply (attend_always_act H1 Hm). Qed. - Lemma attention_at_most_once e: ~ ~ (exists s, forall s', s < s' -> ~ attention e s'). + Lemma attend_at_most_once e: ~ ~ (exists s, forall s', s < s' -> ~ attend e s'). Proof. - ccase (exists k, attention e k) as [[w Hw]|H]. + ccase (exists k, attend e k) as [[w Hw]|H]. - intros H. eapply H. exists w. - now eapply attention_always_not_attention. + now eapply attend_always_not_attend. - intros Hk. apply Hk. exists 0. intros k' Hk' Ha. apply H. now exists k'. Qed. - Lemma attention_uni e k1 k2 : attention e k1 -> attention e k2 -> k1 = k2. + Lemma attend_uni e: unique (attend e). Proof. - intros H1 H2. - specialize (fun a b => @active_not_attention _ _ (@attention_always_active _ _ H1 a b)) as H1'. - specialize (fun a b => @active_not_attention _ _ (@attention_always_active _ _ H2 a b)) as H2'. + intros k1 k2 H1 H2. + specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H1 a b)) as H1'. + specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H2 a b)) as H2'. enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. intro Hk. eapply H1'. apply Hk. easy. intro Hk. eapply H2'. apply Hk. easy. Qed. - Lemma pick_el_uni e x1 x2: pick_el e x1 -> pick_el e x2 -> x1 = x2 . + Lemma pick_el_uni e: unique (pick_el e). Proof. - intros [k [Ht Hk]] [k' [Ht' Hk']]. - assert (k=k') as <-. eapply attention_uni; eauto. - eapply (@extend_uni simple_extendsion); cbn; eauto. + intros x1 x2 [k [Ht Hk]] [k' [Ht' Hk']]. + assert (k=k') as <-. eapply attend_uni; eauto. + eapply (@extend_uni simple_extension); cbn; eauto. Qed. End Requirements_Facts. @@ -351,12 +361,13 @@ Section Assume_EA. - clear IHL Hin. destruct (F_pick Hk) as [k' [Hk' [e He]]]. exists e. split. - exists k'. split; unfold attention. - + destruct He; intuition eauto. enough (Pf_ k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + exists e. unfold ext_choice. destruct He; intuition eauto. - enough (Pf_ k' = L) as -> by eauto. + exists k'. split; unfold attend. + + destruct He; intuition eauto. enough (P_func k' = L) as -> by eauto. eapply F_uni. apply F_func_correctness. easy. + + exists e. unfold ext_choice. + enough (P_func k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. + destruct He; intuition eauto. + destruct He. destruct H0. destruct H1. destruct H1. lia. - destruct (F_pop Hk) as [m' Hm']. @@ -414,7 +425,7 @@ Section Assume_EA. Lemma PredNoDupListTo_NNExist p: forall b, ~~ exists L, PredListTo p L b /\ NoDup L. Proof. - destruct (F_computable simple_extendsion) as [f Hf]. + destruct (F_computable simple_extension) as [f Hf]. induction b; intros H. - ccase (p 0) as [H0 | H0]; apply H. + exists [0]. split; try split. @@ -508,19 +519,40 @@ Section Assume_EA. - rewrite firstn_length. lia. - now eapply firstn_NoDup. - intros ? ? % firstn_In. now eapply H2. - Qed. + Qed. End Compl_P_non_finite. Section Meet_Requirement. - Lemma attention_at_most_once_bound k: - ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attention e s'). + Lemma wall_of_wall' e: exists w, forall x, wall e (P_func x) x < w. + Proof. + destruct (wall_spec e) as [v [k Hk]]. + destruct (@extra_bounded (fun k => wall e (P_func k) k) k) as [w Hw]. + exists (S (max v w)). intros x. + destruct (Dec (x < k)). apply Hw in l. lia. + assert (wall e (P_func x) x = v). apply Hk. lia. lia. + Qed. + + Lemma wall_of_wall e: exists w, forall i x, i <= e -> wall i (P_func x) x < w. + Proof. + induction e. + - destruct (wall_of_wall' 0) as [w Hw]. exists w. + intros i x ?. inv H. trivial. + - destruct IHe as [w IHe]. + destruct (wall_of_wall' (S e)) as [w' Hw]. + exists (S (max w w')). intros i x H. + inv H. specialize (Hw x). lia. + specialize (IHe i x H1). lia. + Qed. + + Lemma attend_at_most_once_bound k: + ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). Proof. induction k. - intros H. apply H. exists 42. intros ??. lia. - intros H. apply IHk. intros [s Hs]; clear IHk. - specialize (@attention_at_most_once k) as Hk. + specialize (@attend_at_most_once k) as Hk. apply Hk. intros [sk Hsk]; clear Hk. set (max sk s) as N. eapply H. exists N. intros e He. @@ -530,22 +562,25 @@ Section Assume_EA. Qed. Lemma non_finite_not_bounded e: - non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x. + non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ + (forall n, forall i, i <= e -> wall i (P_func n) n < x). Proof. intro H. unfold non_finite in H. intros He. rewrite non_finite_nat in H. - specialize (H (2 * e + 1)). + destruct (wall_of_wall e) as [w Hw]. + pose (N := max (2*e + 1) w). specialize (H N). apply H. intros [m [Hm1 [k Hmk]]]. - apply He. exists k, m; split; last lia. - now exists k. + apply He. exists k, m. + repeat split. now exists k. + lia. intros n i Hie. specialize (Hw i n Hie). lia. Qed. - Lemma ext_pick_attention N e: - e < N -> (exists w, w < e /\ ext_pick (Pf_ N) N w) -> - (exists w, w < e /\ attention w N). + Lemma ext_pick_attend N e: + e < N -> (exists w, w < e /\ ext_pick (P_func N) N w) -> + (exists w, w < e /\ attend w N). Proof. intros HeN [w (Hw1 & Hw2)]. - assert (exists w, ext_pick (Pf_ N) N w) by now exists w. + assert (exists w, ext_pick (P_func N) N w) by now exists w. eapply least_ex in H; last eauto. destruct H as [k Hk]. assert (k <= w). { enough (~ k > w) by lia. intro Hkw. @@ -554,30 +589,31 @@ Section Assume_EA. exists k. do 2 (split; first lia). eapply Hk. Qed. - Lemma non_finite_attention e: - non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (Pf_ k) k e \/ attention e k) . + Lemma non_finite_attend e: + non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ attend e k) . Proof. intros H He. - eapply (non_finite_not_bounded H); intros (b & x & Hxb1 & Hxb2). - eapply (@attention_at_most_once_bound e); intros [s Hs]. - eapply He. - pose (N := S (max (max b s) e)). - destruct (Dec (ext_intersect_W (Pf_ N) N e)) as [He'|He']. + eapply (non_finite_not_bounded H); intros (b & x & (Hx1 & Hx2 & Hx3)). + eapply (@attend_at_most_once_bound e); intros [s Hs]. + pose (N := S (max (max b s) e)). apply He. + destruct (Dec (ext_intersect_W (P_func N) N e)) as [He'|He']. - exists N. right. - assert (ext_pick (Pf_ N) N e). - { split; first easy. exists x. split; last easy. - eapply W_incl; last apply Hxb1. lia. } + assert (ext_pick (P_func N) N e). + { split; first easy. exists x. split. + eapply W_incl; last apply Hx1. lia. + split; first lia. eapply Hx3. + } split. lia. split. easy. intros w HEw Hw. - assert (exists w, w < e /\ ext_pick (Pf_ N) N w). - now exists w. eapply ext_pick_attention in H1. + assert (exists w, w < e /\ ext_pick (P_func N) N w). + now exists w. eapply ext_pick_attend in H1. destruct H1 as [g [HEg Hg]]. eapply (Hs g HEg); last apply Hg. lia. lia. - exists N. now left. Qed. Lemma ext_intersect_W_intersect k e: - ~ (Pf_ k # W[k] e) -> W e #ₚ P -> False. + ~ (P_func k # W[k] e) -> W e #ₚ P -> False. Proof. unfold ext_intersect_W. intros H1 H2. apply H1. @@ -589,10 +625,10 @@ Section Assume_EA. Lemma P_meet_R_simple : forall e, R_simple P e. Proof. intros e He. intros He'. - eapply (non_finite_attention He). + eapply (non_finite_attend He). intros [k [H|H]]. - eapply ext_intersect_W_intersect; eauto. - - eapply attention_active in H. + - eapply attend_act in H. eapply ext_intersect_W_intersect; eauto. Qed. @@ -623,68 +659,71 @@ Section Assume_EA. End Result. - Section Effectively_Simple. - - Definition effectively_simple P := - simple P /\ exists f, - forall e, (forall x, W e x -> (compl P) x) -> forall x, W e x -> x < (f e). - - Lemma attention_pick e k: attention e k -> exists x, x > 2*e /\ P x /\ W e x. - Proof. - intros [He H]. - edestruct (ext_pick_witness) as [x Hx]. - { destruct H. eapply e0. } - assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. - inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. - exists e. assert (Pf_ k = l) as <-. - eapply F_uni. apply F_func_correctness. apply H3. - split; first easy. split; first easy. easy. - assert (x = a) as ->. eapply (@extend_uni simple_extendsion); cbn; eauto. - exists a. split. destruct H4, H0, H1, H4, H4. - assert (x=e) as HE. - { eapply least_unique; last eapply H. - enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } - lia. split. exists (Pf_ (S k)), (S k); split; eauto. now rewrite <- H2. - apply F_func_correctness. destruct H4, H0, H1, H4, H4, H4, H4. - assert (x=e) as HE. - { eapply least_unique; last eapply H. - enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } - exists x0; congruence. - exfalso. eapply (H3 x); exists e; do 2 (split; eauto). - enough ((Pf_ k) = (Pf_ (S k))) as <- by easy. - assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. - eapply F_uni; eauto. - Qed. - - Theorem P_effectively_simple: effectively_simple P. - Proof. - split; first apply P_simple. - exists (fun e => 2 * e + 1). - intros e He x Hex. enough (~ 2 * e < x) by lia. - intros Hex'. - assert (W e #ₚ P). - { intros y Hy1 Hy2. now apply (He y). } - assert (forall k, (Pf_ k) # W[k] e). - { intros k y Hy1 [w [_ Hy2]]. eapply (H y). exists w; eauto. - exists (Pf_ k), k; split; eauto. apply F_func_correctness. } - enough (exists k, attention e k) as [k Hk]. - (* apply H1. intros [k Hk]. *) - eapply attention_pick in Hk. - destruct Hk as [y (Hx1&Hx2&Hx3)]. - eapply (He y); eauto. - { unfold attention. - assert (exists k, least (ext_pick (Pf_ k) k) e /\ e < k). - { destruct Hex as [k Hk]. - pose (S (max e k)) as N. - unfold ext_pick. exists N. split. split. split. admit. - exists x. unfold ext_has_wit. split; last easy. - exists k; split; eauto. lia. 2: {lia. } admit. - } - admit. - } - Admitted. - - End Effectively_Simple. + End Assume_WALL. + + Section Instance_of_Wall. + + Definition low_wall (e: nat) (l: list nat) (n: nat) := 42. + Lemma low_wall_spec: forall e, exists b, lim_to low_wall (low_wall e) b. + Proof. intro e. exists 42. exists 0; intuition. Qed. + + Definition Pw := P low_wall. + + Theorem P_simple_low_wall: simple Pw. + Proof. + eapply P_simple. apply low_wall_spec. + Qed. + + Definition effectively_simple P := + simple P /\ exists f, + forall e, (forall x, W e x -> (compl P) x) -> forall x, W e x -> x < (f e). + + Lemma attend_pick e k: attend low_wall e k -> exists x, x > 2*e /\ Pw x /\ W e x. + Proof. + intros [He H]. + (* edestruct (ext_pick_witness) as [x Hx]. + { destruct H. eapply e0. } + assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. + exists e. assert (Pf_ k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + split; first easy. split; first easy. easy. + assert (x = a) as ->. eapply (@extend_uni simple_extendsion); cbn; eauto. + exists a. split. destruct H4, H0, H1, H4, H4. + assert (x=e) as HE. + { eapply least_unique; last eapply H. + enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } + lia. split. exists (Pf_ (S k)), (S k); split; eauto. now rewrite <- H2. + apply F_func_correctness. destruct H4, H0, H1, H4, H4, H4, H4. + assert (x=e) as HE. + { eapply least_unique; last eapply H. + enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } + exists x0; congruence. + exfalso. eapply (H3 x); exists e; do 2 (split; eauto). + enough ((Pf_ k) = (Pf_ (S k))) as <- by easy. + assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. + eapply F_uni; eauto. *) + Admitted. + + Theorem P_effectively_simple: effectively_simple Pw. + Proof. + split; first apply P_simple. apply low_wall_spec. + exists (fun e => 2 * e + 1). + intros e He x Hex. enough (~ 2 * e < x) by lia. + intros Hex'. + assert (W e #ₚ Pw). + { intros y Hy1 Hy2. now apply (He y). } + (* assert (forall k, (Pf_ low_wall k) # W[k] e). + { intros k y Hy1 [w [_ Hy2]]. eapply (H y). exists w; eauto. + exists (Pf_ k), k; split; eauto. apply F_func_correctness. } + enough (exists k, attend e k) as [k Hk]. + (* apply H1. intros [k Hk]. *) + eapply attend_pick in Hk. + destruct Hk as [y (Hx1&Hx2&Hx3)]. + admit. admit. *) + Admitted. + + End Instance_of_Wall. End Assume_EA. From 078f19c8984202bc1ced89b00da92f4eca2622dd Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 1 Apr 2024 19:43:56 +0200 Subject: [PATCH 21/54] step-indexed Oracle --- theories/ReducibilityDegrees/low_wall.v | 61 ++- .../ReducibilityDegrees/priority_method.v | 24 +- .../ReducibilityDegrees/simple_extension.v | 2 +- theories/TuringReducibility/StepIndex.v | 410 ++++++++++++++++++ 4 files changed, 487 insertions(+), 10 deletions(-) create mode 100644 theories/TuringReducibility/StepIndex.v diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 022846b..990edf7 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -11,10 +11,63 @@ Section Requirements_Lowness_Correctness. Variable P: nat -> Prop. + Notation "f ↓" := (f = Some tt) (at level 30). + + Variable Φ_: (nat -> bool -> Prop) -> nat -> nat -> nat -> option unit. + + Hypothesis Φ_spec: + forall e x, Ξ e (char_rel P) x -> (∞∀ n, Φ_ (char_rel P) e x n ↓). + + Definition Ω e n := Φ_ (char_rel P) e e n. + + Hypothesis N_requirements: + forall e, (∞∃ n, Ω e n ↓) -> Ξ e (char_rel P) e. + + Hypothesis LEM_Σ_2: forall (P: nat -> nat -> Prop), + (forall n m, dec (P n m)) -> dec (exists n, forall m, P n m). + + Lemma limit_case e: (∞∀ n, Ω e n ↓) \/ (∞∀ n, ~ Ω e n ↓). + Proof. + assert (forall m n, dec (m < n -> ~ Ω e n ↓)) as H' by eauto. + destruct (LEM_Σ_2 H'); first now right. + assert (forall m n, dec (m < n -> Ω e n ↓)) as H by eauto. + destruct (LEM_Σ_2 H); first now left. + left. apply Φ_spec. eapply N_requirements. + intros i. + assert (forall n, dec (n > i /\ Ω e n ↓)) as H'' by eauto. + destruct (@LEM_Σ_2 (fun n _ => (n > i /\ Ω e n ↓)) ); first eauto. + destruct e0 as [w Hw]. exists w. apply Hw. exact 42. + exfalso. firstorder. + Qed. + + Definition limit_decider e n: bool := Dec (Ω e n ↓). + Lemma Jump_limit :limit_computable (P´). Proof. - unfold J. - unfold limit_computable. - admit. - Admitted. + exists limit_decider; split; intros. + - unfold J. split. + intros [w Hw]%Φ_spec; exists (S w); intros??. + apply Dec_auto. now eapply Hw. + intros [N HN]. eapply N_requirements. + intros m. exists (S N + m); split; first lia. + eapply Dec_true. eapply HN. lia. + - unfold J. split; intros H. + destruct (limit_case x) as [[k Hk]|h2]. + enough (Ξ x (char_rel P) x) by easy. + eapply N_requirements. intros m. exists (S k + m). + split; first lia. eapply Hk. lia. + destruct h2 as [w Hw]. exists (S w). + intros. specialize (Hw n H0). unfold limit_decider. + destruct (Dec _); eauto. + destruct H as [w Hw]. + intros [k Hneg]%Φ_spec. + set (N := S (max w k)). + assert (Φ_ (char_rel P) x x N ↓). + { eapply Hneg. lia. } + enough (~ Φ_ (char_rel P) x x N ↓) by eauto. + eapply Dec_false. eapply Hw. lia. + Qed. + End Requirements_Lowness_Correctness. + + diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index 7c28898..178aa76 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -12,11 +12,27 @@ Notation "'Σ' x .. y , p" := format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") : type_scope. +Definition inf_exists P := forall n, exists m, m > n /\ P m. +Notation "'∞∃' x .. y , p" := + (inf_exists (fun x => .. (inf_exists (fun y => p)) ..)) + (at level 200, x binder, right associativity, + format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") + : type_scope. + +Definition inf_forall (P: nat -> Prop) := exists n, forall m, m > n -> P m. +Notation "'∞∀' x .. y , p" := + (inf_forall (fun x => .. (inf_forall (fun y => p)) ..)) + (at level 200, x binder, right associativity, + format "'[' '∞∀' '/ ' x .. y , '/ ' p ']'") + : type_scope. + +Notation unique p := (forall x x', p x -> p x' -> x = x'). + Class Extension := { extendP: list nat -> nat -> nat -> Prop; extend_dec: forall l x, (Σ y, extendP l x y) + (forall y, ~ extendP l x y); - extend_uni: forall l x y1 y2, extendP l x y1 -> extendP l x y2 -> y1 = y2 + extend_uni: forall l x, unique (extendP l x) }. Inductive F_ (E: Extension) : nat -> list nat -> Prop := @@ -28,9 +44,9 @@ Section Construction. Variable E: Extension. - Lemma F_uni: forall n l1 l2, F_ E n l1 -> F_ E n l2 -> l1 = l2 . + Lemma F_uni n: unique (F_ E n). Proof. - intros n l1 l2. + intros l1 l2. dependent induction n. - intros H1 H2. inv H1. now inv H2. - intros H1 H2. inv H1; inv H2. @@ -194,8 +210,6 @@ Section EWO. End EWO. -Notation unique p := (forall x x', p x -> p x' -> x = x'). - Section LeastWitness. Definition safe p n := forall k, k < n -> ~ p k. diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 3dcfa6b..fc997c5 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -322,7 +322,7 @@ Section Assume_EA. apply (attend_always_act H1 Hm). Qed. - Lemma attend_at_most_once e: ~ ~ (exists s, forall s', s < s' -> ~ attend e s'). + Lemma attend_at_most_once e: ~ ~ (∞∀ s, ~ attend e s). Proof. ccase (exists k, attend e k) as [[w Hw]|H]. - intros H. eapply H. exists w. diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v new file mode 100644 index 0000000..c500cb9 --- /dev/null +++ b/theories/TuringReducibility/StepIndex.v @@ -0,0 +1,410 @@ +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions. +Require Import SyntheticComputability.Synthetic.DecidabilityFacts. +Require Export SyntheticComputability.Shared.FinitenessFacts. +Require Export SyntheticComputability.Shared.Pigeonhole. +Require Export SyntheticComputability.Shared.ListAutomation. +From SyntheticComputability Require Import partial Dec. +Require Import Coq.Program.Equality. +From stdpp Require Import list. +Import PartialTactics. + +Notation "'Σ' x .. y , p" := + (sigT (fun x => .. (sigT (fun y => p)) ..)) + (at level 200, x binder, right associativity, + format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") + : type_scope. + +Definition inf_forall (P: nat -> Prop) := exists n, forall m, n ≤ m → P m. + Notation "'∞∀' x .. y , p" := + (inf_forall (fun x => .. (inf_forall (fun y => p)) ..)) + (at level 200, x binder, right associativity, + format "'[' '∞∀' '/ ' x .. y , '/ ' p ']'") + : type_scope. + +Lemma list_cons_or {X} (A B: list X) a : + A `prefix_of` B ++ [a] → A `prefix_of` B ∨ A = B ++ [a]. +Proof. + induction A in B |-*; intros. + { left. eapply prefix_nil. } + destruct B. + { right. list_simplifier. + set (H' := H). + apply prefix_cons_inv_2, prefix_nil_inv in H' as ->. + by apply prefix_cons_inv_1 in H as ->. } + list_simplifier. + set (H' := H). + apply prefix_cons_inv_2 in H'. + apply prefix_cons_inv_1 in H as ->. + destruct (IHA _ H') as [h1 | ->]. + + left. by eapply prefix_cons. + + by right. +Qed. + +Section Step_Eval. + + Context {Part : partiality}. + Context {Q A O: Type}. + Definition tree := (list A) ↛ (Q + O). + + Notation Ask q := (inl q). + Notation Output o := (inr o). + + Fixpoint evalt_comp (τ : tree) (f : Q → A) (step depth: nat): option (Q + O) := + match (seval (τ []) depth) with + | Some x => + match step, x with + | 0, Ask q => Some (Ask q) + | S n, Ask q => evalt_comp (λ l, τ (f q :: l)) f n depth + | _, Output o => Some (Output o) + end + | None => None + end. + + Lemma evalt_comp_ext (τ τ': tree) f n m: + (∀ l n, seval (τ l) n = seval (τ' l) n) → + evalt_comp τ f n m = evalt_comp τ' f n m. + Proof. + intro Heq; induction n in τ, τ', Heq |- *; cbn. + - by rewrite <- Heq. + - rewrite <- Heq. + destruct (seval (A:=Q + O) (τ []) m); eauto. + destruct s; eauto. + Qed. + + Lemma interrogation_ter (τ: tree) f l lv v: + interrogation τ (λ x y, f x = y) l lv → τ lv =! v → + ∃ m, ∀ ans_, ans_ `prefix_of` lv → ∃ v, seval (τ ans_) m = Some v. + Proof. + intros H1 H2. + induction H1 in H2, v |-*. + - rewrite seval_hasvalue in H2. + destruct H2 as [m Hm]. exists m. + intros ans_ Hans_. exists v. + apply prefix_nil_inv in Hans_. + rewrite Hans_. easy. + - rewrite seval_hasvalue in H2. + destruct H2 as [m Hm]. + destruct (IHinterrogation (Ask q) H) as [m' Hm']. + exists (max m m'). + intros ans_ Hans_. + destruct (list_cons_or Hans_) as [h| ->]. + + destruct (Hm' ans_ h) as [v' Hv']. + exists v'. eapply seval_mono. + eauto. lia. + + exists v. eapply seval_mono; [eauto| lia]. + Qed. + + (** Basic property of evalt_comp **) + + Lemma evalt_comp_depth_mono (τ: tree) f n m o: + evalt_comp τ f n m = Some o → + ∀ m', m ≤ m' → evalt_comp τ f n m' = Some o. + Proof. + intros H m' Hm'. + induction n in H, τ, o |-*; cbn in *. + - destruct (seval (A:=Q + O) (τ []) m) eqn: E; try congruence. + assert (seval (A:=Q + O) (τ []) m' = (Some s)) as ->. + eapply seval_mono. exact E. lia. done. + - destruct (seval (A:=Q + O) (τ []) m) eqn: E; last congruence. + assert (seval (A:=Q + O) (τ []) m' = Some s) as ->. + eapply seval_mono. exact E. lia. + destruct s. by apply IHn. exact H. + Qed. + + Lemma interrogation_plus_evalt_comp (τ: tree) f n m l lv v2: + interrogation τ (λ x y, f x = y) l lv → + (∀ ans_, ans_ `prefix_of` lv -> ∃ v, seval (τ ans_) m = Some v) → + evalt_comp (λ l', τ (lv ++ l')) f n m = Some v2 ↔ + evalt_comp τ f (length l + n) m = Some v2. + Proof. + intros H H1. split; revert n; dependent induction H; try eauto. + - intros. cbn -[evalt]. rewrite app_length. cbn -[evalt]. + replace (length qs + 1 + n) with (length qs + (S n)) by lia. + eapply IHinterrogation. intros; apply H2. + etransitivity. exact H4. by eapply prefix_app_r. + cbn. rewrite app_nil_r. + destruct (H2 ans); first by eapply prefix_app_r. + assert (exists m, seval (τ ans) m = Some x). + by exists m. rewrite <- seval_hasvalue in H5. + assert (x = Ask q). { eapply hasvalue_det; eauto. } + rewrite H4, H6, H1, <- H3. eapply evalt_comp_ext. + intros; by list_simplifier. + - intros. rewrite app_length in H3. cbn in H3. + replace (length qs + 1 + n) with (length qs + (S n)) in H3 by lia. + eapply IHinterrogation in H3. + 2: { intros; apply H2. etransitivity. exact H4. + by eapply prefix_app_r. } + cbn in H3. rewrite app_nil_r in H3. + destruct (H2 ans); first by eapply prefix_app_r. + assert (exists m, seval (τ ans) m = Some x) by by exists m. + rewrite <- seval_hasvalue in H5. + assert (x = Ask q). { eapply hasvalue_det; eauto. } + rewrite H4, H6, H1 in H3. + rewrite <- H3. eapply evalt_comp_ext. + intros; by list_simplifier. + Qed. + + Lemma evalt_comp_step_mono (τ: tree) f qs ans o: + interrogation τ (λ x y, f x = y) qs ans → + τ ans =! Output o → + ∃ depth step, ∀ g, interrogation τ (λ x y, g x = y) qs ans → + ∀ n', step ≤ n' → evalt_comp τ g n' depth = Some (Output o). + Proof. + intros H1 H2. + destruct (interrogation_ter H1 H2) as [step Hstep]. + exists step. exists (length qs). intros g Hg n' Hn'. + assert (exists v, seval (τ ans) step = Some v) as [v Hv]. + { eapply Hstep; naive_solver. } + assert (v = Output o). + { eapply hasvalue_det; [|eapply H2]. rewrite seval_hasvalue. eauto. } + eapply Nat.le_exists_sub in Hn' as [k [-> _]]. + rewrite Nat.add_comm. + eapply interrogation_plus_evalt_comp ; eauto. + induction k. all: cbn; rewrite app_nil_r; by rewrite Hv, H. + Qed. + + Lemma evalt_comp_oracle_approx (τ: tree) f l lv v: + interrogation τ (λ x y, f x = y) l lv → + τ lv =! v → + ∃ step depth, ∀ g, interrogation τ (λ x y, g x = y) l lv → + evalt_comp τ g step depth = Some v. + Proof. + intros H1 h2. + destruct (interrogation_ter H1 h2) as [step Hstep]. + exists (length l + 0), step. + intros. eapply interrogation_plus_evalt_comp; eauto. + cbn. rewrite app_nil_r. + destruct (Hstep lv) as [v' Hv']; first done. + assert (∃ k, seval (A:=Q + O) (τ lv) k = Some v') by by exists step. + rewrite <- seval_hasvalue in H0. + assert (v' = v) by by eapply hasvalue_det. + rewrite Hv', H2. by destruct v. + Qed. + + Lemma interrogation_evalt_comp_limit (τ: tree) f l lv v: + (∞∀ k, interrogation τ (λ x y, f k x = y) l lv) → + τ lv =! Output v → + ∞∀ n, evalt_comp τ (f n) n n = Some (Output v). + Proof. + intros [k h1] h2. + assert (interrogation τ (λ x y, f k x = y) l lv) as H. + apply h1. lia. + destruct (evalt_comp_step_mono H h2) as (a' & b' & Hs). + destruct (evalt_comp_oracle_approx H h2) as (a & b & Hab). + exists (max b'(max a' (max (max a b) k))). + intros n Hn. eapply evalt_comp_depth_mono. + eapply (Hs (f n)); eauto. eapply h1. + all: lia. + Qed. + + Lemma evalt_comp_to_interrogation (τ : tree) (f : Q → A) o n depth: + evalt_comp τ f n depth = Some (Output o) → + Σ (qs : list Q) (ans : list A), + length qs ≤ n ∧ + interrogation τ (λ q a, f q = a) qs ans ∧ + τ ans =! Output o. + Proof. + intros H. + induction n in τ, H |- *. + * cbn in *. destruct (seval (τ []) depth) eqn: E. + exists [], []. repeat split. eauto. econstructor. + destruct s. congruence. rewrite seval_hasvalue. + by exists depth; injection H as ->. congruence. + * cbn in *. destruct (seval (τ []) depth) eqn: E; try congruence. + destruct s; try congruence. + - eapply (IHn (λ l, τ (f q :: l))) in H as (qs & ans & H3 & H1 & H2). + exists (q :: qs), (f q :: ans). split; eauto. cbn; lia. repeat split. + eapply interrogation_app with (q1 := [q]) (a1 := [f q]). + eapply Interrogate with (qs := []) (ans := []); eauto. econstructor. + rewrite seval_hasvalue. by exists depth. + eauto. eauto. + - exists [], []. repeat split. cbn. lia. eauto. econstructor. + rewrite seval_hasvalue. + by exists depth; injection H as ->. + Qed. + + Lemma evalt_comp_limit_interrogation (τ: tree) f v: + (∞∀ n, evalt_comp τ (f n) n n = Some (Output v)) → + (∞∀ k, ∃ l lv, τ lv =! Output v ∧ interrogation τ (λ x y, f k x = y) l lv). + Proof. + intros [w Hw]. + exists w. intros m Hm. + specialize (Hw m Hm) as H. + destruct (evalt_comp_to_interrogation H) as (qs&ans&Hl&Hans&Hans'). + exists qs, ans. split; eauto. + Qed. + +End Step_Eval. + +Section Use_Function. + + Definition list_interp (L: list nat) (q: nat): bool := Dec (In q L). + + Lemma try (O: Type) (τ: tree) n m (v: nat + O) L: + evalt_comp τ (list_interp L) n m = Some v → Σ k, + ∀ x, k < x → evalt_comp τ (list_interp (x::L)) n m = Some v. + Proof. + induction n in v |-*. + - exists 42. intros x Hx. + unfold evalt_comp in *. by destruct (seval (τ []) m). + - intros H. destruct (evalt_comp τ (list_interp L) n m) eqn: E; last first. + admit. + destruct s; last first. + (* specialize (evalt_comp_step_mono ). *) + (* exists 42. intros H. + destruct v as [i|]. + exists (S (max i k)). intros Hi x Hx. + *) + Abort. + +End Use_Function. + +Section Limit_Interrogation. + + Variable Q: Type. + Variable P: Q → Prop. + + Definition stable (f: nat → Q → bool) := + ∀ q n m, n ≤ m → f n q = true → f m q = true. + + Fixpoint stabilize_step (f: Q -> nat -> bool) x n := + match n with + | O => false + | S n => if f x n then true else stabilize_step f x n + end. + + Definition stable_semi_decider (f: nat → Q → bool) := + semi_decider (λ x n, f n x) P ∧ stable f. + + Fact semi_decider_to_stable: ∀ f, semi_decider f P → Σ g, stable_semi_decider g. + Proof. + intros f S_P. exists (λ n x, stabilize_step f x n); split. + - intro x; split; intro h. + rewrite (S_P x) in h. + destruct h as [c Hc]. + by exists (S c); cbn; rewrite Hc. + rewrite (S_P x). + destruct h as [c Hc]. + induction c; cbn in Hc; [congruence|]. + destruct (f x c) eqn: E; [now exists c|now apply IHc]. + - intros x n m Lenm Hn. + induction Lenm; [trivial|]. + cbn; destruct (f x m) eqn: E; [trivial|assumption]. + Qed. + + Definition approx_Σ1 O (f: nat → Q → bool) := + ∀ (τ: list bool ↛ (Q + O)) qs ans, + interrogation τ (char_rel P) qs ans → + ∞∀ m, interrogation τ (λ q a, f m q = a) qs ans. + + Definition approx_Σ1_rev O (f: nat → Q → bool) := + ∀ (τ: list bool ↛ (Q + O)) qs ans, + (∞∀ m, interrogation τ (λ q a, f m q = a) qs ans) → + interrogation τ (char_rel P) qs ans. + + Definition approx_list (f: Q → bool) L := + ∀ i, In i L → P i ↔ f i = true. + + Definition approx_list_func (f g: Q → bool) L := + ∀ i, In i L → f i ↔ g i. + + Variable g: nat → Q → bool. + Hypothesis S_P: stable_semi_decider g. + + Definition S_approx_Σ1: ∀ O, approx_Σ1 O g. + Proof. + destruct S_P as [H1 H2]; intros O τ qs ans Hτ. + induction Hτ. + - exists 0; intros; econstructor. + - destruct IHHτ as [w Hw]. + destruct a; cbn in H0. + + rewrite (H1 q) in H0. destruct H0 as [m Hm]. + exists (S (max w m)). intros m' Hm'. + econstructor. eapply Hw; first lia. + assumption. eapply H2; last apply Hm. lia. + + exists w; intros m Hm. + econstructor. eapply Hw; first done. + assumption. destruct (g m q) eqn: E; last done. + enough (P q) by done. { rewrite (H1 q). by exists m. } + Qed. + + Lemma approx_Σ1_list: definite P → ∀ L, ∞∀ c, approx_list (g c) L. + Proof. + destruct S_P as [S_p HS]. + intros def_p l. induction l as [|a l [c Hc]]. + - exists 42; firstorder. + - destruct (def_p a) as [h1| h2]. + + destruct (S_p a) as [H _]. + destruct (H h1) as [N HN]. + exists (max c N); intros c' Hc' e [->| He]. + split; [intros _|easy]. + eapply HS; [|eapply HN]; lia. + unfold approx_list in Hc. + rewrite (Hc c'); [trivial|lia | assumption]. + + exists c; intros c' Hc' e [->| He]. + split; [easy| intros h']. + unfold semi_decider in S_p. + rewrite S_p. by exists c'. + unfold approx_list in Hc. + rewrite Hc; eauto. + Qed. + + Definition S_approx_Σ1_rev: definite P → ∀ O, approx_Σ1_rev O g. + Proof. + intros defp O τ qs ans [w Hw]. + assert (∞∀ k, ∀ q, In q qs → P q ↔ g k q = true) as [k Hk] by by apply approx_Σ1_list. + assert (interrogation τ (λ q a, g (max w k) q = a) qs ans) as H by (apply Hw; lia). + clear Hw. induction H; first econstructor. + econstructor; [|done|]. + eapply IHinterrogation. + { intros. rewrite Hk; try done. apply in_app_iff. by left. } + destruct a; cbn; rewrite Hk; try (done||lia). + rewrite in_app_iff; right; econstructor. done. + intro H'. enough (g (w`max`k) q = true) by congruence. + destruct S_P as [H1' H2]. eapply H2; last apply H'. lia. + rewrite in_app_iff; right; econstructor. done. + Qed. + +End Limit_Interrogation. + +Section Step_Eval_Spec. + + Variable P: nat → Prop. + Variable f_: nat → nat → bool. + Hypothesis Hf_ : semi_decider f_ P. + Hypothesis D_P: definite P. + + Definition f := projT1 (semi_decider_to_stable Hf_). + Fact S_P: stable_semi_decider P f. + Proof. unfold f. by destruct (semi_decider_to_stable Hf_). Qed. + + Definition Φ_ (f: nat → nat → bool) (e x n: nat): option () := + match evalt_comp (ξ () e x) (f n) n n with + | Some (inr ()) => Some () + | _ => None + end. + + Lemma Φ_spec_1 e x: + Ξ e (char_rel P) x → (∞∀ n, Φ_ f e x n = Some ()). + Proof. + unfold Ξ, rel. intros (qs & ans & Ha & Hb). + specialize (@S_approx_Σ1 _ _ _ S_P () (ξ _ e x) qs ans Ha) as H. + eapply interrogation_evalt_comp_limit in H; last apply Hb. + destruct H as [w Hw]. + exists w; intros m Hm. unfold Φ_. specialize (Hw m Hm). + destruct (evalt_comp (ξ () e x) (f m) m m). + destruct s. by injection Hw. + by destruct u. congruence. + Qed. + + Lemma Φ_spec: + Σ Φ_, ∀ e x, Ξ e (char_rel P) x → (∞∀ n, Φ_ f e x n = Some ()). + Proof. exists Φ_; intros e x; apply Φ_spec_1. Qed. + +End Step_Eval_Spec. + + + + + + From 5d84cf510b3e230169fcc4164786b74004d0396b Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 2 Apr 2024 01:15:42 +0200 Subject: [PATCH 22/54] done Phi --- theories/ReducibilityDegrees/low_wall.v | 68 +++++++++++-------- .../ReducibilityDegrees/priority_method.v | 29 ++++---- .../ReducibilityDegrees/simple_extension.v | 2 +- theories/TuringReducibility/StepIndex.v | 10 ++- 4 files changed, 60 insertions(+), 49 deletions(-) diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 990edf7..145f03c 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -1,4 +1,4 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions StepIndex Limit simple. Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. @@ -7,37 +7,50 @@ Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. Require Import SyntheticComputability.ReducibilityDegrees.priority_method. Require Import SyntheticComputability.ReducibilityDegrees.simple_extension. +Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. + Notation "'∞∃' x .. y , p" := + (inf_exists (λ x, .. (inf_exists (λ y, p)) ..)) + (at level 200, x binder, right associativity, + format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") + : type_scope. + Section Requirements_Lowness_Correctness. - Variable P: nat -> Prop. + Variable P: nat → Prop. + Hypothesis S_P: Σ f, semi_decider f P. + + Notation "f ↓" := (f = Some ()) (at level 30). - Notation "f ↓" := (f = Some tt) (at level 30). + Definition Φ_ := projT1 (Φ_spec S_P). - Variable Φ_: (nat -> bool -> Prop) -> nat -> nat -> nat -> option unit. + Fact Φ_spec e x: Ξ e (char_rel P) x → ∞∀ n, Φ_ (StepIndex.f S_P) e x n ↓. + Proof. + intro H. unfold Φ_. destruct (Φ_spec S_P) as [Φ Φ_spec]. + cbn. destruct (Φ_spec e x H) as [w Hw]. + exists w. intros. eapply Hw. easy. + Qed. - Hypothesis Φ_spec: - forall e x, Ξ e (char_rel P) x -> (∞∀ n, Φ_ (char_rel P) e x n ↓). + Definition Ω e n := Φ_ (StepIndex.f S_P) e e n. - Definition Ω e n := Φ_ (char_rel P) e e n. + Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. - Hypothesis N_requirements: - forall e, (∞∃ n, Ω e n ↓) -> Ξ e (char_rel P) e. + Hypothesis LEM_Σ_2: + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → dec (∃ n, ∀ m, P n m). - Hypothesis LEM_Σ_2: forall (P: nat -> nat -> Prop), - (forall n m, dec (P n m)) -> dec (exists n, forall m, P n m). + #[export]Instance dec_le: ∀ m n, dec (m ≤ n). + Proof. intros n m; destruct (le_gt_dec n m); [by left|right; lia]. Qed. - Lemma limit_case e: (∞∀ n, Ω e n ↓) \/ (∞∀ n, ~ Ω e n ↓). + Lemma limit_case e: (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). Proof. - assert (forall m n, dec (m < n -> ~ Ω e n ↓)) as H' by eauto. - destruct (LEM_Σ_2 H'); first now right. - assert (forall m n, dec (m < n -> Ω e n ↓)) as H by eauto. - destruct (LEM_Σ_2 H); first now left. - left. apply Φ_spec. eapply N_requirements. - intros i. - assert (forall n, dec (n > i /\ Ω e n ↓)) as H'' by eauto. - destruct (@LEM_Σ_2 (fun n _ => (n > i /\ Ω e n ↓)) ); first eauto. - destruct e0 as [w Hw]. exists w. apply Hw. exact 42. - exfalso. firstorder. + assert (∀ m n, dec (m ≤ n → ¬ Ω e n ↓)) as H by eauto. + destruct (LEM_Σ_2 H); first by right. + left. apply Φ_spec, N_requirements. intros i. + assert (∀ n, dec (n ≤ i ∧ Ω e n ↓)) as H'' by eauto. + destruct (@LEM_Σ_2 (λ n _, i ≤ n ∧ Ω e n ↓) ) as [[w Hw]|h1]; first eauto. + exists w; apply Hw; exact 42. + assert (∀ n, i ≤ n → ~ Ω e n ↓). + { intros m Hm HM. eapply h1. exists m; eauto. } + exfalso. by eapply n; exists i. Qed. Definition limit_decider e n: bool := Dec (Ω e n ↓). @@ -46,8 +59,8 @@ Section Requirements_Lowness_Correctness. Proof. exists limit_decider; split; intros. - unfold J. split. - intros [w Hw]%Φ_spec; exists (S w); intros??. - apply Dec_auto. now eapply Hw. + intros [w Hw]%Φ_spec; exists w; intros??. + apply Dec_auto. by eapply Hw. intros [N HN]. eapply N_requirements. intros m. exists (S N + m); split; first lia. eapply Dec_true. eapply HN. lia. @@ -56,15 +69,14 @@ Section Requirements_Lowness_Correctness. enough (Ξ x (char_rel P) x) by easy. eapply N_requirements. intros m. exists (S k + m). split; first lia. eapply Hk. lia. - destruct h2 as [w Hw]. exists (S w). + destruct h2 as [w Hw]. exists w. intros. specialize (Hw n H0). unfold limit_decider. destruct (Dec _); eauto. destruct H as [w Hw]. intros [k Hneg]%Φ_spec. set (N := S (max w k)). - assert (Φ_ (char_rel P) x x N ↓). - { eapply Hneg. lia. } - enough (~ Φ_ (char_rel P) x x N ↓) by eauto. + assert (Ω x N ↓). { eapply Hneg. lia. } + enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. Qed. diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index 178aa76..c5d8625 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -12,20 +12,6 @@ Notation "'Σ' x .. y , p" := format "'[' 'Σ' '/ ' x .. y , '/ ' p ']'") : type_scope. -Definition inf_exists P := forall n, exists m, m > n /\ P m. -Notation "'∞∃' x .. y , p" := - (inf_exists (fun x => .. (inf_exists (fun y => p)) ..)) - (at level 200, x binder, right associativity, - format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") - : type_scope. - -Definition inf_forall (P: nat -> Prop) := exists n, forall m, m > n -> P m. -Notation "'∞∀' x .. y , p" := - (inf_forall (fun x => .. (inf_forall (fun y => p)) ..)) - (at level 200, x binder, right associativity, - format "'[' '∞∀' '/ ' x .. y , '/ ' p ']'") - : type_scope. - Notation unique p := (forall x x', p x -> p x' -> x = x'). Class Extension := @@ -134,6 +120,21 @@ Section Construction. apply Hf. Qed. + + Lemma F_with_semi_decider: Σ f, semi_decider f F_with. + Proof. + destruct F_computable as [f Hf ]. + exists (fun x n => (Dec (In x (f n)))). + intros x. split. + - intros (l & n & Hxl & Hl). + exists n. rewrite Dec_auto; first easy. + destruct (Hf n) as [Hf' _]. + now rewrite (F_uni Hf' Hl). + - intros (n & Hn%Dec_true). + exists (f n), n; split; eauto. + apply Hf. + Qed. + End Construction. Section StrongInduction. diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index fc997c5..570f9bc 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -322,7 +322,7 @@ Section Assume_EA. apply (attend_always_act H1 Hm). Qed. - Lemma attend_at_most_once e: ~ ~ (∞∀ s, ~ attend e s). + Lemma attend_at_most_once e: ~ ~ (exists s', forall s, s' < s -> ~ attend e s). Proof. ccase (exists k, attend e k) as [[w Hw]|H]. - intros H. eapply H. exists w. diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v index c500cb9..beef38d 100644 --- a/theories/TuringReducibility/StepIndex.v +++ b/theories/TuringReducibility/StepIndex.v @@ -5,7 +5,7 @@ Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. From SyntheticComputability Require Import partial Dec. Require Import Coq.Program.Equality. -From stdpp Require Import list. +From stdpp Require Export list. Import PartialTactics. Notation "'Σ' x .. y , p" := @@ -370,13 +370,11 @@ End Limit_Interrogation. Section Step_Eval_Spec. Variable P: nat → Prop. - Variable f_: nat → nat → bool. - Hypothesis Hf_ : semi_decider f_ P. - Hypothesis D_P: definite P. + Hypothesis Hf_: Σ f_, semi_decider f_ P. - Definition f := projT1 (semi_decider_to_stable Hf_). + Definition f := projT1 (semi_decider_to_stable (projT2 Hf_)). Fact S_P: stable_semi_decider P f. - Proof. unfold f. by destruct (semi_decider_to_stable Hf_). Qed. + Proof. unfold f. by destruct (semi_decider_to_stable (projT2 Hf_)). Qed. Definition Φ_ (f: nat → nat → bool) (e x n: nat): option () := match evalt_comp (ξ () e x) (f n) n n with From df53ddeee618b8111b46416c6cacb5a2b1effc3d Mon Sep 17 00:00:00 2001 From: Haoyi Date: Fri, 12 Apr 2024 13:20:27 +0200 Subject: [PATCH 23/54] final fact --- theories/ReducibilityDegrees/low_wall.v | 135 +++++++- theories/ReducibilityDegrees/lowsimple.v | 69 +++-- .../ReducibilityDegrees/priority_method.v | 39 +++ .../ReducibilityDegrees/simple_extension.v | 3 +- theories/TuringReducibility/StepIndex.v | 289 ++++++++++++++---- 5 files changed, 453 insertions(+), 82 deletions(-) diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 145f03c..f6c71ef 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -7,6 +7,7 @@ Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. Require Import SyntheticComputability.ReducibilityDegrees.priority_method. Require Import SyntheticComputability.ReducibilityDegrees.simple_extension. + Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. Notation "'∞∃' x .. y , p" := (inf_exists (λ x, .. (inf_exists (λ y, p)) ..)) @@ -14,23 +15,20 @@ Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") : type_scope. + Notation "f ↓" := (f = Some ()) (at level 30). + Section Requirements_Lowness_Correctness. Variable P: nat → Prop. - Hypothesis S_P: Σ f, semi_decider f P. - - Notation "f ↓" := (f = Some ()) (at level 30). + Variable f: nat → nat → bool. + Hypothesis S_P: stable_semi_decider P f. - Definition Φ_ := projT1 (Φ_spec S_P). + Definition Φ_ := (Φ_ f). - Fact Φ_spec e x: Ξ e (char_rel P) x → ∞∀ n, Φ_ (StepIndex.f S_P) e x n ↓. - Proof. - intro H. unfold Φ_. destruct (Φ_spec S_P) as [Φ Φ_spec]. - cbn. destruct (Φ_spec e x H) as [w Hw]. - exists w. intros. eapply Hw. easy. - Qed. + Fact Φ_spec e x: Ξ e (char_rel P) x → ∞∀ n, Φ_ e x n ↓. + Proof. intro H. unfold Φ_. apply (Φ_spec S_P H). Qed. - Definition Ω e n := Φ_ (StepIndex.f S_P) e e n. + Definition Ω e n := Φ_ e e n. Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. @@ -83,3 +81,118 @@ Section Requirements_Lowness_Correctness. End Requirements_Lowness_Correctness. +Section Wall. + + Variable η: nat → nat → option nat. + Hypothesis EA: EA η. + + Instance wall_instance: Wall := λ e L n, φ (λ x, Dec (In x L)) e e n. + Definition P := P η wall. + Definition P_func := P_func η wall. + Instance E: Extension := simple_extension η wall. + + Fact P_semi_decidable: semi_decidable P. + Proof. eapply P_semi_decidable. Qed. + + Lemma eventally_wall e: + ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). + Proof. + intros H_. eapply (@attend_at_most_once_bound η wall (S e)). + intros [s Hs]. eapply H_; clear H_. + exists (S s). intros m Hm x [e_ [He_ He_']]. + destruct (Dec (e_ < e)) as [E|E]. + { exfalso. enough (attend η wall e_ m). + unshelve eapply (Hs e_ _ m); eauto; lia. + split; first lia. destruct He_' as [H _]. + apply H. } + assert (e ≤ e_) by lia; clear E. + destruct He_', H1, H1, H1, H1, H3. + by eapply H5. + Qed. + + Hypothesis NoIdea: ∀ (P: nat → Prop) (k: nat), + (∀ x, k ≤ x → P x) ∨ (∃ x, k ≤ x ∧ ¬ P x). + + Definition χ := χ (simple_extension η wall). + Definition P_Φ := (Φ_ χ). + Definition P_Ω := (Ω χ). + + Fact wall_convergence e : ¬¬ ∃ b : nat, lim_to η wall (wall e) b. + Proof. + intros H_. + destruct (@eventally_wall e). intros [N HN]. apply H_; clear H_. + destruct (NoIdea (λ m, wall e (P_func m) m = 0) N). + - exists 0. exists N. intros. by apply H. + - destruct H as [x [H1 H2]]. + destruct (wall e (P_func x) x) as [|k] eqn: H; first done; clear H2. + exists (S k), x. intros t Ht. induction Ht; first done. + rewrite <- (φ_spec1 (F_with_χ (simple_extension η wall)) IHHt). + reflexivity. + intros; split. + + intros K%Dec_true. apply Dec_auto. + enough (incl (F_func (simple_extension η wall) m) (F_func (simple_extension η wall) (S m))). + eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension η wall) (S m)) as HE. + inv HE. + * assert (wall e (P_func m) m < a). + { eapply HN. lia. enough (P_func m = l) as ->. apply H5. + eapply F_uni; [apply F_func_correctness|apply H4]. } + assert (wall e (P_func m) m = S k). { rewrite <-IHHt. reflexivity. } + rewrite H6 in H2. destruct (Dec (a = x0)). + lia. apply Dec_auto. rewrite <- H3 in K. + destruct K; first done. + enough ((F_func (simple_extension η wall) m) = l) as -> by done. + eapply F_uni; last apply H4; apply F_func_correctness. + * apply Dec_auto. + enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. + eapply F_uni; first apply F_func_correctness. + assumption. + Qed. + + Hypothesis DN: ∀ P, ¬ ¬ P → P. + Fact P_simple: simple P. + Proof. eapply P_simple; first apply EA. intro e. apply DN, wall_convergence. Qed. + + Lemma N_requirements: ∀ e, (∞∃ n, P_Ω e n ↓) → ¬ ¬ Ξ e (char_rel P) e. + Proof. + intros e He H_. + apply (@eventally_wall e). intros [N HN]. + apply (@wall_convergence e). intros [B [b Hb]]. + apply H_; clear H_. + set (M := max N b). destruct (He M) as [k [Hk Hk']]. + eapply (@φ_spec χ e e k); first apply Hk'. + intros x Hx. unfold P, simple_extension.P. + rewrite F_with_top. split. + - intros (L & m & HL & HLs &HP). + assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } + assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } + apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. + enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. + eauto. eapply F_mono; last apply E_; apply F_func_correctness. + assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). + assert (k ≤ m) as Ek by lia. enough (x ≤ φ (χ m) e e m). + exfalso. assert (φ (χ m) e e m < x). + apply HN. lia. enough (φ (χ m) e e m = φ (χ k) e e k) by congruence. + assert (φ (χ m) e e m = B). + { rewrite <- (Hb m). reflexivity. lia. } + assert (φ (χ k) e e k = B). + { rewrite <- (Hb k). reflexivity. lia. } + congruence. + - intros H%Dec_true. + eapply F_with_top. exists (F_func _ k), k; split; eauto. + eapply F_func_correctness. + Qed. + + Hypothesis LEM_Σ_2: + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → dec (∃ n, ∀ m, P n m). + + Fact jump_P_limit: limit_computable (P´). + Proof. + eapply Jump_limit; last done. apply F_with_χ. + intros e He. eapply DN, N_requirements; eauto. + Qed. + +End Wall. + + + diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index a6ecd79..d1b800e 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -3,6 +3,7 @@ Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. +Require Export SyntheticComputability.ReducibilityDegrees.low_wall. Require Import Arith. Require Import Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. @@ -21,9 +22,22 @@ Post's Problem in Turing degree. **) (* Definition of low *) Definition low (P: nat -> Prop) := P´ ⪯ᴛ K. - Section LowFacts. - (* If the turing jump of A is low, then A is not reduciable to K *) + + Variable vec_to_nat : forall k, vec nat k -> nat. + Variable nat_to_vec : forall k, nat -> vec nat k. + Variable vec_nat_inv : forall k v, nat_to_vec k (vec_to_nat v) = v. + Variable nat_vec_inv : forall k n, vec_to_nat (nat_to_vec k n) = n. + + Variable list_vec_to_nat : forall k, list (vec nat k) -> nat. + Variable nat_to_list_vec : forall k, nat -> list (vec nat k). + Variable list_vec_nat_inv : forall k v, nat_to_list_vec k (list_vec_to_nat v) = v. + Variable nat_list_vec_inv : forall k n, list_vec_to_nat (nat_to_list_vec k n) = n. + + Variable nat_to_list_bool : nat -> list bool. + Variable list_bool_to_nat : list bool -> nat. + Variable list_bool_nat_inv : forall l, nat_to_list_bool (list_bool_to_nat l) = l. + Variable nat_list_bool_inv : forall n, list_bool_to_nat (nat_to_list_bool n) = n. Lemma lowness (P: nat -> Prop) : low P -> ~ K ⪯ᴛ P. @@ -39,27 +53,46 @@ Section LowFacts. limit_computable (A´) -> ~ K ⪯ᴛ A. Proof. intros LEM defK H IH. - apply lowness with (P := A); [|apply IH]. - (* specialize (limit_turing_red_K _ _ _ _ defK H). *) - Admitted. + apply lowness with (P := A); [|apply IH]. + eapply limit_turing_red_K; eauto. exact 42. + Qed. -End LowFacts. + Definition low_simple P := low P /\ simple P. + Definition sol_Post's_problem (P: nat -> Prop) := + (~ decidable P) /\ (enumerable P) /\ ~ (K ⪯ᴛ P). -Section LowSimplePredicate. + Fact low_simple_correct: + forall P, low_simple P -> sol_Post's_problem P. + Proof. + intros P [H1 H2]; split; [now apply simple_undecidable|]. + split; [destruct H2 as [H2 _]; eauto| now apply lowness]. + Qed. + + (*** Instance of low simple predicate ***) + + Section LowSimplePredicate. -Definition low_simple P := low P /\ simple P. + Variable η: nat -> nat -> option nat. + Hypothesis EA: + forall P, semi_decidable P -> exists e, forall x, P x <-> exists n, η e n = Some x. + Hypothesis LEM_Σ_1: LEM_Σ 1. + Hypothesis def_K: definite K. + Hypothesis sth: forall (P: nat -> Prop) (k: nat), + (forall x, k <= x -> P x) \/ (exists x, k <= x /\ ~ P x). + Hypothesis DN: forall P, ~ ~ P -> P. + Hypothesis LEM_Σ_2: + forall (P: nat -> nat -> Prop), (forall n m, dec (P n m)) -> dec (exists n, forall m, P n m). -Definition sol_Post's_problem (P: nat -> Prop) := - (~ decidable P) /\ (enumerable P) /\ ~ (K ⪯ᴛ P). + Theorem a_sol_Post's_problem: exists P, sol_Post's_problem P. + Proof. + eexists. eapply low_simple_correct; split. + - eapply limit_turing_red_K; eauto. exact 42. + apply jump_P_limit; eauto. + - eapply P_simple; eauto. + Qed. -Fact low_simple_correct: - forall P, low_simple P -> sol_Post's_problem P. -Proof. - intros P [H1 H2]; split; [now apply simple_undecidable|]. - split; [destruct H2 as [H2 _]; eauto| now apply lowness]. -Qed. + End LowSimplePredicate. -End LowSimplePredicate. +End LowFacts. -(*** Instance of low simple predicate ***) \ No newline at end of file diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index c5d8625..8ea67c1 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -1,4 +1,5 @@ From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +From stdpp Require Export list. Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. @@ -105,6 +106,16 @@ Section Construction. Definition F_with x := exists l n, In x l /\ (F_ E n l). + Lemma F_func_F_with x: F_with x ↔ ∃ n, In x (F_func n). + Proof. + split. intros (l&n&H1&H2). + exists n. assert (l = F_func n) as ->. + { eapply F_uni. apply H2. apply F_func_correctness. } + done. intros [n Hn]. exists (F_func n), n; split; first done. + apply F_func_correctness. + Qed. + + Lemma F_with_semi_decidable: semi_decidable F_with. Proof. unfold semi_decidable, semi_decider. @@ -120,6 +131,18 @@ Section Construction. apply Hf. Qed. + Lemma F_with_top x: F_with x <-> exists l n, (F_ E n l) /\ (F_ E (S n) (x::l)) /\ extendP l n x. + Proof. + split; last first. + { intros (l&n&H1&H2&H3). exists (x::l), (S n). eauto. } + intros [l [n [Hln Hn]]]. + induction Hn. inversion Hln. + - destruct (Dec (x=a)) as [->|Ex]. + exists l, n; split; first easy. split. + econstructor; easy. easy. + eapply IHHn. destruct Hln; congruence. + - now eapply IHHn. + Qed. Lemma F_with_semi_decider: Σ f, semi_decider f F_with. Proof. @@ -135,6 +158,22 @@ Section Construction. apply Hf. Qed. + Definition χ n x: bool := Dec (In x (F_func n)). + Definition stable {Q} (f: nat → Q → bool) := + ∀ q n m, n ≤ m → f n q = true → f m q = true. + Definition stable_semi_decider {Q} P (f: nat → Q → bool) := + semi_decider (λ x n, f n x) P ∧ stable f. + + Lemma F_with_χ : stable_semi_decider F_with χ. + Proof. + split. split. + - intros [n Hn]%F_func_F_with. by exists n; apply Dec_auto. + - intros [n Hn%Dec_true]. rewrite F_func_F_with. by exists n. + - intros x n m Hn Hm%Dec_true. + apply Dec_auto. enough (incl (F_func n) (F_func m)). eauto. + eapply F_mono; last exact Hn; apply F_func_correctness. + Qed. + End Construction. Section StrongInduction. diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 570f9bc..86a5d3e 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -138,7 +138,8 @@ Section Assume_EA. Qed. Section Assume_WALL. - Variable wall: nat -> list nat -> nat -> nat. + Class Wall := wall: nat → list nat → nat → nat. + Variable wall_instance: Wall. Section Extension. diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v index beef38d..52bdfbb 100644 --- a/theories/TuringReducibility/StepIndex.v +++ b/theories/TuringReducibility/StepIndex.v @@ -49,6 +49,54 @@ Section Step_Eval. Notation Ask q := (inl q). Notation Output o := (inr o). + Lemma prefix_lookup_Some {X} (l1 l2: list X)i x : + l1 !! i = Some x → l1 `prefix_of` l2 → l2 !! i = Some x. + Proof. intros ? [k ->]. rewrite lookup_app_l; eauto using lookup_lt_Some. Qed. + + Lemma prefix_length_eq {X} (l1 l2: list X) : + l1 `prefix_of` l2 → length l2 ≤ length l1 → l1 = l2. + Proof. + intros Hprefix Hlen. assert (length l1 = length l2). + { apply prefix_length in Hprefix. lia. } + eapply list_eq_same_length with (length l1); [done..|]. + intros i x y _ ??. assert (l2 !! i = Some x) by eauto using prefix_lookup_Some. + congruence. + Qed. + + Lemma interrogation_le_eq (q: Q) (a: A) (τ: tree) R use use' ans ans' v : + functional R → + interrogation τ R use ans → τ ans =! Output v → + interrogation τ R use' ans' → τ ans' =! Output v → + length use ≤ length use' → + use = use'. + Proof. + intros funR H1 H1a H2 H2a E. + specialize (interrogation_det _ _ _ _ _ _ funR H1 H2 E) as [HE1 HE2]. + enough (length use' ≤ length use). + { eapply prefix_length_eq; eauto. } + enough (¬ length use < length use') by lia. intros H. + unshelve rewrite (interrogation_quantifiers) in H2; eauto. + destruct H2 as (H11&H12). + assert (length ans < length ans') as H'. + { unshelve rewrite (interrogation_quantifiers) in H1; eauto. + destruct H1 as [-> _]. lia. } induction HE2 as [sth Hs]. + destruct (H12 _ H') as [H12' _]. rewrite Hs in H12'. + rewrite take_app in H12'. + specialize (hasvalue_det H12' H1a). congruence. + Qed. + + Lemma interrogation_eq (q: Q) (a: A) (τ: tree) R use use' ans ans' v : + functional R → + interrogation τ R use ans → τ ans =! Output v → + interrogation τ R use' ans' → τ ans' =! Output v → + use = use'. + Proof. + intros funR H1 H1a H2 H2a. + assert (length use ≤ length use' ∨ length use' ≤ length use) as [E|E] by lia. + eapply interrogation_le_eq; eauto. + symmetry. eapply interrogation_le_eq; eauto. + Qed. + Fixpoint evalt_comp (τ : tree) (f : Q → A) (step depth: nat): option (Q + O) := match (seval (τ []) depth) with | Some x => @@ -144,6 +192,29 @@ Section Step_Eval. intros; by list_simplifier. Qed. + Lemma evalt_comp_step_mono' τ f n m o: + evalt_comp τ f n m = Some (Output o) → + evalt_comp τ f (S n) m = Some (Output o) . + Proof. + induction n in o, τ |-*. + - cbn. destruct (seval (τ []) m); last done. + by destruct s. + - intros H. cbn in H. + cbn. destruct (seval (τ []) m); last done. + destruct s; last done. + apply (IHn _ _ H). + Qed. + + Lemma evalt_comp_step_mono'' τ f n m o: + evalt_comp τ f n m = Some (Output o) → + ∀ n', n ≤ n' → evalt_comp τ f n' m = Some (Output o) . + Proof. + intros H n' Hn'. + induction Hn'; first done. + by eapply evalt_comp_step_mono'. + Qed. + + Lemma evalt_comp_step_mono (τ: tree) f qs ans o: interrogation τ (λ x y, f x = y) qs ans → τ ans =! Output o → @@ -159,10 +230,11 @@ Section Step_Eval. { eapply hasvalue_det; [|eapply H2]. rewrite seval_hasvalue. eauto. } eapply Nat.le_exists_sub in Hn' as [k [-> _]]. rewrite Nat.add_comm. - eapply interrogation_plus_evalt_comp ; eauto. + eapply interrogation_plus_evalt_comp; eauto. induction k. all: cbn; rewrite app_nil_r; by rewrite Hv, H. Qed. + Lemma evalt_comp_oracle_approx (τ: tree) f l lv v: interrogation τ (λ x y, f x = y) l lv → τ lv =! v → @@ -198,64 +270,67 @@ Section Step_Eval. Qed. Lemma evalt_comp_to_interrogation (τ : tree) (f : Q → A) o n depth: - evalt_comp τ f n depth = Some (Output o) → - Σ (qs : list Q) (ans : list A), - length qs ≤ n ∧ - interrogation τ (λ q a, f q = a) qs ans ∧ - τ ans =! Output o. + eq_dec O → eq_dec Q → + (Σ qs ans, evalt_comp τ f n depth = Some (Output o) ∧ + length qs ≤ n ∧ interrogation τ (λ q a, f q = a) qs ans ∧ τ ans =! Output o) + + + (evalt_comp τ f n depth ≠ Some (Output o)). Proof. - intros H. - induction n in τ, H |- *. + intros eq_dec_O eq_dec_Q. + destruct (Dec (evalt_comp τ f n depth = Some (Output o))) as [H|H]; last by right. + left. induction n in τ, H |- *. * cbn in *. destruct (seval (τ []) depth) eqn: E. - exists [], []. repeat split. eauto. econstructor. - destruct s. congruence. rewrite seval_hasvalue. + exists [], []. repeat split; eauto. econstructor. + destruct s. congruence. rewrite seval_hasvalue. by exists depth; injection H as ->. congruence. * cbn in *. destruct (seval (τ []) depth) eqn: E; try congruence. destruct s; try congruence. - - eapply (IHn (λ l, τ (f q :: l))) in H as (qs & ans & H3 & H1 & H2). - exists (q :: qs), (f q :: ans). split; eauto. cbn; lia. repeat split. + - eapply (IHn (λ l, τ (f q :: l))) in H as (qs & ans & H3 & H1 & H2 & H'). + exists (q :: qs), (f q :: ans). split; last split; eauto. cbn; lia. repeat split. eapply interrogation_app with (q1 := [q]) (a1 := [f q]). eapply Interrogate with (qs := []) (ans := []); eauto. econstructor. rewrite seval_hasvalue. by exists depth. eauto. eauto. - - exists [], []. repeat split. cbn. lia. eauto. econstructor. + - exists [], []. repeat split; try split; cbn. easy. lia. eauto. econstructor. rewrite seval_hasvalue. by exists depth; injection H as ->. Qed. - Lemma evalt_comp_limit_interrogation (τ: tree) f v: - (∞∀ n, evalt_comp τ (f n) n n = Some (Output v)) → - (∞∀ k, ∃ l lv, τ lv =! Output v ∧ interrogation τ (λ x y, f k x = y) l lv). + Lemma final_fact (τ: tree) n m f g use ans v: + evalt_comp τ f n m = Some (Output v) → + interrogation τ (λ x y, f x = y) use ans → + interrogation τ (λ x y, g x = y) use ans → + τ ans =! Output v → + evalt_comp τ g n m = Some (Output v). Proof. - intros [w Hw]. - exists w. intros m Hm. - specialize (Hw m Hm) as H. - destruct (evalt_comp_to_interrogation H) as (qs&ans&Hl&Hans&Hans'). - exists qs, ans. split; eauto. - Qed. + intros Hf H1 H2 Ho. + Admitted. End Step_Eval. Section Use_Function. - Definition list_interp (L: list nat) (q: nat): bool := Dec (In q L). - - Lemma try (O: Type) (τ: tree) n m (v: nat + O) L: - evalt_comp τ (list_interp L) n m = Some v → Σ k, - ∀ x, k < x → evalt_comp τ (list_interp (x::L)) n m = Some v. + Lemma use_function {Q O: Type} (τ: (list bool) ↛ (Q + O)) (f: Q → bool) n m v: + eq_dec O → eq_dec Q → + (Σ use, + evalt_comp τ f n m = Some (inr v) ∧ + ∀ P, (∀ q, q ∈ use → P q ↔ f q = true) → + ∃ ans, interrogation τ (char_rel P) use ans ∧ τ ans =! inr v) + + (evalt_comp τ f n m ≠ Some (inr v)). Proof. - induction n in v |-*. - - exists 42. intros x Hx. - unfold evalt_comp in *. by destruct (seval (τ []) m). - - intros H. destruct (evalt_comp τ (list_interp L) n m) eqn: E; last first. - admit. - destruct s; last first. - (* specialize (evalt_comp_step_mono ). *) - (* exists 42. intros H. - destruct v as [i|]. - exists (S (max i k)). intros Hi x Hx. - *) - Abort. + intros h1 h2. + destruct (evalt_comp_to_interrogation τ f v n m) as [(qs&ans&H3&_&H1&H2)|H]; try done. + left. exists qs. split; first easy. intros P HqP. + exists ans; split; last done. + eapply interrogation_ext; [eauto| |apply H1]. + intros q' [|] Hqa'%elem_of_list_In; cbn; first by rewrite HqP. + specialize (HqP q' Hqa'). split. + - intros H. destruct (f q') eqn: E; last done. + enough (P q') by easy. by rewrite HqP. + - intros H H'%HqP. congruence. + - right. done. + Qed. + End Use_Function. @@ -365,39 +440,148 @@ Section Limit_Interrogation. rewrite in_app_iff; right; econstructor. done. Qed. + End Limit_Interrogation. Section Step_Eval_Spec. - Variable P: nat → Prop. - Hypothesis Hf_: Σ f_, semi_decider f_ P. - - Definition f := projT1 (semi_decider_to_stable (projT2 Hf_)). - Fact S_P: stable_semi_decider P f. - Proof. unfold f. by destruct (semi_decider_to_stable (projT2 Hf_)). Qed. - Definition Φ_ (f: nat → nat → bool) (e x n: nat): option () := match evalt_comp (ξ () e x) (f n) n n with | Some (inr ()) => Some () | _ => None end. - Lemma Φ_spec_1 e x: - Ξ e (char_rel P) x → (∞∀ n, Φ_ f e x n = Some ()). + Definition φ (f: nat → bool) (e x n: nat) := + if use_function (ξ () e x) f n n () unit_eq_dec nat_eq_dec + is inl H then S (list_max (projT1 H)) + else 0. + + Variable P: nat → Prop. + Variable decider: nat → nat → bool. + Hypothesis S_P: stable_semi_decider P decider. + + Fact phi_iff_evalt f e x n : + Φ_ f e x n = Some () ↔ evalt_comp (ξ () e x) (f n) n n = Some (inr ()). + Proof. + unfold Φ_. destruct (evalt_comp (ξ () e x) (f n) n n) eqn: E; [destruct s|]. + - split; congruence. + - destruct u. done. + - split; congruence. + Qed. + + Theorem Φ_spec e x: + Ξ e (char_rel P) x → (∞∀ n, Φ_ decider e x n = Some ()). Proof. unfold Ξ, rel. intros (qs & ans & Ha & Hb). specialize (@S_approx_Σ1 _ _ _ S_P () (ξ _ e x) qs ans Ha) as H. eapply interrogation_evalt_comp_limit in H; last apply Hb. destruct H as [w Hw]. exists w; intros m Hm. unfold Φ_. specialize (Hw m Hm). - destruct (evalt_comp (ξ () e x) (f m) m m). + destruct (evalt_comp (ξ () e x) (decider m) m m). destruct s. by injection Hw. by destruct u. congruence. + Qed. + + Notation "A ≡{ k }≡ B" := (∀ x, x ≤ k → A x ↔ B x) (at level 30). + Definition to_pred (f: nat → bool) x := f x = true. + + Theorem φ_spec e x n p: + Φ_ decider e x n = Some () → + p ≡{φ (decider n) e x n}≡ to_pred (decider n) → + Ξ e (char_rel p) x. + Proof. + intros H2 H1. rewrite phi_iff_evalt in H2. unfold φ in H1. + destruct (use_function (ξ () e x) (decider n) n n _ _ _) as [(ans&Hans)|H]; last done. + exists ans. eapply Hans. + intros q [i Hq]%elem_of_list_lookup_1. rewrite H1; first done. + simpl. enough (q ≤ list_max (ans)) by lia. + eapply implementation.list_max_lookup. + eapply Hq. + Qed. + + Lemma φ_spec0_1 e x n: + φ (decider n) e x n ≠ 0 → Φ_ decider e x n = Some (). + Proof. + unfold φ, Φ_. intros H. + destruct (use_function (ξ () e x) (decider n) n n () unit_eq_dec nat_eq_dec). + - clear H. by destruct s as (_&->&_). + - congruence. + Qed. + + Lemma φ_spec0_2 e x n: + Φ_ decider e x n = Some () → φ (decider n) e x n ≠ 0. + Proof. + unfold φ, Φ_. intros H. + destruct (use_function (ξ () e x) (decider n) n n () unit_eq_dec nat_eq_dec). + - lia. + - destruct (evalt_comp (ξ () e x) (decider n) n n); last eauto. + destruct s; eauto. destruct u. congruence. + Qed. + + Theorem φ_spec0 e x n: φ (decider n) e x n ≠ 0 ↔ Φ_ decider e x n = Some (). + Proof. split; [apply φ_spec0_1|apply φ_spec0_2]. Qed. + + Theorem φ_spec0' e x n: φ (decider n) e x n = 0 ↔ Φ_ decider e x n = None. + Proof. + destruct (φ_spec0 e x n) as [H1 H2]. split; intros H. + - destruct (φ (decider n) e x n); eauto. + destruct (Φ_ decider e x n); eauto. + destruct u; eauto. enough (0≠0) by congruence. + by eapply H2. + - destruct (Φ_ decider e x n); try congruence. + destruct (φ (decider n) e x n); eauto. + enough (None = Some ()) by congruence. + by eapply H1. Qed. - Lemma Φ_spec: - Σ Φ_, ∀ e x, Ξ e (char_rel P) x → (∞∀ n, Φ_ f e x n = Some ()). - Proof. exists Φ_; intros e x; apply Φ_spec_1. Qed. + Fact char_rel_boring n: + ∀ q a, char_rel (decider n) q a ↔ (λ x y, decider n x = y) q a. + Proof. intros. unfold char_rel. destruct a, (decider n q); intuition. Qed. + + Theorem φ_spec1 e x n k : + φ (decider n) e x n = S k → + to_pred (decider n) ≡{k}≡ to_pred (decider (S n)) → + φ (decider (S n)) e x (S n) = S k. + Proof. + intros H H2. unfold φ in *. + destruct (use_function (ξ () e x) (decider n) n n) + as [(use & Hu1 & Hu2)|]; last congruence. simpl in *. + + assert (∀ q, q ∈ use → decider (S n) q ↔ decider n q = true) as boring1. + { intros q Hq. destruct (H2 q). apply elem_of_list_lookup_1 in Hq. + destruct Hq as [i Hi]. injection H; intros <-. + by eapply implementation.list_max_lookup. + unfold to_pred in *. destruct (decider (S n) q); intuition. } + + assert (∀ q a, In q use → char_rel (decider n) q a ↔ char_rel (decider (S n)) q a) as boring2. + { intros q a Hq. destruct (H2 q). rewrite <-elem_of_list_In in Hq. + apply elem_of_list_lookup_1 in Hq. destruct Hq as [i Hi]. + injection H; intros <-. by eapply implementation.list_max_lookup. + unfold to_pred, char_rel in *. + destruct a, (decider (S n) q), (decider n q); intuition. } + + destruct (Hu2 (decider (S n)) boring1) as [ans (Hans & Hans1)]. + + destruct (use_function (ξ () e x) (decider (S n)) (S n) (S n)) as [(use' & Hu1' & Hu2')|HSn]. + + destruct (Hu2' (decider (S n))) as [ans' (Hans' & Hans1')]. + { intros; destruct (decider (S n) q); intuition. } + enough (use = use') as Hanseq. + { cbn. rewrite <- Hanseq. apply H. } + assert (functional (char_rel (decider (S n)))) as _H_. + { intros ?[][]; unfold to_pred, char_rel; eauto. } + by edestruct (interrogation_eq 42 true _H_ Hans Hans1 Hans' Hans1'). + + exfalso; apply HSn. + assert (interrogation (ξ () e x) (λ q a, decider n q = a) use ans) as Hansn. + { eapply interrogation_ext; last apply Hans; first done. intros. + rewrite <- char_rel_boring. by apply boring2. } + assert (n ≤ S n) as _H_ by lia. + unshelve eapply (evalt_comp_depth_mono _ _H_). + eapply (evalt_comp_step_mono'). + assert (interrogation (ξ () e x) (λ q a, decider (S n) q = a) use ans) as HansSn. + { eapply interrogation_ext; last apply Hans; first done. intros; by rewrite char_rel_boring. } + clear Hu2 S_P HSn boring1 boring2 H H2 _H_ Hans. + eapply final_fact; eauto. + Qed. End Step_Eval_Spec. @@ -406,3 +590,4 @@ End Step_Eval_Spec. + From 6b826b4a2f2915e988f2f2ffbdd970b069bdfcc6 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Fri, 12 Apr 2024 13:37:05 +0200 Subject: [PATCH 24/54] fix error --- theories/Basic/Limit.v | 23 +++++++++++++++++++++-- theories/ReducibilityDegrees/low_wall.v | 2 +- 2 files changed, 22 insertions(+), 3 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index 64fcea0..bbaa618 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -102,10 +102,11 @@ Section LimitLemma1. all: eauto. Qed. +About Σ_semi_decidable_jump. (** TODO: LEM_Σ 1 <-> definite K **) (* First part of limit lemma *) - Lemma limit_turing_red_K {k: nat} (P: vec nat k -> Prop) : + Lemma limit_turing_red_K' {k: nat} (P: vec nat k -> Prop) : LEM_Σ 1 -> definite K -> limit_computable P -> @@ -116,8 +117,26 @@ Section LimitLemma1. apply Dec.nat_eq_dec. Qed. -End LimitLemma1. + Search (vec) "hd". + Fact elim_vec (P: nat -> Prop): + P ⪯ₘ (fun x: vec nat 1 => P (hd x)) . + Proof. exists (fun x => [x]). now intros x. Qed. + Lemma limit_turing_red_K {k: nat} (P: nat -> Prop) : + LEM_Σ 1 -> + definite K -> + limit_computable P -> + P ⪯ᴛ K. + Proof. + intros Hc HK [h Hh]. + eapply Turing_transitive; last eapply (@limit_turing_red_K' 1); eauto. + eapply red_m_impl_red_T. apply elim_vec. + exists (fun v n => h (hd v) n). + intros x; split; + destruct (Hh (hd x)) as [Hh1 Hh2]; eauto. + Qed. + +End LimitLemma1. Section Σ1Approximation. diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index f6c71ef..6d4e88b 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -126,7 +126,7 @@ Section Wall. - destruct H as [x [H1 H2]]. destruct (wall e (P_func x) x) as [|k] eqn: H; first done; clear H2. exists (S k), x. intros t Ht. induction Ht; first done. - rewrite <- (φ_spec1 (F_with_χ (simple_extension η wall)) IHHt). + rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). reflexivity. intros; split. + intros K%Dec_true. apply Dec_auto. From a8f78d574217155d06c5dbfdab6473ccbdabf7e8 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sat, 20 Apr 2024 18:22:34 +0200 Subject: [PATCH 25/54] final fact gen --- theories/TuringReducibility/StepIndex.v | 98 ++++++++++++++++++++++++- 1 file changed, 97 insertions(+), 1 deletion(-) diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v index 52bdfbb..f4240f1 100644 --- a/theories/TuringReducibility/StepIndex.v +++ b/theories/TuringReducibility/StepIndex.v @@ -192,6 +192,30 @@ Section Step_Eval. intros; by list_simplifier. Qed. + Lemma interrogation_plus_evalt_comp_1 (τ: tree) f n m l lv v: + interrogation τ (λ x y, f x = y) l lv → + evalt_comp τ f (length l + n) m = Some v → + evalt_comp (λ l', τ (lv ++ l')) f n m = Some v. + Proof. + intros H. revert n; dependent induction H; try eauto. + intros. rewrite app_length in H2. cbn in H2. + replace (length qs + 1 + n) with (length qs + (S n)) in H2 by lia. + eapply IHinterrogation in H2. + cbn in H2. rewrite app_nil_r in H2. + destruct (seval (τ ans) m) eqn: E; last done. + destruct s. + 2: { exfalso. enough (Output o = Ask q) by done. + eapply hasvalue_det; eauto. + eapply seval_hasvalue. by exists m. } + rewrite <- H2. eapply evalt_comp_ext. + assert (q = q0) as [=<-]. + { assert (τ ans =! Ask q0) by (eapply seval_hasvalue; by exists m). + by specialize (hasvalue_det H0 H3) as [=<-]. } + intros l_ n_. rewrite H1. + by replace ((ans ++ [a]) ++ l_) with (ans ++ a :: l_) by by list_simplifier. + Qed. + + Lemma evalt_comp_step_mono' τ f n m o: evalt_comp τ f n m = Some (Output o) → evalt_comp τ f (S n) m = Some (Output o) . @@ -296,6 +320,77 @@ Section Step_Eval. by exists depth; injection H as ->. Qed. + Lemma evalt_comp_step (τ: tree) f n m v qs ans q: + interrogation τ (λ x y, f x = y) qs ans → + length qs = n → + τ ans =! Ask q → + evalt_comp τ f n m = Some (Ask q) ∧ seval (τ (f q::ans)) m = Some v → + evalt_comp τ f (S n) m = Some v. + Proof. + intros HIn HE Hans [H1 H2]. + rewrite <- HE in *. + replace (length qs) with (length qs + 0) in H1 by lia. + eapply interrogation_plus_evalt_comp_1 in H1; last done. + induction HIn. + - simpl in *. admit. + - admit. + Admitted. + + Lemma evalt_comp_roll_back (τ: tree) f n m v qs ans q a: + interrogation τ (λ x y, f x = y) (qs++[q]) (ans++[a]) → + length qs = n → + evalt_comp τ f (S n) m = Some v ↔ + evalt_comp τ f n m = Some (Ask q) ∧ seval (τ (f q::ans)) m = Some v. + Proof. + intro H. inv H. + { exfalso. by apply app_cons_not_nil in H1. } + apply app_inj_tail in H0. destruct H0 as [H61 H62]. + apply app_inj_tail in H1. destruct H1 as [H71 H72]. subst. + intros H. subst. + split; cbn. + Admitted. + + + + Lemma final_fact_gen (τ: tree) m f g use ans: + interrogation τ (λ x y, f x = y) use ans → + interrogation τ (λ x y, g x = y) use ans → + ∀ v n, + evalt_comp τ f n m = Some v → + length use = n → + (∀ u, u ∈ use -> f u = g u) → + evalt_comp τ g n m = Some v. + Proof. + induction 1; intros. + { subst. simpl in *. done. } + assert (interrogation τ (λ x y, f x = y) (qs ++ [q]) (ans ++ [a])) by + (econstructor; eauto). + + (* prepare the hypothesis *) + rewrite last_length in *. + destruct n; first congruence. + destruct (evalt_comp_roll_back m v H6 eq_refl) as [Hc1 _]. + destruct (evalt_comp_roll_back m v H2 eq_refl) as [_ Hc2]. + assert (length qs = n) as H' by lia. subst. + destruct (Hc1 H3) as (Hc&Hcq). + + inv H2. + (* inverse the second interrogation *) + { exfalso. by apply app_cons_not_nil in H7. } + apply app_inj_tail in H1. destruct H1 as [H61 H62]. + apply app_inj_tail in H7. destruct H7 as [H71 H72]. subst. + specialize (IHinterrogation H8 (Ask q) (length qs) Hc eq_refl). + + (* inverse the goal *) + apply Hc2. split. + + eapply IHinterrogation. intros. eapply H5. + rewrite elem_of_app; by left. + + enough (g q = f q) as -> by done. + symmetry. eapply H5. + rewrite elem_of_app. right. + by rewrite elem_of_list_singleton. + Qed. + Lemma final_fact (τ: tree) n m f g use ans v: evalt_comp τ f n m = Some (Output v) → interrogation τ (λ x y, f x = y) use ans → @@ -304,6 +399,7 @@ Section Step_Eval. evalt_comp τ g n m = Some (Output v). Proof. intros Hf H1 H2 Ho. + eapply (final_fact_gen H1 H2 Hf). Admitted. End Step_Eval. @@ -579,7 +675,7 @@ Section Step_Eval_Spec. eapply (evalt_comp_step_mono'). assert (interrogation (ξ () e x) (λ q a, decider (S n) q = a) use ans) as HansSn. { eapply interrogation_ext; last apply Hans; first done. intros; by rewrite char_rel_boring. } - clear Hu2 S_P HSn boring1 boring2 H H2 _H_ Hans. + clear Hu2 S_P HSn boring1 boring2 H H2 _H_ Hans. eapply final_fact; eauto. Qed. From 7ffa62c67f05ea3ab8025c653ed57027753789c7 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 6 May 2024 20:03:27 +0200 Subject: [PATCH 26/54] state the lemma --- theories/Basic/Limit.v | 42 +++++- theories/ReducibilityDegrees/low_wall.v | 129 +++++++++++++++--- theories/ReducibilityDegrees/lowsimple.v | 28 +++- .../ReducibilityDegrees/simple_extension.v | 30 ++-- theories/TuringReducibility/StepIndex.v | 53 ++++++- 5 files changed, 231 insertions(+), 51 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index bbaa618..1797791 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -1,4 +1,4 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic PostsTheorem reductions SemiDec TuringJump OracleComputability Definitions partial. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic PostsTheorem reductions SemiDec TuringJump OracleComputability Definitions partial Pigeonhole. Require Import Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. @@ -34,7 +34,10 @@ Convention: Definition char_rel_limit_computable {X} (P: X -> bool -> Prop) := exists f: X -> nat -> bool, forall x y, P x y <-> exists N, forall n, n >= N -> f x n = y. - Lemma char_rel_limit_equv {X} (P: X -> Prop): + Definition char_rel_limit_computable' {X} (P: X -> bool -> Prop) := + exists f: X -> nat -> bool, forall x y, P x y -> exists N, forall n, n >= N -> f x n = y. + + Lemma char_rel_limit_equiv {X} (P: X -> Prop): char_rel_limit_computable (char_rel P) <-> limit_computable P. Proof. split; intros [f Hf]; exists f; intros x. @@ -42,6 +45,31 @@ Convention: - intros []; destruct (Hf x) as [h1 h2]; eauto. Qed. + Lemma char_rel_limit_equiv' {X} (P: X -> Prop): + definite P -> char_rel_limit_computable (char_rel P) <-> char_rel_limit_computable' (char_rel P) . + Proof. + intros HP; split. + - intros [f Hf]. exists f; intros. + destruct (Hf x y) as [Hf' _]. + now apply Hf'. + - intros [f Hf]. exists f. intros x y. + split. intro H. now apply Hf. + intros [N HN]. destruct (HP x). + destruct y; [easy|]. + destruct (Hf x true H) as [N' HfN]. + intros _. enough (true = false) by congruence. + specialize (HN (max N N')). + specialize (HfN (max N N')). + rewrite <- HN, <- HfN; eauto; lia. + destruct y; [|easy]. + destruct (Hf x false H) as [N' HfN]. + enough (true = false) by congruence. + specialize (HN (max N N')). + specialize (HfN (max N N')). + rewrite <- HN, <- HfN; eauto; lia. + Qed. + + (* Naming the halting problem as K *) Notation K := (­{0}^(1)). @@ -117,8 +145,6 @@ About Σ_semi_decidable_jump. apply Dec.nat_eq_dec. Qed. - Search (vec) "hd". - Fact elim_vec (P: nat -> Prop): P ⪯ₘ (fun x: vec nat 1 => P (hd x)) . Proof. exists (fun x => [x]). now intros x. Qed. @@ -192,6 +218,10 @@ Section Σ1Approximation. (forall L, exists c, forall c', c' >= c -> approximation_list P (P_ c') L) /\ (forall tau q a, @interrogation _ _ _ bool tau (char_rel P) q a -> exists n, forall m, m >= n -> interrogation tau (fun q a => P_ m q = a) q a). + Definition approximation_Σ1_weak {A} (P: A -> Prop) := + exists P_ : nat -> A -> bool, + (forall tau q a, @interrogation _ _ _ bool tau (char_rel P) q a -> exists n, forall m, m >= n -> interrogation tau (fun q a => P_ m q = a) q a). + Lemma semi_dec_approximation_Σ1 {X} (P: X -> Prop) : definite P -> semi_decidable P -> approximation_Σ1 P. @@ -277,8 +307,8 @@ Section LimitLemma2. limit_computable P. Proof. intros [F [H HF]] defK defP. - rewrite <- char_rel_limit_equv. - destruct (approximation_Σ1_halting_strong defK) as [k_ [Hk_1 Hk_2]]. + rewrite <- char_rel_limit_equiv. + destruct (approximation_Σ1_halting_strong defK) as [k_ [_ Hk_2]]. destruct H as [tau Htau]. pose (char_K_ n := char_K_ k_ n). pose (K_ n := K_ k_ n). diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 6d4e88b..02c5262 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -17,6 +17,9 @@ Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. Notation "f ↓" := (f = Some ()) (at level 30). + Global Instance dec_le: ∀ m n, dec (m ≤ n). + Proof. intros n m; destruct (le_gt_dec n m); [by left|right; lia]. Qed. + Section Requirements_Lowness_Correctness. Variable P: nat → Prop. @@ -30,13 +33,12 @@ Section Requirements_Lowness_Correctness. Definition Ω e n := Φ_ e e n. + Section classic_logic. + Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. Hypothesis LEM_Σ_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → dec (∃ n, ∀ m, P n m). - - #[export]Instance dec_le: ∀ m n, dec (m ≤ n). - Proof. intros n m; destruct (le_gt_dec n m); [by left|right; lia]. Qed. + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). Lemma limit_case e: (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). Proof. @@ -48,12 +50,12 @@ Section Requirements_Lowness_Correctness. exists w; apply Hw; exact 42. assert (∀ n, i ≤ n → ~ Ω e n ↓). { intros m Hm HM. eapply h1. exists m; eauto. } - exfalso. by eapply n; exists i. - Qed. + exfalso. eapply H0. by exists i. + Qed. Definition limit_decider e n: bool := Dec (Ω e n ↓). - Lemma Jump_limit :limit_computable (P´). + Lemma Jump_limit : limit_computable (P´). Proof. exists limit_decider; split; intros. - unfold J. split. @@ -77,6 +79,80 @@ Section Requirements_Lowness_Correctness. enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. Qed. + + End classic_logic. + + Section intuitionistic_logic. + Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → ¬¬ Ξ e (char_rel P) e. + + Lemma N_requirements': ∀ e, ¬¬ ((∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e). + Proof. firstorder. Qed. + + Fact dn_or (R Q: Prop): (¬¬ R ∨ ¬¬ Q) → ¬¬ (R ∨ Q). + Proof. firstorder. Qed. + Lemma not_not_limit_case e: ~ ~ ((∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓)). + Proof. + ccase (∃ n, ∀ m, n ≤ m → ¬ Ω e m ↓) as [H1|H1]. + apply dn_or. { right. eauto. } + ccase (∀ i, ∃ n, i ≤ n ∧ Ω e n ↓) as [H2|H2]. + intros H_. apply (@N_requirements' e). + intros N_requirements'. + apply H_. left. apply Φ_spec, N_requirements'. intros i. + { destruct (H2 i) as [w Hw]. exists w. apply Hw. } + exfalso. clear P S_P N_requirements. + apply H2. intros i. + Admitted. + + Definition not_not_limit_computable {X} (P: X -> Prop) := + exists f: X -> nat -> bool, + forall x, ~~ + ((P x <-> exists N, forall n, n >= N -> f x n = true) /\ + (~ P x <-> exists N, forall n, n >= N -> f x n = false)). + + Fact dn_and (R Q: Prop): (¬¬ R ∧ ¬¬ Q) → ¬¬ (R ∧ Q). + Proof. firstorder. Qed. + + Lemma not_not_Jump_limit : not_not_limit_computable (P´). + Proof. + exists limit_decider; intro x. + apply dn_and. + split; intros. + - unfold J. intros H_. + eapply (@N_requirements' x). intros N_requirements'. + apply H_. split. + intros [w Hw]%Φ_spec; exists w; intros??. + apply Dec_auto. by eapply Hw. + intros [N HN]. + eapply N_requirements'. + intros m. exists (S N + m); split; first lia. + eapply Dec_true. eapply HN. lia. + - unfold J. intros H_. + eapply (@N_requirements' x). intros N_requirements'. + eapply (@not_not_limit_case x). + intros [[k Hk]|h2]. + apply H_. split. + enough (Ξ x (char_rel P) x) by easy. + eapply N_requirements'. intros m. exists (S k + m). + split; first lia. eapply Hk. lia. + intro H. destruct H as [w Hw]. + intros [k' Hneg]%Φ_spec. + set (N := S (max w k')). + assert (Ω x N ↓). { eapply Hneg. lia. } + enough (¬ Ω x N ↓) by eauto. + eapply Dec_false. eapply Hw. lia. + apply H_. split. + destruct h2 as [w Hw]. exists w. + intros. specialize (Hw n H0). unfold limit_decider. + destruct (Dec _); eauto. + intro H. destruct H as [w Hw]. + intros [k Hneg]%Φ_spec. + set (N := S (max w k)). + assert (Ω x N ↓). { eapply Hneg. lia. } + enough (¬ Ω x N ↓) by eauto. + eapply Dec_false. eapply Hw. lia. + Qed. + +End intuitionistic_logic. End Requirements_Lowness_Correctness. @@ -110,9 +186,6 @@ Section Wall. by eapply H5. Qed. - Hypothesis NoIdea: ∀ (P: nat → Prop) (k: nat), - (∀ x, k ≤ x → P x) ∨ (∃ x, k ≤ x ∧ ¬ P x). - Definition χ := χ (simple_extension η wall). Definition P_Φ := (Φ_ χ). Definition P_Ω := (Ω χ). @@ -120,10 +193,14 @@ Section Wall. Fact wall_convergence e : ¬¬ ∃ b : nat, lim_to η wall (wall e) b. Proof. intros H_. - destruct (@eventally_wall e). intros [N HN]. apply H_; clear H_. - destruct (NoIdea (λ m, wall e (P_func m) m = 0) N). - - exists 0. exists N. intros. by apply H. - - destruct H as [x [H1 H2]]. + destruct (@eventally_wall e). intros [N HN]. + assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. + ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. + - apply H_; clear H_. exists 0. exists N. intros. by apply H'. + - enough (~~∃ x, N ≤ x ∧ wall e (P_func x) x ≠ 0) as H__. + apply H__. intros H''. + clear H'. destruct H'' as [x [H1 H2]]. + apply H_; clear H_. destruct (wall e (P_func x) x) as [|k] eqn: H; first done; clear H2. exists (S k), x. intros t Ht. induction Ht; first done. rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). @@ -147,12 +224,15 @@ Section Wall. enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. eapply F_uni; first apply F_func_correctness. assumption. + + intro H. + apply H'. intros x Hx. + assert (∀ n m: nat, ~~n=m → n=m). + { intros n m Hnm. destruct (Dec (n=m)); eauto. + exfalso. by apply Hnm. } + apply H0. intros Hmn. + apply H. now exists x; split. Qed. - Hypothesis DN: ∀ P, ¬ ¬ P → P. - Fact P_simple: simple P. - Proof. eapply P_simple; first apply EA. intro e. apply DN, wall_convergence. Qed. - Lemma N_requirements: ∀ e, (∞∃ n, P_Ω e n ↓) → ¬ ¬ Ξ e (char_rel P) e. Proof. intros e He H_. @@ -166,7 +246,7 @@ Section Wall. - intros (L & m & HL & HLs &HP). assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } - apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. + apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. eauto. eapply F_mono; last apply E_; apply F_func_correctness. assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). @@ -181,17 +261,22 @@ Section Wall. - intros H%Dec_true. eapply F_with_top. exists (F_func _ k), k; split; eauto. eapply F_func_correctness. - Qed. + Qed. + + (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both + eventally_wall and wall_convergence *) + Fact P_simple: simple P. + Proof. eapply P_simple; first apply EA. intro e. apply wall_convergence. Qed. Hypothesis LEM_Σ_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → dec (∃ n, ∀ m, P n m). + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). + Hypothesis DN: ∀ P, ¬ ¬ P → P. Fact jump_P_limit: limit_computable (P´). Proof. eapply Jump_limit; last done. apply F_with_χ. intros e He. eapply DN, N_requirements; eauto. Qed. - End Wall. diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index d1b800e..88b88b1 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -47,6 +47,15 @@ Section LowFacts. eapply Turing_transitive; [apply H| easy]. Qed. + Lemma DN_lowness (P: nat -> Prop) : + ~ ~ low P -> ~ K ⪯ᴛ P. + Proof. + intros H_ IH. + apply H_. intros H. + eapply not_turing_red_J with (Q := P). + eapply Turing_transitive; [apply H| easy]. + Qed. + Lemma limit_jump_lowness (A: nat -> Prop) : LEM_Σ 1 -> definite K -> @@ -76,19 +85,26 @@ Section LowFacts. Variable η: nat -> nat -> option nat. Hypothesis EA: forall P, semi_decidable P -> exists e, forall x, P x <-> exists n, η e n = Some x. - Hypothesis LEM_Σ_1: LEM_Σ 1. - Hypothesis def_K: definite K. - Hypothesis sth: forall (P: nat -> Prop) (k: nat), - (forall x, k <= x -> P x) \/ (exists x, k <= x /\ ~ P x). + + (* Use to eliminate ~~Φ *) + (* should be able to weaker to DNE_Σ_2 *) Hypothesis DN: forall P, ~ ~ P -> P. + + (* Use to prove limit computable from N requirements *) Hypothesis LEM_Σ_2: - forall (P: nat -> nat -> Prop), (forall n m, dec (P n m)) -> dec (exists n, forall m, P n m). + forall (P: nat -> nat -> Prop), + (forall n m, dec (P n m)) -> + (exists n, forall m, P n m) \/ ~ (exists n, forall m, P n m). + + (* Use to prove Limit Lemma *) + Hypothesis LEM_Σ_1: LEM_Σ 1. + Hypothesis def_K: definite K. Theorem a_sol_Post's_problem: exists P, sol_Post's_problem P. Proof. eexists. eapply low_simple_correct; split. - eapply limit_turing_red_K; eauto. exact 42. - apply jump_P_limit; eauto. + apply jump_P_limit; eauto. - eapply P_simple; eauto. Qed. diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 86a5d3e..99dfa22 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -243,7 +243,7 @@ Section Assume_EA. End Simple_Extension. - Hypothesis wall_spec: forall e, exists b, lim_to (wall e) b. + Hypothesis wall_spec: forall e, ~~ exists b, lim_to (wall e) b. Section Requirements. @@ -526,25 +526,27 @@ Section Assume_EA. Section Meet_Requirement. - Lemma wall_of_wall' e: exists w, forall x, wall e (P_func x) x < w. + Lemma wall_of_wall' e: ~~ exists w, forall x, wall e (P_func x) x < w. Proof. - destruct (wall_spec e) as [v [k Hk]]. + intro H_. apply (@wall_spec e). intros [v [k Hk]]. + apply H_. destruct (@extra_bounded (fun k => wall e (P_func k) k) k) as [w Hw]. exists (S (max v w)). intros x. destruct (Dec (x < k)). apply Hw in l. lia. assert (wall e (P_func x) x = v). apply Hk. lia. lia. Qed. - Lemma wall_of_wall e: exists w, forall i x, i <= e -> wall i (P_func x) x < w. + Lemma wall_of_wall e: + ~~ exists w, forall i x, i <= e -> wall i (P_func x) x < w. Proof. - induction e. - - destruct (wall_of_wall' 0) as [w Hw]. exists w. + intro H_. induction e. + - apply (@wall_of_wall' 0). intros [w Hw]. apply H_. exists w. intros i x ?. inv H. trivial. - - destruct IHe as [w IHe]. - destruct (wall_of_wall' (S e)) as [w' Hw]. - exists (S (max w w')). intros i x H. - inv H. specialize (Hw x). lia. - specialize (IHe i x H1). lia. + - apply IHe. intros [w IHe_]. + apply (@wall_of_wall' (S e)). intros [w' Hw']. + apply H_. exists (S (max w w')). intros i x H. + inv H. specialize (Hw' x). lia. + specialize (IHe_ i x H1). lia. Qed. Lemma attend_at_most_once_bound k: @@ -568,7 +570,7 @@ Section Assume_EA. Proof. intro H. unfold non_finite in H. intros He. rewrite non_finite_nat in H. - destruct (wall_of_wall e) as [w Hw]. + apply (@wall_of_wall e). intros [w Hw]. pose (N := max (2*e + 1) w). specialize (H N). apply H. intros [m [Hm1 [k Hmk]]]. apply He. exists k, m. @@ -665,8 +667,8 @@ Section Assume_EA. Section Instance_of_Wall. Definition low_wall (e: nat) (l: list nat) (n: nat) := 42. - Lemma low_wall_spec: forall e, exists b, lim_to low_wall (low_wall e) b. - Proof. intro e. exists 42. exists 0; intuition. Qed. + Lemma low_wall_spec: forall e, ~~ exists b, lim_to low_wall (low_wall e) b. + Proof. intro e. intros H_; apply H_. exists 42. exists 0; intuition. Qed. Definition Pw := P low_wall. diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v index f4240f1..f70de22 100644 --- a/theories/TuringReducibility/StepIndex.v +++ b/theories/TuringReducibility/StepIndex.v @@ -46,6 +46,35 @@ Section Step_Eval. Context {Q A O: Type}. Definition tree := (list A) ↛ (Q + O). + Print red_Turing. + + Definition red_Turing' (X Y : Type) (p : X → Prop) (q : Y → Prop) := + ∃ r : Functional Y bool X bool, OracleComputable r + ∧ (∀ (x : X) (b : bool), char_rel p x b -> r (char_rel q) x b). + + Lemma red_Turing_equive (X Y : Type) (p : X → Prop) (q : Y → Prop): + definite p → red_Turing p q ↔ red_Turing' p q. + Proof. + intros H. split. + - intros [r [Hr1 Hr2]]. exists r; split; eauto. + intros x b Hxb. by apply Hr2. + - intros [r [Hr1 Hr2]]. exists r; split; eauto. + intros x b. + assert (functional (r (char_rel q))). + apply OracleComputable_functional; eauto. + apply char_rel_functional. + specialize (H0 x). + split. by apply Hr2. + intros Hr. destruct b. + destruct (H x); first done. + specialize (Hr2 x false H1). + firstorder. + destruct (H x); last done. + specialize (Hr2 x true H1). + firstorder. + Qed. + + Notation Ask q := (inl q). Notation Output o := (inr o). @@ -349,9 +378,15 @@ Section Step_Eval. intros H. subst. split; cbn. Admitted. - - + Lemma evalt_comp_step_length (τ: tree) f n m v qs ans: + interrogation τ (λ x y, f x = y) qs ans → + τ ans =! Output v → + evalt_comp τ f n m = Some (Output v) → + length ans ≤ n. + Proof. + Admitted. + Lemma final_fact_gen (τ: tree) m f g use ans: interrogation τ (λ x y, f x = y) use ans → interrogation τ (λ x y, g x = y) use ans → @@ -399,6 +434,7 @@ Section Step_Eval. evalt_comp τ g n m = Some (Output v). Proof. intros Hf H1 H2 Ho. + specialize (evalt_comp_step_length H1 Ho Hf) as Hn. eapply (final_fact_gen H1 H2 Hf). Admitted. @@ -520,6 +556,17 @@ Section Limit_Interrogation. rewrite Hc; eauto. Qed. + Lemma approx_Σ1_list_rev: (∀ L, ∞∀ c, approx_list (g c) L) → definite P. + Proof. + intros H x. + destruct (H [x]) as [k Hk]. + destruct (Hk k (le_n k) x); first done. + destruct (g k x) eqn: E. + left. apply H1. easy. + right. intro H'. + by apply H0 in H'. + Qed. + Definition S_approx_Σ1_rev: definite P → ∀ O, approx_Σ1_rev O g. Proof. intros defp O τ qs ans [w Hw]. @@ -536,9 +583,9 @@ Section Limit_Interrogation. rewrite in_app_iff; right; econstructor. done. Qed. - End Limit_Interrogation. + Section Step_Eval_Spec. Definition Φ_ (f: nat → nat → bool) (e x n: nat): option () := From 0eafc42cdf47ccc3510cdfc28934e2b58129d5d4 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 12 Aug 2024 00:55:38 +0200 Subject: [PATCH 27/54] LPO -> exists p, p is low simple --- theories/Basic/Limit.v | 43 ++- theories/ReducibilityDegrees/low_wall.v | 232 ++++++++++- theories/ReducibilityDegrees/lowsimple.v | 30 +- .../ReducibilityDegrees/priority_method.v | 61 ++- .../ReducibilityDegrees/simple_extension.v | 66 +++- .../TuringReducibility/OracleComputability.v | 103 ++--- theories/TuringReducibility/StepIndex.v | 359 +++++++++++++++++- 7 files changed, 790 insertions(+), 104 deletions(-) diff --git a/theories/Basic/Limit.v b/theories/Basic/Limit.v index 1797791..66bf7f3 100644 --- a/theories/Basic/Limit.v +++ b/theories/Basic/Limit.v @@ -92,6 +92,41 @@ Section LimitLemma1. Variable list_bool_nat_inv : forall l, nat_to_list_bool (list_bool_to_nat l) = l. Variable nat_list_bool_inv : forall n, list_bool_to_nat (nat_to_list_bool n) = n. + + Section def_K. + + Hypothesis LEM_Σ_1: LEM_Σ 1. + + Lemma semi_dec_def {X} (p: X -> Prop): + semi_decidable p -> definite p. + Proof. + intros [f Hf]. unfold semi_decider in Hf. + destruct level1 as (_&H2&_). + assert principles.LPO as H by now rewrite <- H2. + intro x. destruct (H (f x)). + left. now rewrite Hf. + right. intros [k Hk]%Hf. + apply H0. now exists k. + Qed. + + Lemma def_K: definite K. + Proof. + apply semi_dec_def. + assert (isΣsem 1 (@jumpNK _ 1 1)). + eapply jump_in_Σn; eauto. + assert (@jumpNK _ 1 1 ≡ₘ ­{0}^(1)). + apply jumpNKspec. + rewrite <- semi_dec_iff_Σ1 in H. + destruct H0 as [_ [f Hf]]. + unfold reduces_m in Hf. + destruct H as [g Hg]. + unfold semi_decider in Hg. + exists (fun x => g (f x)). + split. now intros H%Hf%Hg. now intros H%Hg%Hf. + Qed. + + End def_K. + (* Extensionality of Σ2, i.e. P t iff ∃ x. ∀ y. f(x, y, t) = true *) Lemma char_Σ2 {k: nat} (P: vec nat k -> Prop) : @@ -130,10 +165,6 @@ Section LimitLemma1. all: eauto. Qed. -About Σ_semi_decidable_jump. - (** TODO: LEM_Σ 1 <-> definite K **) - (* First part of limit lemma *) - Lemma limit_turing_red_K' {k: nat} (P: vec nat k -> Prop) : LEM_Σ 1 -> definite K -> @@ -150,11 +181,11 @@ About Σ_semi_decidable_jump. Proof. exists (fun x => [x]). now intros x. Qed. Lemma limit_turing_red_K {k: nat} (P: nat -> Prop) : LEM_Σ 1 -> - definite K -> limit_computable P -> P ⪯ᴛ K. Proof. - intros Hc HK [h Hh]. + intros Hc [h Hh]. + specialize (def_K Hc) as Hk. eapply Turing_transitive; last eapply (@limit_turing_red_K' 1); eauto. eapply red_m_impl_red_T. apply elim_vec. exists (fun v n => h (hd v) n). diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 02c5262..9467fa7 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -7,7 +7,6 @@ Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. Require Import SyntheticComputability.ReducibilityDegrees.priority_method. Require Import SyntheticComputability.ReducibilityDegrees.simple_extension. - Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. Notation "'∞∃' x .. y , p" := (inf_exists (λ x, .. (inf_exists (λ y, p)) ..)) @@ -15,6 +14,7 @@ Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") : type_scope. + Notation "f ↓" := (f = Some ()) (at level 30). Global Instance dec_le: ∀ m n, dec (m ≤ n). @@ -36,9 +36,15 @@ Section Requirements_Lowness_Correctness. Section classic_logic. Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. + Definition limit_decider e n: bool := Dec (Ω e n ↓). + + Section with_LEM_2. Hypothesis LEM_Σ_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). + + Hypothesis LEM_Π_2: + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∀ n, ∃ m, P n m) ∨ ¬ (∀ n, ∃ m, P n m). Lemma limit_case e: (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). Proof. @@ -51,9 +57,7 @@ Section Requirements_Lowness_Correctness. assert (∀ n, i ≤ n → ~ Ω e n ↓). { intros m Hm HM. eapply h1. exists m; eauto. } exfalso. eapply H0. by exists i. - Qed. - - Definition limit_decider e n: bool := Dec (Ω e n ↓). + Qed. Lemma Jump_limit : limit_computable (P´). Proof. @@ -79,9 +83,40 @@ Section Requirements_Lowness_Correctness. enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. Qed. + End with_LEM_2. + + Section with_LEM_1. + + Hypothesis convergent: ∀ e, (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). + + Lemma Jump_limit_1 : limit_computable (P´). + Proof. + exists limit_decider; split; intros. + - unfold J. split. + intros [w Hw]%Φ_spec; exists w; intros??. + apply Dec_auto. by eapply Hw. + intros [N HN]. eapply N_requirements. + intros m. exists (S N + m); split; first lia. + eapply Dec_true. eapply HN. lia. + - unfold J. split; intros H. + destruct (convergent x) as [[k Hk]|h2]. + enough (Ξ x (char_rel P) x) by easy. + eapply N_requirements. intros m. exists (S k + m). + split; first lia. eapply Hk. lia. + destruct h2 as [w Hw]. exists w. + intros. specialize (Hw n H0). unfold limit_decider. + destruct (Dec _); eauto. + destruct H as [w Hw]. + intros [k Hneg]%Φ_spec. + set (N := S (max w k)). + assert (Ω x N ↓). { eapply Hneg. lia. } + enough (¬ Ω x N ↓) by eauto. + eapply Dec_false. eapply Hw. lia. + Qed. + End with_LEM_1. End classic_logic. - +(* Section intuitionistic_logic. Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → ¬¬ Ξ e (char_rel P) e. @@ -115,7 +150,7 @@ Section Requirements_Lowness_Correctness. Lemma not_not_Jump_limit : not_not_limit_computable (P´). Proof. exists limit_decider; intro x. - apply dn_and. + apply dn_and. split; intros. - unfold J. intros H_. eapply (@N_requirements' x). intros N_requirements'. @@ -152,7 +187,7 @@ Section Requirements_Lowness_Correctness. eapply Dec_false. eapply Hw. lia. Qed. -End intuitionistic_logic. +End intuitionistic_logic. *) End Requirements_Lowness_Correctness. @@ -170,6 +205,160 @@ Section Wall. Fact P_semi_decidable: semi_decidable P. Proof. eapply P_semi_decidable. Qed. + Definition χ := χ (simple_extension η wall). + Definition P_Φ := (Φ_ χ). + Definition P_Ω := (Ω χ). + + Section TEST. + + Hypothesis Σ_1_lem: LEM_Σ 1. + + Lemma attend_at_most_once_bound_constructive: + ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ attend η wall e s'. + Proof. by apply attend_at_most_once_bound_test. Qed. + + Lemma eventally_wall_test: + ∀ e, (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). + Proof. + intros e. + destruct (@attend_at_most_once_bound_constructive (S e)) as [s Hs]. + exists (S s). intros m Hm x [e_ [He_ He_']]. + destruct (Dec (e_ < e)) as [E|E]. + { exfalso. enough (attend η wall e_ m). + unshelve eapply (Hs e_ _ m); eauto; lia. + split; first lia. destruct He_' as [H _]. + apply H. } + assert (e ≤ e_) by lia; clear E. + destruct He_', H1, H1, H1, H1, H3. + by eapply H5. + Qed. + + Fact wall_convergence_test e : ∃ b : nat, lim_to η wall (wall e) b. + Proof. + destruct (@eventally_wall_test e) as [N HN]. + assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. + assert (pdec (∀ x, N ≤ x → wall e (P_func x) x = 0)) as [H'|H']. + { apply assume_Π_1_lem. apply Σ_1_lem. intros x. eauto. } + + (* ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. *) + - exists 0. exists N. intros. by apply H'. + - enough (∃ x, N ≤ x ∧ wall e (P_func x) x ≠ 0) as H''. + clear H'. destruct H'' as [x [H1 H2]]. + destruct (wall e (P_func x) x) as [|k] eqn: H; first done; clear H2. + exists (S k), x. intros t Ht. induction Ht; first done. + rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). + reflexivity. + intros; split. + + intros K%Dec_true. apply Dec_auto. + enough (incl (F_func (simple_extension η wall) m) (F_func (simple_extension η wall) (S m))). + eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension η wall) (S m)) as HE. + inv HE. + * assert (wall e (P_func m) m < a). + { eapply HN. lia. enough (P_func m = l) as ->. apply H5. + eapply F_uni; [apply F_func_correctness|apply H4]. } + assert (wall e (P_func m) m = S k). { rewrite <-IHHt. reflexivity. } + rewrite H6 in H2. destruct (Dec (a = x0)). + lia. apply Dec_auto. rewrite <- H3 in K. + destruct K; first done. + enough ((F_func (simple_extension η wall) m) = l) as -> by done. + eapply F_uni; last apply H4; apply F_func_correctness. + * apply Dec_auto. + enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. + eapply F_uni; first apply F_func_correctness. + assumption. + + eapply assume_Σ_1_dne. apply Σ_1_lem. + { intro x. eauto. } + intro H. apply H'. intros x Hx. + assert (∀ n m: nat, ~~n=m → n=m). + { intros n m Hnm. destruct (Dec (n=m)); eauto. + exfalso. by apply Hnm. } + apply H0. intros Hmn. + apply H. now exists x; split. + Qed. + + Lemma N_requirements_test: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. + Proof. + intros e He. + destruct (@eventally_wall_test e) as [N HN]. + destruct (@wall_convergence_test e) as [B [b Hb]]. + set (M := max N b). destruct (He M) as [k [Hk Hk']]. + eapply (@φ_spec χ e e k); first apply Hk'. + intros x Hx. unfold P, simple_extension.P. + rewrite F_with_top. split. + - intros (L & m & HL & HLs &HP). + assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } + assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } + apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. + enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. + eauto. eapply F_mono; last apply E_; apply F_func_correctness. + assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). + assert (k ≤ m) as Ek by lia. enough (x ≤ φ (χ m) e e m). + exfalso. assert (φ (χ m) e e m < x). + apply HN. lia. enough (φ (χ m) e e m = φ (χ k) e e k) by congruence. + assert (φ (χ m) e e m = B). + { rewrite <- (Hb m). reflexivity. lia. } + assert (φ (χ k) e e k = B). + { rewrite <- (Hb k). reflexivity. lia. } + congruence. + - intros H%Dec_true. + eapply F_with_top. exists (F_func _ k), k; split; eauto. + eapply F_func_correctness. + Qed. + + Lemma convergent e : (∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓). + Proof. + destruct (@eventally_wall_test e) as [N HN]. + assert (pdec (∃ k, N ≤ k ∧ P_Ω e k ↓)) as [[k [Hk HNk]]|H'] by (apply assume_Σ_1_lem; eauto). + - left. exists k. intros n Hm. + induction Hm; first done. + unfold P_Ω, Ω, Φ_ in *. + destruct (φ (χ m) e e m) eqn: E. + { eapply φ_spec0' in E. congruence. } + (* Check φ_spec2. *) + eapply (@φ_spec2 χ _ ); eauto. + intros x Hx; split. + + intros K%Dec_true. apply Dec_auto. + enough (incl (F_func (simple_extension η wall) m) (F_func (simple_extension η wall) (S m))). + eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension η wall) (S m)) as HE. + inv HE. + * assert (wall e (P_func m) m < a). + { eapply HN. lia. enough (P_func m = l) as ->. apply H2. + eapply F_uni; [apply F_func_correctness|apply H1]. } + assert (wall e (P_func m) m = S n). { unfold wall, wall_instance. + rewrite <-E. reflexivity. } + rewrite H3 in H. destruct (Dec (a = x)). + lia. apply Dec_auto. rewrite <- H0 in K. + destruct K; first done. + enough ((F_func (simple_extension η wall) m) = l) as -> by done. + eapply F_uni; last apply H1; apply F_func_correctness. + * apply Dec_auto. + enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. + eapply F_uni; first apply F_func_correctness. + assumption. + - right. exists N. intros m Hm. + destruct (Dec (P_Ω e m ↓)); eauto. + Qed. + + (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both + eventally_wall and wall_convergence *) + Fact P_simple_test: simple P. + Proof. eapply P_simple; first apply EA. intro e. + intros H. apply H. apply wall_convergence_test. Qed. + + Hypothesis Σ_2_LEM: + ∀ (P: nat → nat → Prop), + (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). + + Fact jump_P_limit_test: limit_computable (P´). + Proof. + eapply Jump_limit; last done. apply F_with_χ. + intros e He. eapply N_requirements_test; eauto. + Qed. + End TEST. + + Lemma eventally_wall e: ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). Proof. @@ -186,10 +375,6 @@ Section Wall. by eapply H5. Qed. - Definition χ := χ (simple_extension η wall). - Definition P_Φ := (Φ_ χ). - Definition P_Ω := (Ω χ). - Fact wall_convergence e : ¬¬ ∃ b : nat, lim_to η wall (wall e) b. Proof. intros H_. @@ -268,8 +453,10 @@ Section Wall. Fact P_simple: simple P. Proof. eapply P_simple; first apply EA. intro e. apply wall_convergence. Qed. + Section with_LEM_2. + Hypothesis LEM_Σ_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). + ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). Hypothesis DN: ∀ P, ¬ ¬ P → P. Fact jump_P_limit: limit_computable (P´). @@ -277,7 +464,26 @@ Section Wall. eapply Jump_limit; last done. apply F_with_χ. intros e He. eapply DN, N_requirements; eauto. Qed. + + End with_LEM_2. + + Section with_LEM_1. + + Hypothesis LEM_Σ_1: LEM_Σ 1. + + Fact jump_P_limit_2: limit_computable (P´). + Proof. + eapply Jump_limit_1; first apply F_with_χ. + - intros e He. eapply N_requirements_test; eauto. + - eapply convergent; eauto. + Qed. + + End with_LEM_1. + End Wall. +(* Check jump_P_limit_2. *) + + diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index 88b88b1..ad4d4f0 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -5,12 +5,15 @@ Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Export SyntheticComputability.ReducibilityDegrees.low_wall. Require Import Arith. +Require Import SyntheticComputability.PostsTheorem.PostsTheorem. Require Import Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. +Require Import stdpp.list. Local Notation vec := Vector.t. + (* ########################################################################## *) (** * Low Simple Set *) (* ########################################################################## *) @@ -78,6 +81,8 @@ Section LowFacts. split; [destruct H2 as [H2 _]; eauto| now apply lowness]. Qed. + + (*** Instance of low simple predicate ***) Section LowSimplePredicate. @@ -86,17 +91,17 @@ Section LowFacts. Hypothesis EA: forall P, semi_decidable P -> exists e, forall x, P x <-> exists n, η e n = Some x. - (* Use to eliminate ~~Φ *) + (* Used to eliminate ~~Φ *) (* should be able to weaker to DNE_Σ_2 *) Hypothesis DN: forall P, ~ ~ P -> P. - (* Use to prove limit computable from N requirements *) + (* Used to prove limit computable from N requirements *) Hypothesis LEM_Σ_2: forall (P: nat -> nat -> Prop), (forall n m, dec (P n m)) -> (exists n, forall m, P n m) \/ ~ (exists n, forall m, P n m). - (* Use to prove Limit Lemma *) + (* Used to prove Limit Lemma *) Hypothesis LEM_Σ_1: LEM_Σ 1. Hypothesis def_K: definite K. @@ -110,5 +115,22 @@ Section LowFacts. End LowSimplePredicate. -End LowFacts. + Section LowSimplePredicate2. + + Variable η: nat → nat → option nat. + Hypothesis EA: + forall P, semi_decidable P → exists e, forall x, P x ↔ exists n, η e n = Some x. + + Hypothesis LEM_Σ_1: LEM_Σ 1. + + Theorem a_sol_Post's_problem_2: exists P, sol_Post's_problem P. + Proof. + eexists. eapply low_simple_correct; split. + - eapply limit_turing_red_K; eauto. exact 42. + apply jump_P_limit_2; eauto. + - eapply P_simple; eauto. + Qed. + + End LowSimplePredicate2. +End LowFacts. \ No newline at end of file diff --git a/theories/ReducibilityDegrees/priority_method.v b/theories/ReducibilityDegrees/priority_method.v index 8ea67c1..d67b4a5 100644 --- a/theories/ReducibilityDegrees/priority_method.v +++ b/theories/ReducibilityDegrees/priority_method.v @@ -69,6 +69,14 @@ Section Construction. - eapply IHF_; eauto. Qed. + Lemma F_pick' n x l: F_ E n (x::l) -> exists m, m < n ∧ F_ E m l ∧ extendP l m x. + Proof. + intros H. dependent induction H. + - exists n; eauto. + - destruct (IHF_ x l eq_refl) as [m (Hm1&Hm2&Hm3)]. + exists m; eauto. + Qed. + Lemma F_computable : Σ f: nat -> list nat, forall n, F_ E n (f n) /\ length (f n) <= n. Proof. @@ -444,4 +452,55 @@ Section LeastWitness. right. intro H'. apply H1. intros. apply H'. lia. Qed. -End LeastWitness. \ No newline at end of file +End LeastWitness. + + +Section logic. + + Definition pdec p := p ∨ ¬ p. + + Definition Π_1_lem := ∀ p : nat -> Prop, + (∀ x, dec (p x)) -> pdec (∀ x, p x). + Definition Σ_1_dne := ∀ p : nat -> Prop, + (∀ x, dec (p x)) -> (¬¬ ∃ x, p x) → ∃ x, p x. + Definition Σ_1_lem := ∀ p: nat → Prop, + (∀ x, dec (p x)) -> pdec (∃ x, p x). + + Hypothesis LEM1: LEM_Σ 1. + + Fact assume_Σ_1_lem: Σ_1_lem . + Proof. + intros p Hp. + destruct level1 as (_&H2&_). + assert principles.LPO as H by by rewrite <- H2. + destruct (H Hp) as [[k Hk]|H1]. + - left. exists k. destruct (Hp k); eauto. + cbn in Hk. congruence. + - right. intros [k Hk]. apply H1. + exists k. destruct (Hp k); eauto. + Qed. + + Fact assume_Σ_1_dne: Σ_1_dne. + Proof. + intros p Hp H. + destruct (assume_Σ_1_lem Hp) as [H1|H1]; eauto. + exfalso. by apply H. + Qed. + + Fact assume_Π_1_lem: Π_1_lem. + Proof. + intros p Hp. + destruct level1 as (_&H2&_). + assert principles.LPO as H by by rewrite <- H2. + apply principles.LPO_to_WLPO in H. + assert (∀ x : nat, dec (¬ p x)) as Hp' by eauto. + destruct (H Hp') as [H1|H1]. + - left. intro x. + specialize (H1 x). + apply Dec_false in H1. + destruct (Hp x); firstorder. + - right. intros H3. apply H1. intros n. + specialize (H3 n). destruct (Hp' n); firstorder. + Qed. + +End logic. diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 99dfa22..8103d78 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -147,7 +147,7 @@ Section Assume_EA. Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. - Definition ext_least_choice L n x := exists e, ext_choice L n e x. + Definition ext_least_choice L n x := exists e, ext_choice L n e x. End Extension. @@ -251,6 +251,14 @@ Section Assume_EA. Definition attend e n := e < n /\ least (ext_pick (P_func n) n) e. Definition act e n := ~ (P_func n) # W[n] e. Definition pick_el e x := exists k, attend e k /\ ext_least_choice (P_func k) k x. + Definition finish e n := ∀ s, n < s → ¬ attend e s. + + #[export]Instance attend_dec e n: dec (attend e n). + Proof. + unfold attend. apply and_dec; first eauto. + eapply least_dec. intros y. + eapply ext_pick_dec. + Qed. End Requirements. @@ -333,6 +341,20 @@ Section Assume_EA. apply H. now exists k'. Qed. + Definition done e s := ∀ s', s < s' → ¬ attend e s'. + + Lemma attend_at_most_once_test: + LEM_Σ 1 → ∀ e, ∃ s, done e s. + Proof. + intros Hlem e. + assert (pdec (exists k, attend e k)) as [[w Hw]|H]. + { eapply assume_Σ_1_lem. apply Hlem. eauto. } + - exists w. unfold done. by eapply attend_always_not_attend. + - exists 0. + intros k' Hk' Ha. + apply H. now exists k'. + Qed. + Lemma attend_uni e: unique (attend e). Proof. intros k1 k2 H1 H2. @@ -564,6 +586,21 @@ Section Assume_EA. intros s' Hs'. eapply Hs; lia. Qed. + + Lemma attend_at_most_once_bound_test: + LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k -> ∀ s', s < s' -> ~ attend e s'). + Proof. + intros Hlem. induction k. + - exists 42. intros ??. lia. + - destruct IHk as [s Hs]. + specialize (@attend_at_most_once_test Hlem k) as [sk Hsk]. + set (max sk s) as N. + exists N. intros e He. + assert (e = k \/ e < k) as [->|Hek] by lia. + intros s' Hs'. eapply Hsk. lia. + intros s' Hs'. eapply Hs; lia. + Qed. + Lemma non_finite_not_bounded e: non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ (forall n, forall i, i <= e -> wall i (P_func n) n < x). @@ -579,18 +616,18 @@ Section Assume_EA. Qed. Lemma ext_pick_attend N e: - e < N -> (exists w, w < e /\ ext_pick (P_func N) N w) -> - (exists w, w < e /\ attend w N). + e < N -> ext_pick (P_func N) N e -> + (exists w, w ≤ e /\ attend w N). Proof. - intros HeN [w (Hw1 & Hw2)]. - assert (exists w, ext_pick (P_func N) N w) by now exists w. + intros HeN H1. + assert (exists w, ext_pick (P_func N) N w) by now exists e. eapply least_ex in H; last eauto. - destruct H as [k Hk]. assert (k <= w). - { enough (~ k > w) by lia. intro Hkw. + destruct H as [k Hk]. assert (k <= e). + { enough (~ k > e) by lia. intro Hkw. destruct Hk. rewrite safe_char in H0. - specialize (H0 w Hw2). lia. } + specialize (H0 e H1). lia. } exists k. do 2 (split; first lia). eapply Hk. - Qed. + Qed. Lemma non_finite_attend e: non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ attend e k) . @@ -608,10 +645,9 @@ Section Assume_EA. } split. lia. split. easy. intros w HEw Hw. - assert (exists w, w < e /\ ext_pick (P_func N) N w). - now exists w. eapply ext_pick_attend in H1. - destruct H1 as [g [HEg Hg]]. - eapply (Hs g HEg); last apply Hg. lia. lia. + eapply ext_pick_attend in Hw. + destruct Hw as [g [HEg Hg]]. + eapply Hs; last apply Hg; lia. lia. - exists N. now left. Qed. @@ -663,7 +699,7 @@ Section Assume_EA. End Result. End Assume_WALL. - +(* Section Instance_of_Wall. Definition low_wall (e: nat) (l: list nat) (n: nat) := 42. @@ -726,7 +762,7 @@ Section Assume_EA. admit. admit. *) Admitted. - End Instance_of_Wall. + End Instance_of_Wall. *) End Assume_EA. diff --git a/theories/TuringReducibility/OracleComputability.v b/theories/TuringReducibility/OracleComputability.v index e59ed86..5d24455 100644 --- a/theories/TuringReducibility/OracleComputability.v +++ b/theories/TuringReducibility/OracleComputability.v @@ -828,11 +828,11 @@ Qed. (** A total computable version of evalt **) Fixpoint evalt_comp {Q A O} (tau : list A ↛ (Q + O)) - (f : Q -> A) (n : nat) (step: nat): option (Q + O) := - match (seval (tau []) step) with - | Some x => match n, x with + (f : Q -> A) (step : nat) (depth: nat): option (Q + O) := + match (seval (tau []) depth) with + | Some x => match step, x with | 0, inl q => Some (inl q) - | S n, inl q => evalt_comp (fun l => tau (f q :: l)) f n step + | S n, inl q => evalt_comp (fun l => tau (f q :: l)) f n depth | _, inr o => Some (inr o) end | None => None end. @@ -895,7 +895,7 @@ Qed. (** Basic property of evalt_comp **) -Lemma evalt_comp_step_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : +Lemma evalt_comp_depth_mono {Q A O} (tau: (list A) ↛ (Q + O)) f n m o : evalt_comp tau f n m = Some o -> forall m', m' >= m -> evalt_comp tau f n m' = Some o. Proof. @@ -916,7 +916,7 @@ Proof. + congruence. Qed. -Lemma interrogation_plus_comp {Q A O} tau f n m l lv v2: +Lemma interrogation_plus_evalt_comp {Q A O} tau f n m l lv v2: @interrogation Q A O tau (fun x y => f x = y) l lv -> (forall ans_, prefix ans_ lv -> exists v, seval (tau ans_) m = Some v) -> evalt_comp (fun l' => tau (lv ++ l')) f n m = Some v2 <-> @@ -965,12 +965,12 @@ Proof. intros; now list_simplifier. Qed. -Lemma evalt_comp_step_length_mono {Q A O} (tau: (list A) ↛ (Q + O)) f qs ans o: +Lemma evalt_comp_step_mono {Q A O} (tau: (list A) ↛ (Q + O)) f qs ans o: @interrogation Q A O tau (fun x y => f x = y) qs ans -> tau ans =! Output o -> - exists step n, + exists depth step, forall g, @interrogation Q A O tau (fun x y => g x = y) qs ans -> - forall n', n <= n' -> evalt_comp tau g n' step = Some (inr o). + forall n', step <= n' -> evalt_comp tau g n' depth = Some (Output o). Proof. intros H1 H2. destruct (interrogation_ter _ _ _ _ _ H1 H2) as [step Hstep]. @@ -981,57 +981,38 @@ Proof. { eapply hasvalue_det; [|eapply H2]. rewrite seval_hasvalue. eauto. } eapply Nat.le_exists_sub in Hn' as [k [-> _]]. rewrite Nat.add_comm. - eapply interrogation_plus_comp; eauto. + eapply interrogation_plus_evalt_comp; eauto. induction k. all: cbn; rewrite app_nil_r; by rewrite Hv, H. Qed. -Lemma interrogation_evalt_comp {Q A O} tau f l lv v: +Lemma evalt_comp_oracle_approx {Q A O} tau f l lv v: @interrogation Q A O tau (fun x y => f x = y) l lv -> tau lv =! v -> - exists n step, evalt_comp tau f n step = Some v. -Proof. - intros H1 h2. - destruct (interrogation_ter _ _ _ _ _ H1 h2) as [step Hstep]. - exists (length l + 0). - exists step. eapply interrogation_plus_comp; eauto. - cbn. rewrite app_nil_r. - destruct (Hstep lv) as [v' Hv']. - reflexivity. - assert (exists k, seval (A:=Q + O) (tau lv) k = Some v'). - exists step. easy. - rewrite <- seval_hasvalue in H. - assert (v' = v). - eapply hasvalue_det; eauto. - rewrite Hv', H0. - destruct v; done. -Qed. - -Lemma evalt_comp_index_mono {Q A O} tau f l lv v: - @interrogation Q A O tau (fun x y => f x = y) l lv -> - tau lv =! v -> - exists a b, + exists step depth, forall g, @interrogation Q A O tau (fun x y => g x = y) l lv -> - evalt_comp tau g a b = Some v. + evalt_comp tau g step depth = Some v. Proof. intros H1 h2. destruct (interrogation_ter _ _ _ _ _ H1 h2) as [step Hstep]. exists (length l + 0). exists step. intros. - eapply interrogation_plus_comp; eauto. - cbn. rewrite app_nil_r. - destruct (Hstep lv) as [v' Hv']. - reflexivity. - assert (exists k, seval (A:=Q + O) (tau lv) k = Some v'). - exists step. easy. - rewrite <- seval_hasvalue in H0. - assert (v' = v). - eapply hasvalue_det; eauto. - rewrite Hv', H2. - destruct v; done. + eapply interrogation_plus_evalt_comp; eauto. + cbn. rewrite app_nil_r. + destruct (Hstep lv) as [v' Hv']. + reflexivity. + assert (exists k, seval (A:=Q + O) (tau lv) k = Some v'). + exists step. easy. + rewrite <- seval_hasvalue in H0. + assert (v' = v). + eapply hasvalue_det; eauto. + rewrite Hv', H2. + destruct v; done. Qed. + + Lemma interrogation_evalt_comp_limit {Q A O} tau f l lv v: (exists K, forall k, k >= K -> @interrogation Q A O tau (fun x y => (f k) x = y) l lv) -> @@ -1041,16 +1022,42 @@ Proof. intros [k h1] h2. assert (interrogation tau (fun x y => f k x = y) l lv) as H. apply h1. lia. - destruct (evalt_comp_step_length_mono _ _ _ _ _ H h2) as (a' & b' & Hs). - destruct (evalt_comp_index_mono _ _ _ _ _ H h2) as (a & b & Hab). + destruct (evalt_comp_step_mono _ _ _ _ _ H h2) as (a' & b' & Hs). + destruct (evalt_comp_oracle_approx _ _ _ _ _ H h2) as (a & b & Hab). exists (max b'(max a' (max (max a b) k))). intros n Hn. - eapply evalt_comp_step_mono. + eapply evalt_comp_depth_mono. eapply (Hs (f n)); eauto. eapply h1. all: lia. Qed. +Lemma evalt_comp_to_interrogation: + ∀ {Q A I O : Type} (tau : I → (list A) ↛ (Q + O)) (f : Q -> A) (i : I) (o : O) (n depth: nat), + evalt_comp (tau i) f n depth = Some (Output o) → + ∃ (qs : list Q) (ans : list A), + length qs <= n /\ interrogation (tau i) (λ (x : Q) (y : A), f x = y) qs ans ∧ + tau i ans =! Output o. +Proof. + intros Q A I O tau f i o n depth H. + induction n in tau, H |- *. + * cbn in *. destruct (seval (tau i []) depth) eqn: E. + exists [], []. repeat split. eauto. econstructor. + destruct s. congruence. rewrite seval_hasvalue. + by exists depth; injection H as ->. congruence. + * cbn in *. destruct (seval (tau i []) depth) eqn: E; try congruence. + destruct s; try congruence. + -- eapply (IHn (fun i l => tau i (f q :: l))) in H as (qs & ans & H3 & H1 & H2). + exists (q :: qs), (f q :: ans). split; eauto. cbn; lia. repeat split. + eapply interrogation_app with (q1 := [q]) (a1 := [f q]). + eapply Interrogate with (qs := []) (ans := []); eauto. + rewrite seval_hasvalue. by exists depth. + eauto. eauto. + -- exists [], []. repeat split. cbn. lia. eauto. + rewrite seval_hasvalue. + by exists depth; injection H as ->. +Qed. + (** ** Closure properties of Oracle computability *) diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v index f70de22..9c91f0e 100644 --- a/theories/TuringReducibility/StepIndex.v +++ b/theories/TuringReducibility/StepIndex.v @@ -137,6 +137,14 @@ Section Step_Eval. | None => None end. + Fact evalt_comp_tail (τ: tree) f m ans v: + seval (τ (ans ++ [])) m = Some (Output v) → + evalt_comp (λ l' : list A, τ (ans ++ l')) f 0 m = Some (Output v) . + Proof. + intros H. cbn; destruct (seval (τ (ans ++ [])) m) eqn: E; last done. + destruct s; done. + Qed. + Lemma evalt_comp_ext (τ τ': tree) f n m: (∀ l n, seval (τ l) n = seval (τ' l) n) → evalt_comp τ f n m = evalt_comp τ' f n m. @@ -266,7 +274,6 @@ Section Step_Eval. induction Hn'; first done. by eapply evalt_comp_step_mono'. Qed. - Lemma evalt_comp_step_mono (τ: tree) f qs ans o: interrogation τ (λ x y, f x = y) qs ans → @@ -287,7 +294,6 @@ Section Step_Eval. induction k. all: cbn; rewrite app_nil_r; by rewrite Hv, H. Qed. - Lemma evalt_comp_oracle_approx (τ: tree) f l lv v: interrogation τ (λ x y, f x = y) l lv → τ lv =! v → @@ -314,11 +320,11 @@ Section Step_Eval. intros [k h1] h2. assert (interrogation τ (λ x y, f k x = y) l lv) as H. apply h1. lia. - destruct (evalt_comp_step_mono H h2) as (a' & b' & Hs). destruct (evalt_comp_oracle_approx H h2) as (a & b & Hab). - exists (max b'(max a' (max (max a b) k))). + exists (max b (max a k)). intros n Hn. eapply evalt_comp_depth_mono. - eapply (Hs (f n)); eauto. eapply h1. + eapply evalt_comp_step_mono''. + eapply (Hab (f n)); eauto. eapply h1. all: lia. Qed. @@ -349,7 +355,7 @@ Section Step_Eval. by exists depth; injection H as ->. Qed. - Lemma evalt_comp_step (τ: tree) f n m v qs ans q: + (* Lemma evalt_comp_step (τ: tree) f n m v qs ans q: interrogation τ (λ x y, f x = y) qs ans → length qs = n → τ ans =! Ask q → @@ -363,30 +369,191 @@ Section Step_Eval. induction HIn. - simpl in *. admit. - admit. - Admitted. + Admitted. *) + Fact sub_tree (τ: tree) f use ans ans_: + interrogation τ (λ x y, f x = y) use ans → + ans_ `prefix_of` ans → + ∃ use_ : list Q, interrogation τ (λ x y, f x = y) use_ ans_ ∧ + length use_ = length ans_. + Proof. + intros H1 [Hl]. + dependent induction H1. + - assert (ans_ = []). + { eapply prefix_nil_inv; eauto. exists (Hl). done. } subst. + exists []. split. econstructor. done. + - destruct Hl. list_simplifier. exists (qs ++ [q] ). + split. econstructor; eauto. rewrite !app_length. + cbn. f_equal. eapply interrogation_length. done. + list_simplifier. + apply app_eq_inv in H2. destruct H2 as [[k [Hk Hk']]|[k [Hk Hk']]]. + + eapply (IHinterrogation k). done. + + assert (k = []). destruct k; first done. + list_simplifier. by eapply app_cons_not_nil in H2. + list_simplifier. eapply (IHinterrogation []). + list_simplifier. done. + Qed. + + Fact sub_tree_2 (τ: tree) f use ans use_: + interrogation τ (λ x y, f x = y) use ans → + use_ `prefix_of` use → + ∃ ans_ : list A, interrogation τ (λ x y, f x = y) use_ ans_ ∧ + length use_ = length ans_. + Proof. + intros H1 [Hl]. + dependent induction H1. + - assert (use_ = []). + { eapply prefix_nil_inv; eauto. exists (Hl). done. } subst. + exists []. split. econstructor. done. + - destruct Hl. list_simplifier. exists (ans ++ [f q] ). + split. econstructor; eauto. rewrite !app_length. + cbn. f_equal. eapply interrogation_length. done. + list_simplifier. + apply app_eq_inv in H2. destruct H2 as [[k [Hk Hk']]|[k [Hk Hk']]]. + + eapply (IHinterrogation k). done. + + assert (k = []). destruct k; first done. + list_simplifier. by eapply app_cons_not_nil in H2. + list_simplifier. eapply (IHinterrogation []). + list_simplifier. done. + Qed. + + Fact sub_tree_3 (τ: tree) f use ans use_: + interrogation τ (λ x y, f x = y) use ans → + use_ `prefix_of` use → + length use_ < length use → + ∃ ans_ : list A, interrogation τ (λ x y, f x = y) use_ ans_ ∧ + length use_ = length ans_ + ∧ ∃ q, τ ans_ =! Ask q. + Proof. + intros H1 [Hl] Hlen. + dependent induction H1. + - assert (use_ = []). + { eapply prefix_nil_inv; eauto. exists (Hl). done. } subst. + cbn in Hlen. lia. + - destruct Hl. list_simplifier. lia. + list_simplifier. + apply app_eq_inv in H2. destruct H2 as [[k [Hk Hk']]|[k [Hk Hk']]]. + + rewrite Hk in Hlen. destruct k. + 2:{ apply (IHinterrogation (q1::k)). done. list_simplifier. + rewrite app_length. cbn. lia. } + list_simplifier. subst. exists ans. split. done. + split. eapply interrogation_length. done. + exists q. done. + + assert (k = []). destruct k; first done. + list_simplifier. by eapply app_cons_not_nil in H2. + list_simplifier. + exists ans. split. done. + split. eapply interrogation_length. done. + exists q0. done. + Qed. + + + Fact sub_computation_ter (τ: tree) n m f use ans v : + length ans ≤ n → + evalt_comp τ f n m = Some v → + interrogation τ (λ x y, f x = y) use ans → + ∀ ans_, ans_ `prefix_of` ans → + ∃ v : Q + O, seval (τ ans_) m = Some v. + Proof. + intros Hn Hcomp Hin ans_ Hans_. + assert (∃ k, length ans_ + k = n) as [k Hk]. + { apply prefix_length in Hans_. exists (n - length ans_). lia. } + rewrite <-Hk in Hcomp. + specialize (sub_tree Hin Hans_) as (use_ & Huse_ & HEU). + rewrite <- HEU in Hcomp. + specialize (interrogation_plus_evalt_comp_1 Huse_ Hcomp) as H. + { destruct k. cbn in H. destruct (seval (τ (ans_ ++ [])) m) eqn: E; last done. + list_simplifier. by exists s. + cbn in H. destruct (seval (τ (ans_ ++ [])) m) eqn: E; last done. + list_simplifier. by exists s. + } + Qed. Lemma evalt_comp_roll_back (τ: tree) f n m v qs ans q a: interrogation τ (λ x y, f x = y) (qs++[q]) (ans++[a]) → length qs = n → evalt_comp τ f (S n) m = Some v ↔ - evalt_comp τ f n m = Some (Ask q) ∧ seval (τ (f q::ans)) m = Some v. + evalt_comp τ f n m = Some (Ask q) ∧ seval (τ (ans ++ [f q])) m = Some v. Proof. intro H. inv H. { exfalso. by apply app_cons_not_nil in H1. } apply app_inj_tail in H0. destruct H0 as [H61 H62]. apply app_inj_tail in H1. destruct H1 as [H71 H72]. subst. - intros H. subst. - split; cbn. - Admitted. + intros H. subst. split. + - intros H. + assert (evalt_comp τ f (S (length qs)) m = Some v) as Hcopy by eauto. + replace (S (length qs)) with (length qs + 1) in H by lia. + eapply interrogation_plus_evalt_comp_1 in H; eauto. + replace ((length qs)) with (length qs + 0) by lia. + rewrite <- interrogation_plus_evalt_comp; eauto. + cbn in *. destruct (seval (τ (ans ++ [])) m) eqn: E; last done. + split. + destruct s; eauto. + + { f_equal. eapply hasvalue_det; eauto. eapply seval_hasvalue. exists m. + list_simplifier. done. } + { f_equal. eapply hasvalue_det; eauto. eapply seval_hasvalue. exists m. + list_simplifier. done. } + destruct s; last first. + { exfalso. list_simplifier. assert (Output o = Ask q). + eapply hasvalue_det. eapply seval_hasvalue. exists m. done. done. congruence. } + enough (q = q0) as <-. + destruct (seval (τ (ans ++ [f q])) m); last done. + destruct s; eauto. + { list_simplifier. assert (τ ans =! Ask q0). eapply seval_hasvalue; eauto. + specialize (hasvalue_det H3 H0). intros [=]. eauto. } + unshelve eapply (sub_computation_ter _ Hcopy); eauto. + rewrite (interrogation_length H2); lia. + + - intros [He Ho]. + replace (S (length qs)) with (length qs + 1) by lia. + eapply interrogation_plus_evalt_comp; eauto. + unshelve eapply (sub_computation_ter _ He); eauto. + rewrite (interrogation_length H2); lia. + replace ((length qs)) with (length qs + 0) in He by lia. + eapply interrogation_plus_evalt_comp_1 in He ; eauto. + cbn in *. destruct (seval (τ (ans ++ [])) m); last done. + destruct s; last congruence. + injection He. intros ->. + destruct (seval (τ (ans ++ [f q])) m); last done. + destruct s; last congruence. + congruence. + Qed. + + Fact boring_decomposition {X} (qs: list X) n: + n < length qs → + ∃ q other qs_, length qs_ = n ∧ qs = qs_ ++ q :: other. + Proof. + intros H. + induction qs in n, H |- *. + - cbn in *. lia. + - destruct n. + + cbn. exists a, qs, []. done. + + cbn. destruct (IHqs n) as [q [other [qs_ [H1 H2]]]]. cbn in H. lia. + exists q, other, (a :: qs_); split. + cbn. lia. list_simplifier. done. + Qed. Lemma evalt_comp_step_length (τ: tree) f n m v qs ans: interrogation τ (λ x y, f x = y) qs ans → τ ans =! Output v → evalt_comp τ f n m = Some (Output v) → - length ans ≤ n. + length qs ≤ n. Proof. - Admitted. - + intros Hin Hans Hev. + assert (length qs ≤ n ∨ length qs > n) as [H|H] by lia; first done. + destruct (@boring_decomposition _ qs n) as (q&other&qs_&Hqs_&Hq). lia. + assert (qs_ `prefix_of` qs) as Hqs. + { exists (q :: other). list_simplifier. done. } + replace n with (length qs_ + 0) in Hev by lia. + destruct (sub_tree_3 Hin Hqs) as [ans_ [Ha1 [Ha2 [q' Hq']]]]; first lia. + eapply interrogation_plus_evalt_comp_1 in Hev; eauto. + cbn in Hev. destruct (seval (τ (ans_ ++ [])))eqn: E; last done. + destruct s. congruence. + enough (Output o = Ask q') by congruence. + { eapply hasvalue_det. eapply seval_hasvalue. exists m. done. + list_simplifier. apply Hq'. } + Qed. + Lemma final_fact_gen (τ: tree) m f g use ans: interrogation τ (λ x y, f x = y) use ans → interrogation τ (λ x y, g x = y) use ans → @@ -426,6 +593,53 @@ Section Step_Eval. by rewrite elem_of_list_singleton. Qed. + + + Fact there_is_a_computation (τ: tree) n m f use ans v: + interrogation τ (λ x y, f x = y) use ans → + τ ans =! Output v → + evalt_comp τ f n m = Some (Output v) → + evalt_comp τ f (length use) m = Some (Output v). + Proof. + intros H1 H2 H3. + specialize (evalt_comp_step_length H1 H2 H3) as Hn. + assert (∃ k, length use + k = n) as [k Hk] by (exists (n - length use); lia). + clear Hn. rewrite <- Hk in H3. + specialize (interrogation_plus_evalt_comp_1 H1 H3) as H. + assert (seval ( (λ l', τ (ans ++ l')) []) m = Some (Output v)). + { destruct k. cbn in H. destruct (seval (τ (ans ++ [])) m) eqn: E; last done. + destruct s; congruence. cbn in H. + destruct (seval (τ (ans ++ [])) m) eqn: E; last done. destruct s; eauto. + list_simplifier . + assert (τ ans =! Ask q). by eapply seval_hasvalue'. + specialize (hasvalue_det H2 H0). congruence. } + specialize (evalt_comp_tail f H0) as H'. + assert (length use = length use + 0) as -> by lia. + rewrite <- interrogation_plus_evalt_comp; eauto. + eapply (sub_computation_ter _ H3); eauto. + Unshelve. rewrite (interrogation_length H1). lia. + Qed. + + Lemma final_fact' (τ: tree) m f g use ans v: + evalt_comp τ f (length use) m = Some (Output v) → + interrogation τ (λ x y, f x = y) use ans → + interrogation τ (λ x y, g x = y) use ans → + τ ans =! Output v → + evalt_comp τ g (length use) m = Some (Output v). + Proof. + intros Hf H1 H2 Ho. + eapply (final_fact_gen H1 H2 Hf); eauto. + intros u Hu. clear Ho Hf. + induction H1; first by apply not_elem_of_nil in Hu. + inv H2; first by apply app_cons_not_nil in H4. + apply app_inj_tail in H3. + apply app_inj_tail in H4. + destruct H3, H4; subst. + apply elem_of_app in Hu. + destruct Hu as [Hu1| Hu2%elem_of_list_singleton]; last by rewrite Hu2, H4. + apply IHinterrogation; eauto. + Qed. + Lemma final_fact (τ: tree) n m f g use ans v: evalt_comp τ f n m = Some (Output v) → interrogation τ (λ x y, f x = y) use ans → @@ -435,13 +649,57 @@ Section Step_Eval. Proof. intros Hf H1 H2 Ho. specialize (evalt_comp_step_length H1 Ho Hf) as Hn. - eapply (final_fact_gen H1 H2 Hf). - Admitted. + apply (there_is_a_computation H1) in Hf; last done. + eapply evalt_comp_step_mono''. + eapply final_fact'; eauto. done. + Qed. End Step_Eval. Section Use_Function. +Lemma extract_computation {Q O: Type} (τ: (list bool) ↛ (Q + O)) (f: Q → bool) n m v: + eq_dec O → eq_dec Q → + evalt_comp τ f n m = Some (inr v) → + Σ use ans, + τ ans =! inr v ∧ + ∀ P, (∀ q, q ∈ use → P q ↔ f q = true) → + interrogation τ (char_rel P) use ans . +Proof. + intros He H_ H_'. + destruct (evalt_comp_to_interrogation τ f v n m) as [(qs&ans&H3&_&H1&H2)|H]; try done. + exists qs, ans. + split; first done. + intros p Hp. + eapply interrogation_ext; last exact H1; eauto. + intros q' [|] Hqa'%elem_of_list_In; cbn; first by rewrite Hp. + specialize (Hp q' Hqa'). split. + - intros H. destruct (f q') eqn: E; last done. + enough (p q') by easy. by rewrite Hp. + - intros H H'%Hp. congruence. +Qed. + +Lemma use_function' {Q O: Type} (τ: (list bool) ↛ (Q + O)) (f: Q → bool) n m v: + eq_dec O → eq_dec Q → + (evalt_comp τ f n m = Some (inr v)) + (evalt_comp τ f n m ≠ Some (inr v)). +Proof. + intros h1 h2. + destruct (evalt_comp_to_interrogation τ f v n m) as [(qs&ans&H3&_&H1&H2)|H]; intuition done. +Qed. + +(* Lemma evalt_comp_ext_strong (τ τ': tree) f n m: +(∀ l n, seval (τ l) n = seval (τ' l) n) → +(∀ (q_ : Q) (a0 : A), In q_ q → f q_ a0 ↔ f' q_ a0) → +evalt_comp τ f n m = evalt_comp τ' f n m. +Proof. +intro Heq; induction n in τ, τ', Heq |- *; cbn. +- by rewrite <- Heq. +- rewrite <- Heq. + destruct (seval (A:=Q + O) (τ []) m); eauto. + destruct s; eauto. +Qed. *) + + Lemma use_function {Q O: Type} (τ: (list bool) ↛ (Q + O)) (f: Q → bool) n m v: eq_dec O → eq_dec Q → (Σ use, @@ -599,6 +857,11 @@ Section Step_Eval_Spec. is inl H then S (list_max (projT1 H)) else 0. + Definition φ' (f: nat → bool) (e x n: nat) := + if (use_function' (ξ () e x) f n n () unit_eq_dec nat_eq_dec) + is inl H then S (list_max (projT1 (extract_computation unit_eq_dec nat_eq_dec H))) + else 0. + Variable P: nat → Prop. Variable decider: nat → nat → bool. Hypothesis S_P: stable_semi_decider P decider. @@ -613,7 +876,7 @@ Section Step_Eval_Spec. Qed. Theorem Φ_spec e x: - Ξ e (char_rel P) x → (∞∀ n, Φ_ decider e x n = Some ()). + Ξ e (char_rel P) x → (∞∀ n, Φ_ decider e x n = Some ()). Proof. unfold Ξ, rel. intros (qs & ans & Ha & Hb). specialize (@S_approx_Σ1 _ _ _ S_P () (ξ _ e x) qs ans Ha) as H. @@ -624,6 +887,7 @@ Section Step_Eval_Spec. destruct s. by injection Hw. by destruct u. congruence. Qed. + Notation "A ≡{ k }≡ B" := (∀ x, x ≤ k → A x ↔ B x) (at level 30). Definition to_pred (f: nat → bool) x := f x = true. @@ -642,6 +906,23 @@ Section Step_Eval_Spec. eapply Hq. Qed. + Theorem φ'_spec e x n p: + Φ_ decider e x n = Some () → + p ≡{φ' (decider n) e x n}≡ to_pred (decider n) → + Ξ e (char_rel p) x. + Proof. + intros H2 H1. rewrite phi_iff_evalt in H2. unfold φ' in H1. + destruct (use_function' (ξ () e x) (decider n) n n () _ _) as [H'|H]; last done. + destruct (extract_computation unit_eq_dec nat_eq_dec H') as [qs [ans [Hans1 Hans2]]]. + simpl in H1. + exists qs, ans; split; eauto. + eapply Hans2. + intros q [i Hq]%elem_of_list_lookup_1. rewrite H1; first done. + simpl. enough (q ≤ list_max (qs)) by lia. + eapply implementation.list_max_lookup. + eapply Hq. + Qed. + Lemma φ_spec0_1 e x n: φ (decider n) e x n ≠ 0 → Φ_ decider e x n = Some (). Proof. @@ -681,6 +962,7 @@ Section Step_Eval_Spec. ∀ q a, char_rel (decider n) q a ↔ (λ x y, decider n x = y) q a. Proof. intros. unfold char_rel. destruct a, (decider n q); intuition. Qed. + Theorem φ_spec1 e x n k : φ (decider n) e x n = S k → to_pred (decider n) ≡{k}≡ to_pred (decider (S n)) → @@ -726,6 +1008,49 @@ Section Step_Eval_Spec. eapply final_fact; eauto. Qed. + Theorem φ_spec2 e x n k : + φ (decider n) e x n = S k → + to_pred (decider n) ≡{k}≡ to_pred (decider (S n)) → + Φ_ decider e x (S n) = Some (). + Proof. + intros H H2. unfold φ in *. + destruct (use_function (ξ () e x) (decider n) n n) + as [(use & Hu1 & Hu2)|]; last congruence. simpl in *. + + (* eapply interrogation_ext; last apply Hans; first done. + intros. by rewrite char_rel_boring. } *) + destruct (evalt_comp (ξ () e x) (decider n) n n) eqn: E1; last congruence. + destruct s eqn: E2; first congruence. + destruct u; clear Hu1 E2. + unfold Φ_; enough (evalt_comp (ξ () e x) (decider (S n)) (S n) (S n) = Some (inr ())) as H' + by by rewrite H'. + eapply evalt_comp_step_mono'. + eapply evalt_comp_depth_mono; last exact (le_S _ _ (le_n n)). + + + assert (∀ q, q ∈ use → decider (S n) q ↔ decider n q = true) as boring1. + { intros q Hq. destruct (H2 q). apply elem_of_list_lookup_1 in Hq. + destruct Hq as [i Hi]. injection H; intros <-. + by eapply implementation.list_max_lookup. + unfold to_pred in *. destruct (decider (S n) q); intuition. } + + assert (∀ q a, In q use → char_rel (decider n) q a ↔ char_rel (decider (S n)) q a) as boring2. + { intros q a Hq. destruct (H2 q). rewrite <-elem_of_list_In in Hq. + apply elem_of_list_lookup_1 in Hq. destruct Hq as [i Hi]. + injection H; intros <-. by eapply implementation.list_max_lookup. + unfold to_pred, char_rel in *. + destruct a, (decider (S n) q), (decider n q); intuition. } + + destruct (Hu2 (decider (S n)) boring1) as [ans (Hans & Hans1)]. + assert (interrogation (ξ () e x) (λ q a, decider n q = a) use ans) as Hansn. + { eapply interrogation_ext; last apply Hans; first done. intros. + rewrite <- char_rel_boring. by apply boring2. } + assert (n ≤ S n) as _H_ by lia. + assert (interrogation (ξ () e x) (λ q a, decider (S n) q = a) use ans) as HansSn. + { eapply interrogation_ext; last apply Hans; first done. intros; by rewrite char_rel_boring. } + eapply final_fact; eauto. + Qed. + End Step_Eval_Spec. From 49450621b8c9000f90b45e81d642eabbba20e783 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 12 Aug 2024 17:16:13 +0200 Subject: [PATCH 28/54] error --- theories/ReducibilityDegrees/low_wall.v | 2 -- theories/ReducibilityDegrees/lowsimple.v | 32 ++++++++++++++----- .../ReducibilityDegrees/simple_extension.v | 26 +++++++-------- theories/TuringReducibility/StepIndex.v | 5 --- 4 files changed, 36 insertions(+), 29 deletions(-) diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 9467fa7..17bf7cd 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -341,8 +341,6 @@ Section Wall. destruct (Dec (P_Ω e m ↓)); eauto. Qed. - (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both - eventally_wall and wall_convergence *) Fact P_simple_test: simple P. Proof. eapply P_simple; first apply EA. intro e. intros H. apply H. apply wall_convergence_test. Qed. diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index ad4d4f0..6316cbd 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -4,6 +4,7 @@ Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Export SyntheticComputability.ReducibilityDegrees.low_wall. +Require Export SyntheticComputability.ReducibilityDegrees.simple_extension. Require Import Arith. Require Import SyntheticComputability.PostsTheorem.PostsTheorem. Require Import Vectors.VectorDef Arith.Compare_dec Lia. @@ -110,27 +111,42 @@ Section LowFacts. eexists. eapply low_simple_correct; split. - eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit; eauto. - - eapply P_simple; eauto. + - eapply low_wall.P_simple; eauto. Qed. End LowSimplePredicate. Section LowSimplePredicate2. - Variable η: nat → nat → option nat. - Hypothesis EA: - forall P, semi_decidable P → exists e, forall x, P x ↔ exists n, η e n = Some x. + Lemma EA_correctness: sigT (λ φ, EA φ). + Proof. + Import SyntheticComputability.Axioms.EA.Assume_EA. + exists φ. intros P HP%SyntheticComputability.Axioms.EA.enum_iff. + rewrite W_spec in HP. destruct HP as [c Hc]. + exists c. intros x. unfold W in Hc. + apply Hc. + Qed. - Hypothesis LEM_Σ_1: LEM_Σ 1. + Definition θ := projT1 EA_correctness. + Lemma EA: + ∀ p, semi_decidable p → ∃ e, ∀ x, p x ↔ ∃ n, θ e n = Some x. + Proof. intros; unfold θ; destruct (EA_correctness); eauto. Qed. - Theorem a_sol_Post's_problem_2: exists P, sol_Post's_problem P. + Theorem a_sol_Post's_problem_2 (H: LEM_Σ 1): ∃ P, sol_Post's_problem P. Proof. eexists. eapply low_simple_correct; split. - eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit_2; eauto. - - eapply P_simple; eauto. + - eapply low_wall.P_simple; eauto. exact Qed. + Corollary a_fact `(LEM_Σ 1): + ∃ p: nat → Prop, ¬ decidable p ∧ enumerable p ∧ ¬ K ⪯ᴛ p. + Proof. + by apply a_sol_Post's_problem_2. + Qed. End LowSimplePredicate2. -End LowFacts. \ No newline at end of file +End LowFacts. + +Print Assumptions a_fact. \ No newline at end of file diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index 8103d78..f7ef1b8 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -5,6 +5,7 @@ Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. Require Import SyntheticComputability.ReducibilityDegrees.priority_method. +Import SyntheticComputability.Axioms.EA.Assume_EA. Import ListNotations. Section ComplToBound. @@ -68,10 +69,17 @@ End ComplToBound. Section Assume_EA. - Variable φ: nat -> nat -> option nat. - Definition EA := forall P, - semi_decidable P -> exists e, forall x, P x <-> exists n, φ e n = Some x. - Hypothesis EA: EA. + Definition θ := φ. + + Definition EA_spec := ∀ p, semi_decidable p -> ∃ e, ∀ x, p x <-> ∃ n, φ e n = Some x. + + Lemma EA: EA_spec. + Proof. + intros P HP%SyntheticComputability.Axioms.EA.enum_iff. + rewrite W_spec in HP. destruct HP as [c Hc]. + exists c. intros x. unfold W in Hc. + eapply Hc. + Qed. Definition W_ n e x := φ n e = Some x. Definition W e x := exists n, W_ e n x. @@ -766,13 +774,3 @@ Section Assume_EA. End Assume_EA. -Require SyntheticComputability.Axioms.EA. - -Lemma EA_correctness: Σ φ, EA φ. -Proof. - Import SyntheticComputability.Axioms.EA.Assume_EA. - exists φ. intros P HP%SyntheticComputability.Axioms.EA.enum_iff. - rewrite W_spec in HP. destruct HP as [c Hc]. - exists c. intros x. unfold W in Hc. - apply Hc. -Qed. \ No newline at end of file diff --git a/theories/TuringReducibility/StepIndex.v b/theories/TuringReducibility/StepIndex.v index 9c91f0e..2a8357e 100644 --- a/theories/TuringReducibility/StepIndex.v +++ b/theories/TuringReducibility/StepIndex.v @@ -593,8 +593,6 @@ Section Step_Eval. by rewrite elem_of_list_singleton. Qed. - - Fact there_is_a_computation (τ: tree) n m f use ans v: interrogation τ (λ x y, f x = y) use ans → τ ans =! Output v → @@ -1017,8 +1015,6 @@ Section Step_Eval_Spec. destruct (use_function (ξ () e x) (decider n) n n) as [(use & Hu1 & Hu2)|]; last congruence. simpl in *. - (* eapply interrogation_ext; last apply Hans; first done. - intros. by rewrite char_rel_boring. } *) destruct (evalt_comp (ξ () e x) (decider n) n n) eqn: E1; last congruence. destruct s eqn: E2; first congruence. destruct u; clear Hu1 E2. @@ -1027,7 +1023,6 @@ Section Step_Eval_Spec. eapply evalt_comp_step_mono'. eapply evalt_comp_depth_mono; last exact (le_S _ _ (le_n n)). - assert (∀ q, q ∈ use → decider (S n) q ↔ decider n q = true) as boring1. { intros q Hq. destruct (H2 q). apply elem_of_list_lookup_1 in Hq. destruct Hq as [i Hi]. injection H; intros <-. From 54a37e7d0e482ff0b08adf3dc8f2724d83b2f75d Mon Sep 17 00:00:00 2001 From: Haoyi Date: Mon, 12 Aug 2024 17:40:44 +0200 Subject: [PATCH 29/54] finish proof --- theories/ReducibilityDegrees/low_wall.v | 51 +++++++++---------- theories/ReducibilityDegrees/lowsimple.v | 20 ++------ .../ReducibilityDegrees/simple_extension.v | 3 +- 3 files changed, 29 insertions(+), 45 deletions(-) diff --git a/theories/ReducibilityDegrees/low_wall.v b/theories/ReducibilityDegrees/low_wall.v index 17bf7cd..54885a5 100644 --- a/theories/ReducibilityDegrees/low_wall.v +++ b/theories/ReducibilityDegrees/low_wall.v @@ -194,18 +194,15 @@ End Requirements_Lowness_Correctness. Section Wall. - Variable η: nat → nat → option nat. - Hypothesis EA: EA η. - Instance wall_instance: Wall := λ e L n, φ (λ x, Dec (In x L)) e e n. - Definition P := P η wall. - Definition P_func := P_func η wall. - Instance E: Extension := simple_extension η wall. + Definition P := P wall. + Definition P_func := P_func wall. + Instance E: Extension := simple_extension wall. Fact P_semi_decidable: semi_decidable P. Proof. eapply P_semi_decidable. Qed. - Definition χ := χ (simple_extension η wall). + Definition χ := χ (simple_extension wall). Definition P_Φ := (Φ_ χ). Definition P_Ω := (Ω χ). @@ -214,7 +211,7 @@ Section Wall. Hypothesis Σ_1_lem: LEM_Σ 1. Lemma attend_at_most_once_bound_constructive: - ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ attend η wall e s'. + ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ attend wall e s'. Proof. by apply attend_at_most_once_bound_test. Qed. Lemma eventally_wall_test: @@ -224,7 +221,7 @@ Section Wall. destruct (@attend_at_most_once_bound_constructive (S e)) as [s Hs]. exists (S s). intros m Hm x [e_ [He_ He_']]. destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (attend η wall e_ m). + { exfalso. enough (attend wall e_ m). unshelve eapply (Hs e_ _ m); eauto; lia. split; first lia. destruct He_' as [H _]. apply H. } @@ -233,7 +230,7 @@ Section Wall. by eapply H5. Qed. - Fact wall_convergence_test e : ∃ b : nat, lim_to η wall (wall e) b. + Fact wall_convergence_test e : ∃ b : nat, lim_to wall (wall e) b. Proof. destruct (@eventally_wall_test e) as [N HN]. assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. @@ -250,9 +247,9 @@ Section Wall. reflexivity. intros; split. + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension η wall) m) (F_func (simple_extension η wall) (S m))). + enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension η wall) (S m)) as HE. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. inv HE. * assert (wall e (P_func m) m < a). { eapply HN. lia. enough (P_func m = l) as ->. apply H5. @@ -261,10 +258,10 @@ Section Wall. rewrite H6 in H2. destruct (Dec (a = x0)). lia. apply Dec_auto. rewrite <- H3 in K. destruct K; first done. - enough ((F_func (simple_extension η wall) m) = l) as -> by done. + enough ((F_func (simple_extension wall) m) = l) as -> by done. eapply F_uni; last apply H4; apply F_func_correctness. * apply Dec_auto. - enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. + enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. eapply F_uni; first apply F_func_correctness. assumption. + eapply assume_Σ_1_dne. apply Σ_1_lem. @@ -319,9 +316,9 @@ Section Wall. eapply (@φ_spec2 χ _ ); eauto. intros x Hx; split. + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension η wall) m) (F_func (simple_extension η wall) (S m))). + enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension η wall) (S m)) as HE. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. inv HE. * assert (wall e (P_func m) m < a). { eapply HN. lia. enough (P_func m = l) as ->. apply H2. @@ -331,10 +328,10 @@ Section Wall. rewrite H3 in H. destruct (Dec (a = x)). lia. apply Dec_auto. rewrite <- H0 in K. destruct K; first done. - enough ((F_func (simple_extension η wall) m) = l) as -> by done. + enough ((F_func (simple_extension wall) m) = l) as -> by done. eapply F_uni; last apply H1; apply F_func_correctness. * apply Dec_auto. - enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. + enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. eapply F_uni; first apply F_func_correctness. assumption. - right. exists N. intros m Hm. @@ -342,7 +339,7 @@ Section Wall. Qed. Fact P_simple_test: simple P. - Proof. eapply P_simple; first apply EA. intro e. + Proof. eapply P_simple. intro e. intros H. apply H. apply wall_convergence_test. Qed. Hypothesis Σ_2_LEM: @@ -360,11 +357,11 @@ Section Wall. Lemma eventally_wall e: ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). Proof. - intros H_. eapply (@attend_at_most_once_bound η wall (S e)). + intros H_. eapply (@attend_at_most_once_bound wall (S e)). intros [s Hs]. eapply H_; clear H_. exists (S s). intros m Hm x [e_ [He_ He_']]. destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (attend η wall e_ m). + { exfalso. enough (attend wall e_ m). unshelve eapply (Hs e_ _ m); eauto; lia. split; first lia. destruct He_' as [H _]. apply H. } @@ -373,7 +370,7 @@ Section Wall. by eapply H5. Qed. - Fact wall_convergence e : ¬¬ ∃ b : nat, lim_to η wall (wall e) b. + Fact wall_convergence e : ¬¬ ∃ b : nat, lim_to wall (wall e) b. Proof. intros H_. destruct (@eventally_wall e). intros [N HN]. @@ -390,9 +387,9 @@ Section Wall. reflexivity. intros; split. + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension η wall) m) (F_func (simple_extension η wall) (S m))). + enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension η wall) (S m)) as HE. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. inv HE. * assert (wall e (P_func m) m < a). { eapply HN. lia. enough (P_func m = l) as ->. apply H5. @@ -401,10 +398,10 @@ Section Wall. rewrite H6 in H2. destruct (Dec (a = x0)). lia. apply Dec_auto. rewrite <- H3 in K. destruct K; first done. - enough ((F_func (simple_extension η wall) m) = l) as -> by done. + enough ((F_func (simple_extension wall) m) = l) as -> by done. eapply F_uni; last apply H4; apply F_func_correctness. * apply Dec_auto. - enough (F_func (simple_extension η wall) m = F_func (simple_extension η wall) (S m)) as -> by eauto. + enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. eapply F_uni; first apply F_func_correctness. assumption. + intro H. @@ -449,7 +446,7 @@ Section Wall. (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both eventally_wall and wall_convergence *) Fact P_simple: simple P. - Proof. eapply P_simple; first apply EA. intro e. apply wall_convergence. Qed. + Proof. eapply P_simple. intro e. apply wall_convergence. Qed. Section with_LEM_2. diff --git a/theories/ReducibilityDegrees/lowsimple.v b/theories/ReducibilityDegrees/lowsimple.v index 6316cbd..4502338 100644 --- a/theories/ReducibilityDegrees/lowsimple.v +++ b/theories/ReducibilityDegrees/lowsimple.v @@ -118,26 +118,12 @@ Section LowFacts. Section LowSimplePredicate2. - Lemma EA_correctness: sigT (λ φ, EA φ). - Proof. - Import SyntheticComputability.Axioms.EA.Assume_EA. - exists φ. intros P HP%SyntheticComputability.Axioms.EA.enum_iff. - rewrite W_spec in HP. destruct HP as [c Hc]. - exists c. intros x. unfold W in Hc. - apply Hc. - Qed. - - Definition θ := projT1 EA_correctness. - Lemma EA: - ∀ p, semi_decidable p → ∃ e, ∀ x, p x ↔ ∃ n, θ e n = Some x. - Proof. intros; unfold θ; destruct (EA_correctness); eauto. Qed. - Theorem a_sol_Post's_problem_2 (H: LEM_Σ 1): ∃ P, sol_Post's_problem P. Proof. eexists. eapply low_simple_correct; split. - eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit_2; eauto. - - eapply low_wall.P_simple; eauto. exact + - eapply low_wall.P_simple; eauto. Qed. Corollary a_fact `(LEM_Σ 1): @@ -149,4 +135,6 @@ Section LowFacts. End LowFacts. -Print Assumptions a_fact. \ No newline at end of file +(* Check a_fact. *) +(* Print Assumptions a_fact. *) + diff --git a/theories/ReducibilityDegrees/simple_extension.v b/theories/ReducibilityDegrees/simple_extension.v index f7ef1b8..c2fb29a 100644 --- a/theories/ReducibilityDegrees/simple_extension.v +++ b/theories/ReducibilityDegrees/simple_extension.v @@ -70,8 +70,7 @@ End ComplToBound. Section Assume_EA. Definition θ := φ. - - Definition EA_spec := ∀ p, semi_decidable p -> ∃ e, ∀ x, p x <-> ∃ n, φ e n = Some x. + Definition EA_spec := ∀ p, semi_decidable p → ∃ e, ∀ x, p x ↔ ∃ n, φ e n = Some x. Lemma EA: EA_spec. Proof. From 9d1f69c1189b617d071c4b225e9c176db45ebfbf Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sat, 17 Aug 2024 12:30:51 +0200 Subject: [PATCH 30/54] PostsProblem --- theories/PostsProblem/.DS_Store | Bin 0 -> 6148 bytes theories/{Basic => PostsProblem}/Limit.v | 0 .../StepIndex.v | 0 .../low_wall.v | 0 .../lowsimple.v | 0 .../priority_method.v | 0 .../simple_extension.v | 0 .../ReducibilityDegrees/lowsimple_extension.v | 0 8 files changed, 0 insertions(+), 0 deletions(-) create mode 100644 theories/PostsProblem/.DS_Store rename theories/{Basic => PostsProblem}/Limit.v (100%) rename theories/{TuringReducibility => PostsProblem}/StepIndex.v (100%) rename theories/{ReducibilityDegrees => PostsProblem}/low_wall.v (100%) rename theories/{ReducibilityDegrees => PostsProblem}/lowsimple.v (100%) rename theories/{ReducibilityDegrees => PostsProblem}/priority_method.v (100%) rename theories/{ReducibilityDegrees => PostsProblem}/simple_extension.v (100%) delete mode 100644 theories/ReducibilityDegrees/lowsimple_extension.v diff --git a/theories/PostsProblem/.DS_Store b/theories/PostsProblem/.DS_Store new file mode 100644 index 0000000000000000000000000000000000000000..16e93b579f14f286bbe6926c4c1770f980ff1d3e GIT binary patch literal 6148 zcmeHK!Ab)$5Pi{t)n3$-2a!FB2f;sBq=E%+{eafmf`!#B+FHFW_#a;MBLqRe!tc>H zlUO5bwI@-TfyqlIndI$*BpU#3chubhmH|}hf|Vr}pP2NEmuxOX7Ku*h7-Ab2jVwRv zWTnv-Fa=D3-==`<-31(=j}v;dfAbD| Date: Sun, 18 Aug 2024 18:10:25 +0200 Subject: [PATCH 31/54] Delete theories/PostsProblem/.DS_Store --- theories/PostsProblem/.DS_Store | Bin 6148 -> 0 bytes 1 file changed, 0 insertions(+), 0 deletions(-) delete mode 100644 theories/PostsProblem/.DS_Store diff --git a/theories/PostsProblem/.DS_Store b/theories/PostsProblem/.DS_Store deleted file mode 100644 index 16e93b579f14f286bbe6926c4c1770f980ff1d3e..0000000000000000000000000000000000000000 GIT binary patch literal 0 HcmV?d00001 literal 6148 zcmeHK!Ab)$5Pi{t)n3$-2a!FB2f;sBq=E%+{eafmf`!#B+FHFW_#a;MBLqRe!tc>H zlUO5bwI@-TfyqlIndI$*BpU#3chubhmH|}hf|Vr}pP2NEmuxOX7Ku*h7-Ab2jVwRv zWTnv-Fa=D3-==`<-31(=j}v;dfAbD| Date: Sun, 18 Aug 2024 18:10:55 +0200 Subject: [PATCH 32/54] Rename Limit.v to limit_computability.v --- theories/PostsProblem/{Limit.v => limit_computability.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{Limit.v => limit_computability.v} (100%) diff --git a/theories/PostsProblem/Limit.v b/theories/PostsProblem/limit_computability.v similarity index 100% rename from theories/PostsProblem/Limit.v rename to theories/PostsProblem/limit_computability.v From ec490eb48d88e21f65f9e98fda01d98fdce8820e Mon Sep 17 00:00:00 2001 From: Zeng Haoyi <80993578+HaoyiZeng@users.noreply.github.com> Date: Sun, 18 Aug 2024 18:11:21 +0200 Subject: [PATCH 33/54] Rename StepIndex.v to step_indexing.v --- theories/PostsProblem/{StepIndex.v => step_indexing.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{StepIndex.v => step_indexing.v} (100%) diff --git a/theories/PostsProblem/StepIndex.v b/theories/PostsProblem/step_indexing.v similarity index 100% rename from theories/PostsProblem/StepIndex.v rename to theories/PostsProblem/step_indexing.v From 493e4b4f0b5eaf6d02556a413f5c818155f9030d Mon Sep 17 00:00:00 2001 From: Zeng Haoyi <80993578+HaoyiZeng@users.noreply.github.com> Date: Sun, 18 Aug 2024 18:12:07 +0200 Subject: [PATCH 34/54] Rename low_wall.v to the_low_wall_function.v --- theories/PostsProblem/{low_wall.v => the_low_wall_function.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{low_wall.v => the_low_wall_function.v} (100%) diff --git a/theories/PostsProblem/low_wall.v b/theories/PostsProblem/the_low_wall_function.v similarity index 100% rename from theories/PostsProblem/low_wall.v rename to theories/PostsProblem/the_low_wall_function.v From 9f3e6638c0fd849a7a6875636e3a30c9a350b4c1 Mon Sep 17 00:00:00 2001 From: Zeng Haoyi <80993578+HaoyiZeng@users.noreply.github.com> Date: Sun, 18 Aug 2024 18:12:26 +0200 Subject: [PATCH 35/54] Rename lowsimple.v to low_simple_predicates.v --- theories/PostsProblem/{lowsimple.v => low_simple_predicates.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{lowsimple.v => low_simple_predicates.v} (100%) diff --git a/theories/PostsProblem/lowsimple.v b/theories/PostsProblem/low_simple_predicates.v similarity index 100% rename from theories/PostsProblem/lowsimple.v rename to theories/PostsProblem/low_simple_predicates.v From aa1d3051954d7e3ef62f0d877630c1af96329792 Mon Sep 17 00:00:00 2001 From: Zeng Haoyi <80993578+HaoyiZeng@users.noreply.github.com> Date: Sun, 18 Aug 2024 18:12:47 +0200 Subject: [PATCH 36/54] Rename priority_method.v to the_priority_method.v --- .../PostsProblem/{priority_method.v => the_priority_method.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{priority_method.v => the_priority_method.v} (100%) diff --git a/theories/PostsProblem/priority_method.v b/theories/PostsProblem/the_priority_method.v similarity index 100% rename from theories/PostsProblem/priority_method.v rename to theories/PostsProblem/the_priority_method.v From 5fb992a35f81a050c4255647445b9a7c864653a5 Mon Sep 17 00:00:00 2001 From: Zeng Haoyi <80993578+HaoyiZeng@users.noreply.github.com> Date: Sun, 18 Aug 2024 18:13:37 +0200 Subject: [PATCH 37/54] Rename the_low_wall_function.v to lowness.v --- theories/PostsProblem/{the_low_wall_function.v => lowness.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{the_low_wall_function.v => lowness.v} (100%) diff --git a/theories/PostsProblem/the_low_wall_function.v b/theories/PostsProblem/lowness.v similarity index 100% rename from theories/PostsProblem/the_low_wall_function.v rename to theories/PostsProblem/lowness.v From e69fa226c48597bd58e7ffb52c2e174ed8ddf84a Mon Sep 17 00:00:00 2001 From: Zeng Haoyi <80993578+HaoyiZeng@users.noreply.github.com> Date: Sun, 18 Aug 2024 18:14:03 +0200 Subject: [PATCH 38/54] Rename simple_extension.v to simpleness.v --- theories/PostsProblem/{simple_extension.v => simpleness.v} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename theories/PostsProblem/{simple_extension.v => simpleness.v} (100%) diff --git a/theories/PostsProblem/simple_extension.v b/theories/PostsProblem/simpleness.v similarity index 100% rename from theories/PostsProblem/simple_extension.v rename to theories/PostsProblem/simpleness.v From fac238d85445624b23a884b63adc67acd1e05e4a Mon Sep 17 00:00:00 2001 From: Haoyi Date: Sun, 18 Aug 2024 18:21:05 +0200 Subject: [PATCH 39/54] index --- theories/PostsProblem/.DS_Store | Bin 6148 -> 0 bytes .../{Limit.v => limit_computability.v} | 7 +- .../{lowsimple.v => low_simple_predicates.v} | 10 +-- .../PostsProblem/{low_wall.v => lowness.v} | 62 ++++++++++-------- .../{simple_extension.v => simpleness.v} | 12 +++- .../{StepIndex.v => step_indexing.v} | 29 +++++--- ...riority_method.v => the_priority_method.v} | 6 ++ 7 files changed, 83 insertions(+), 43 deletions(-) delete mode 100644 theories/PostsProblem/.DS_Store rename theories/PostsProblem/{Limit.v => limit_computability.v} (99%) rename theories/PostsProblem/{lowsimple.v => low_simple_predicates.v} (95%) rename theories/PostsProblem/{low_wall.v => lowness.v} (92%) rename theories/PostsProblem/{simple_extension.v => simpleness.v} (98%) rename theories/PostsProblem/{StepIndex.v => step_indexing.v} (98%) rename theories/PostsProblem/{priority_method.v => the_priority_method.v} (98%) diff --git a/theories/PostsProblem/.DS_Store b/theories/PostsProblem/.DS_Store deleted file mode 100644 index 16e93b579f14f286bbe6926c4c1770f980ff1d3e..0000000000000000000000000000000000000000 GIT binary patch literal 0 HcmV?d00001 literal 6148 zcmeHK!Ab)$5Pi{t)n3$-2a!FB2f;sBq=E%+{eafmf`!#B+FHFW_#a;MBLqRe!tc>H zlUO5bwI@-TfyqlIndI$*BpU#3chubhmH|}hf|Vr}pP2NEmuxOX7Ku*h7-Ab2jVwRv zWTnv-Fa=D3-==`<-31(=j}v;dfAbD| Prop): P ⪯ₘ (fun x: vec nat 1 => P (hd x)) . Proof. exists (fun x => [x]). now intros x. Qed. + + (** ** The Limit Lemma 1 *) + Lemma limit_turing_red_K {k: nat} (P: nat -> Prop) : LEM_Σ 1 -> limit_computable P -> @@ -331,6 +334,8 @@ Section LimitLemma2. End Construction. + (** ** The Limit Lemma 2 *) + Theorem turing_red_K_lim (P: nat -> Prop) : P ⪯ᴛ K -> definite K -> diff --git a/theories/PostsProblem/lowsimple.v b/theories/PostsProblem/low_simple_predicates.v similarity index 95% rename from theories/PostsProblem/lowsimple.v rename to theories/PostsProblem/low_simple_predicates.v index 4502338..068537a 100644 --- a/theories/PostsProblem/lowsimple.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -3,8 +3,8 @@ Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. -Require Export SyntheticComputability.ReducibilityDegrees.low_wall. -Require Export SyntheticComputability.ReducibilityDegrees.simple_extension. +Require Export SyntheticComputability.PostsProblem.low_wall. +Require Export SyntheticComputability.PostsProblem.simple_extension. Require Import Arith. Require Import SyntheticComputability.PostsTheorem.PostsTheorem. Require Import Vectors.VectorDef Arith.Compare_dec Lia. @@ -13,10 +13,8 @@ Require Import stdpp.list. Local Notation vec := Vector.t. - - (* ########################################################################## *) -(** * Low Simple Set *) +(** * Low Simple Predicates *) (* ########################################################################## *) (** This file contains the definition of low simple set and proves the @@ -118,6 +116,8 @@ Section LowFacts. Section LowSimplePredicate2. + (** ** A solution to Post's Problem *) + Theorem a_sol_Post's_problem_2 (H: LEM_Σ 1): ∃ P, sol_Post's_problem P. Proof. eexists. eapply low_simple_correct; split. diff --git a/theories/PostsProblem/low_wall.v b/theories/PostsProblem/lowness.v similarity index 92% rename from theories/PostsProblem/low_wall.v rename to theories/PostsProblem/lowness.v index 54885a5..9f54297 100644 --- a/theories/PostsProblem/low_wall.v +++ b/theories/PostsProblem/lowness.v @@ -4,8 +4,13 @@ Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. -Require Import SyntheticComputability.ReducibilityDegrees.priority_method. -Require Import SyntheticComputability.ReducibilityDegrees.simple_extension. +Require Import SyntheticComputability.PostsProblem.priority_method. +Require Import SyntheticComputability.PostsProblem.simple_extension. + + +(* ########################################################################## *) +(** * The Low Wall Function *) +(* ########################################################################## *) Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. Notation "'∞∃' x .. y , p" := @@ -35,6 +40,8 @@ Section Requirements_Lowness_Correctness. Section classic_logic. + (** ** Requirements *) + Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. Definition limit_decider e n: bool := Dec (Ω e n ↓). @@ -86,7 +93,7 @@ Section Requirements_Lowness_Correctness. End with_LEM_2. Section with_LEM_1. - + Hypothesis convergent: ∀ e, (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). Lemma Jump_limit_1 : limit_computable (P´). @@ -194,6 +201,8 @@ End Requirements_Lowness_Correctness. Section Wall. + (** ** Construction *) + Instance wall_instance: Wall := λ e L n, φ (λ x, Dec (In x L)) e e n. Definition P := P wall. Definition P_func := P_func wall. @@ -206,7 +215,9 @@ Section Wall. Definition P_Φ := (Φ_ χ). Definition P_Ω := (Ω χ). - Section TEST. + Section Verification. + + (** ** Verification *) Hypothesis Σ_1_lem: LEM_Σ 1. @@ -214,7 +225,7 @@ Section Wall. ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ attend wall e s'. Proof. by apply attend_at_most_once_bound_test. Qed. - Lemma eventally_wall_test: + Lemma eventally_wall: ∀ e, (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). Proof. intros e. @@ -230,9 +241,9 @@ Section Wall. by eapply H5. Qed. - Fact wall_convergence_test e : ∃ b : nat, lim_to wall (wall e) b. + Fact wall_convergence e : ∃ b : nat, lim_to wall (wall e) b. Proof. - destruct (@eventally_wall_test e) as [N HN]. + destruct (@eventally_wall e) as [N HN]. assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. assert (pdec (∀ x, N ≤ x → wall e (P_func x) x = 0)) as [H'|H']. { apply assume_Π_1_lem. apply Σ_1_lem. intros x. eauto. } @@ -274,11 +285,11 @@ Section Wall. apply H. now exists x; split. Qed. - Lemma N_requirements_test: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. + Lemma N_requirements: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. Proof. intros e He. - destruct (@eventally_wall_test e) as [N HN]. - destruct (@wall_convergence_test e) as [B [b Hb]]. + destruct (@eventally_wall e) as [N HN]. + destruct (@wall_convergence e) as [B [b Hb]]. set (M := max N b). destruct (He M) as [k [Hk Hk']]. eapply (@φ_spec χ e e k); first apply Hk'. intros x Hx. unfold P, simple_extension.P. @@ -305,7 +316,7 @@ Section Wall. Lemma convergent e : (∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓). Proof. - destruct (@eventally_wall_test e) as [N HN]. + destruct (@eventally_wall e) as [N HN]. assert (pdec (∃ k, N ≤ k ∧ P_Ω e k ↓)) as [[k [Hk HNk]]|H'] by (apply assume_Σ_1_lem; eauto). - left. exists k. intros n Hm. induction Hm; first done. @@ -338,9 +349,10 @@ Section Wall. destruct (Dec (P_Ω e m ↓)); eauto. Qed. - Fact P_simple_test: simple P. + Fact P_simple: simple P. Proof. eapply P_simple. intro e. - intros H. apply H. apply wall_convergence_test. Qed. + intros H. apply H. apply wall_convergence. + Qed. Hypothesis Σ_2_LEM: ∀ (P: nat → nat → Prop), @@ -349,12 +361,12 @@ Section Wall. Fact jump_P_limit_test: limit_computable (P´). Proof. eapply Jump_limit; last done. apply F_with_χ. - intros e He. eapply N_requirements_test; eauto. + intros e He. eapply N_requirements; eauto. Qed. - End TEST. + End Verification. - Lemma eventally_wall e: + Lemma eventally_wall_db e: ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). Proof. intros H_. eapply (@attend_at_most_once_bound wall (S e)). @@ -368,12 +380,12 @@ Section Wall. assert (e ≤ e_) by lia; clear E. destruct He_', H1, H1, H1, H1, H3. by eapply H5. - Qed. + Qed. - Fact wall_convergence e : ¬¬ ∃ b : nat, lim_to wall (wall e) b. + Fact wall_convergence_db e : ¬¬ ∃ b : nat, lim_to wall (wall e) b. Proof. intros H_. - destruct (@eventally_wall e). intros [N HN]. + destruct (@eventally_wall_db e). intros [N HN]. assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. - apply H_; clear H_. exists 0. exists N. intros. by apply H'. @@ -413,11 +425,11 @@ Section Wall. apply H. now exists x; split. Qed. - Lemma N_requirements: ∀ e, (∞∃ n, P_Ω e n ↓) → ¬ ¬ Ξ e (char_rel P) e. + Lemma N_requirements_db: ∀ e, (∞∃ n, P_Ω e n ↓) → ¬ ¬ Ξ e (char_rel P) e. Proof. intros e He H_. - apply (@eventally_wall e). intros [N HN]. - apply (@wall_convergence e). intros [B [b Hb]]. + apply (@eventally_wall_db e). intros [N HN]. + apply (@wall_convergence_db e). intros [B [b Hb]]. apply H_; clear H_. set (M := max N b). destruct (He M) as [k [Hk Hk']]. eapply (@φ_spec χ e e k); first apply Hk'. @@ -445,8 +457,6 @@ Section Wall. (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both eventally_wall and wall_convergence *) - Fact P_simple: simple P. - Proof. eapply P_simple. intro e. apply wall_convergence. Qed. Section with_LEM_2. @@ -457,7 +467,7 @@ Section Wall. Fact jump_P_limit: limit_computable (P´). Proof. eapply Jump_limit; last done. apply F_with_χ. - intros e He. eapply DN, N_requirements; eauto. + intros e He. eapply DN, N_requirements_db; eauto. Qed. End with_LEM_2. @@ -469,7 +479,7 @@ Section Wall. Fact jump_P_limit_2: limit_computable (P´). Proof. eapply Jump_limit_1; first apply F_with_χ. - - intros e He. eapply N_requirements_test; eauto. + - intros e He. eapply N_requirements; eauto. - eapply convergent; eauto. Qed. diff --git a/theories/PostsProblem/simple_extension.v b/theories/PostsProblem/simpleness.v similarity index 98% rename from theories/PostsProblem/simple_extension.v rename to theories/PostsProblem/simpleness.v index c2fb29a..929902d 100644 --- a/theories/PostsProblem/simple_extension.v +++ b/theories/PostsProblem/simpleness.v @@ -4,10 +4,14 @@ Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. -Require Import SyntheticComputability.ReducibilityDegrees.priority_method. +Require Import SyntheticComputability.PostsProblem.priority_method. Import SyntheticComputability.Axioms.EA.Assume_EA. Import ListNotations. +(* ########################################################################## *) +(** * The Simple Extension *) +(* ########################################################################## *) + Section ComplToBound. Definition complToBound L b : list nat := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). @@ -149,6 +153,8 @@ Section Assume_EA. Variable wall_instance: Wall. Section Extension. + + (** ** Construction *) Definition ext_intersect_W L n e := L # W[n] e. Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. @@ -254,6 +260,8 @@ Section Assume_EA. Section Requirements. + (** ** Requirements *) + Definition R_simple P e := non_finite e -> ~ (W e #ₚ P). Definition attend e n := e < n /\ least (ext_pick (P_func n) n) e. Definition act e n := ~ (P_func n) # W[n] e. @@ -271,6 +279,8 @@ Section Assume_EA. Section Requirements_Facts. + (** ** Verification *) + Lemma ext_pick_witness L n e: ext_pick L n e -> exists x, least (ext_has_wit L n e) x. Proof. diff --git a/theories/PostsProblem/StepIndex.v b/theories/PostsProblem/step_indexing.v similarity index 98% rename from theories/PostsProblem/StepIndex.v rename to theories/PostsProblem/step_indexing.v index 2a8357e..dac92cc 100644 --- a/theories/PostsProblem/StepIndex.v +++ b/theories/PostsProblem/step_indexing.v @@ -8,6 +8,10 @@ Require Import Coq.Program.Equality. From stdpp Require Export list. Import PartialTactics. +(* ########################################################################## *) +(** * Step-Inedxed Oracle Machines *) +(* ########################################################################## *) + Notation "'Σ' x .. y , p" := (sigT (fun x => .. (sigT (fun y => p)) ..)) (at level 200, x binder, right associativity, @@ -844,21 +848,15 @@ End Limit_Interrogation. Section Step_Eval_Spec. + (** ** Step-Inedxed Execution *) + Definition Φ_ (f: nat → nat → bool) (e x n: nat): option () := match evalt_comp (ξ () e x) (f n) n n with | Some (inr ()) => Some () | _ => None end. - Definition φ (f: nat → bool) (e x n: nat) := - if use_function (ξ () e x) f n n () unit_eq_dec nat_eq_dec - is inl H then S (list_max (projT1 H)) - else 0. - Definition φ' (f: nat → bool) (e x n: nat) := - if (use_function' (ξ () e x) f n n () unit_eq_dec nat_eq_dec) - is inl H then S (list_max (projT1 (extract_computation unit_eq_dec nat_eq_dec H))) - else 0. Variable P: nat → Prop. Variable decider: nat → nat → bool. @@ -886,9 +884,20 @@ Section Step_Eval_Spec. by destruct u. congruence. Qed. + (** ** Use Functions *) + + Notation "A ≡{ k }≡ B" := (∀ x, x ≤ k → A x ↔ B x) (at level 30). + Definition to_pred (f: nat → bool) x := f x = true. + + Definition φ (f: nat → bool) (e x n: nat) := + if use_function (ξ () e x) f n n () unit_eq_dec nat_eq_dec + is inl H then S (list_max (projT1 H)) + else 0. - Notation "A ≡{ k }≡ B" := (∀ x, x ≤ k → A x ↔ B x) (at level 30). - Definition to_pred (f: nat → bool) x := f x = true. + Definition φ' (f: nat → bool) (e x n: nat) := + if (use_function' (ξ () e x) f n n () unit_eq_dec nat_eq_dec) + is inl H then S (list_max (projT1 (extract_computation unit_eq_dec nat_eq_dec H))) + else 0. Theorem φ_spec e x n p: Φ_ decider e x n = Some () → diff --git a/theories/PostsProblem/priority_method.v b/theories/PostsProblem/the_priority_method.v similarity index 98% rename from theories/PostsProblem/priority_method.v rename to theories/PostsProblem/the_priority_method.v index d67b4a5..66c4996 100644 --- a/theories/PostsProblem/priority_method.v +++ b/theories/PostsProblem/the_priority_method.v @@ -7,6 +7,10 @@ Require Export SyntheticComputability.Shared.ListAutomation. Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. Import ListNotations. +(* ########################################################################## *) +(** * The Priority Method *) +(* ########################################################################## *) + Notation "'Σ' x .. y , p" := (sigT (fun x => .. (sigT (fun y => p)) ..)) (at level 200, x binder, right associativity, @@ -29,6 +33,8 @@ Notation unique p := (forall x x', p x -> p x' -> x = x'). Section Construction. + (** ** Basic Properties *) + Variable E: Extension. Lemma F_uni n: unique (F_ E n). From 4bcc0b33267b30769cf9699f0efd72bc507aed54 Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Tue, 10 Dec 2024 14:02:19 +0100 Subject: [PATCH 40/54] refactor to prove Post's problem from ~~ LPO --- theories/PostsProblem/low_simple_predicates.v | 37 ++++---- theories/PostsProblem/lowness.v | 94 +++---------------- theories/PostsProblem/simpleness.v | 68 +------------- theories/PostsProblem/the_priority_method.v | 2 +- theories/PostsTheorem/PostsTheorem.v | 28 ++++++ 5 files changed, 60 insertions(+), 169 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index 4502338..455b176 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -1,10 +1,10 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions limit_computability simple. Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. -Require Export SyntheticComputability.ReducibilityDegrees.low_wall. -Require Export SyntheticComputability.ReducibilityDegrees.simple_extension. +From SyntheticComputability Require Export lowness. +From SyntheticComputability Require Export simpleness. Require Import Arith. Require Import SyntheticComputability.PostsTheorem.PostsTheorem. Require Import Vectors.VectorDef Arith.Compare_dec Lia. @@ -111,30 +111,29 @@ Section LowFacts. eexists. eapply low_simple_correct; split. - eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit; eauto. - - eapply low_wall.P_simple; eauto. + - eapply P_simple. + intros. intros d. apply d. + apply wall_convergence_test. assumption. Qed. End LowSimplePredicate. - Section LowSimplePredicate2. - - Theorem a_sol_Post's_problem_2 (H: LEM_Σ 1): ∃ P, sol_Post's_problem P. + Theorem PostProblem_from_neg_negLPO : + ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (~~ (LEM_Σ 1) -> ¬ K ⪯ᴛ p). Proof. - eexists. eapply low_simple_correct; split. - - eapply limit_turing_red_K; eauto. exact 42. + eexists. + repeat split. + - apply simple_undecidable. + eapply P_simple. apply wall_convergence. + - apply P_semi_decidable. + - intros L. intros G. apply L. clear L. intros L. revert G. + apply lowness. red. + eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit_2; eauto. - - eapply low_wall.P_simple; eauto. - Qed. - - Corollary a_fact `(LEM_Σ 1): - ∃ p: nat → Prop, ¬ decidable p ∧ enumerable p ∧ ¬ K ⪯ᴛ p. - Proof. - by apply a_sol_Post's_problem_2. Qed. - End LowSimplePredicate2. End LowFacts. -(* Check a_fact. *) -(* Print Assumptions a_fact. *) +Check PostProblem_from_neg_negLPO. +Print Assumptions PostProblem_from_neg_negLPO. diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index 54885a5..c5d125f 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -1,11 +1,11 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions StepIndex Limit simple. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions step_indexing limit_computability simple. Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. -Require Import SyntheticComputability.ReducibilityDegrees.priority_method. -Require Import SyntheticComputability.ReducibilityDegrees.simple_extension. +From SyntheticComputability Require Import the_priority_method. +From SyntheticComputability Require Import simpleness. Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. Notation "'∞∃' x .. y , p" := @@ -66,19 +66,19 @@ Section Requirements_Lowness_Correctness. intros [w Hw]%Φ_spec; exists w; intros??. apply Dec_auto. by eapply Hw. intros [N HN]. eapply N_requirements. - intros m. exists (S N + m); split; first lia. + intros m. exists (Datatypes.S N + m)%nat; split; first lia. eapply Dec_true. eapply HN. lia. - unfold J. split; intros H. destruct (limit_case x) as [[k Hk]|h2]. enough (Ξ x (char_rel P) x) by easy. - eapply N_requirements. intros m. exists (S k + m). + eapply N_requirements. intros m. exists (Datatypes.S k + m). split; first lia. eapply Hk. lia. destruct h2 as [w Hw]. exists w. intros. specialize (Hw n H0). unfold limit_decider. destruct (Dec _); eauto. destruct H as [w Hw]. intros [k Hneg]%Φ_spec. - set (N := S (max w k)). + set (N := Datatypes.S (max w k)). assert (Ω x N ↓). { eapply Hneg. lia. } enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. @@ -96,19 +96,19 @@ Section Requirements_Lowness_Correctness. intros [w Hw]%Φ_spec; exists w; intros??. apply Dec_auto. by eapply Hw. intros [N HN]. eapply N_requirements. - intros m. exists (S N + m); split; first lia. + intros m. exists (Datatypes.S N + m); split; first lia. eapply Dec_true. eapply HN. lia. - unfold J. split; intros H. destruct (convergent x) as [[k Hk]|h2]. enough (Ξ x (char_rel P) x) by easy. - eapply N_requirements. intros m. exists (S k + m). + eapply N_requirements. intros m. exists (Datatypes.S k + m). split; first lia. eapply Hk. lia. destruct h2 as [w Hw]. exists w. intros. specialize (Hw n H0). unfold limit_decider. destruct (Dec _); eauto. destruct H as [w Hw]. intros [k Hneg]%Φ_spec. - set (N := S (max w k)). + set (N := Datatypes.S (max w k)). assert (Ω x N ↓). { eapply Hneg. lia. } enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. @@ -116,78 +116,6 @@ Section Requirements_Lowness_Correctness. End with_LEM_1. End classic_logic. -(* - Section intuitionistic_logic. - Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → ¬¬ Ξ e (char_rel P) e. - - Lemma N_requirements': ∀ e, ¬¬ ((∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e). - Proof. firstorder. Qed. - - Fact dn_or (R Q: Prop): (¬¬ R ∨ ¬¬ Q) → ¬¬ (R ∨ Q). - Proof. firstorder. Qed. - Lemma not_not_limit_case e: ~ ~ ((∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓)). - Proof. - ccase (∃ n, ∀ m, n ≤ m → ¬ Ω e m ↓) as [H1|H1]. - apply dn_or. { right. eauto. } - ccase (∀ i, ∃ n, i ≤ n ∧ Ω e n ↓) as [H2|H2]. - intros H_. apply (@N_requirements' e). - intros N_requirements'. - apply H_. left. apply Φ_spec, N_requirements'. intros i. - { destruct (H2 i) as [w Hw]. exists w. apply Hw. } - exfalso. clear P S_P N_requirements. - apply H2. intros i. - Admitted. - - Definition not_not_limit_computable {X} (P: X -> Prop) := - exists f: X -> nat -> bool, - forall x, ~~ - ((P x <-> exists N, forall n, n >= N -> f x n = true) /\ - (~ P x <-> exists N, forall n, n >= N -> f x n = false)). - - Fact dn_and (R Q: Prop): (¬¬ R ∧ ¬¬ Q) → ¬¬ (R ∧ Q). - Proof. firstorder. Qed. - - Lemma not_not_Jump_limit : not_not_limit_computable (P´). - Proof. - exists limit_decider; intro x. - apply dn_and. - split; intros. - - unfold J. intros H_. - eapply (@N_requirements' x). intros N_requirements'. - apply H_. split. - intros [w Hw]%Φ_spec; exists w; intros??. - apply Dec_auto. by eapply Hw. - intros [N HN]. - eapply N_requirements'. - intros m. exists (S N + m); split; first lia. - eapply Dec_true. eapply HN. lia. - - unfold J. intros H_. - eapply (@N_requirements' x). intros N_requirements'. - eapply (@not_not_limit_case x). - intros [[k Hk]|h2]. - apply H_. split. - enough (Ξ x (char_rel P) x) by easy. - eapply N_requirements'. intros m. exists (S k + m). - split; first lia. eapply Hk. lia. - intro H. destruct H as [w Hw]. - intros [k' Hneg]%Φ_spec. - set (N := S (max w k')). - assert (Ω x N ↓). { eapply Hneg. lia. } - enough (¬ Ω x N ↓) by eauto. - eapply Dec_false. eapply Hw. lia. - apply H_. split. - destruct h2 as [w Hw]. exists w. - intros. specialize (Hw n H0). unfold limit_decider. - destruct (Dec _); eauto. - intro H. destruct H as [w Hw]. - intros [k Hneg]%Φ_spec. - set (N := S (max w k)). - assert (Ω x N ↓). { eapply Hneg. lia. } - enough (¬ Ω x N ↓) by eauto. - eapply Dec_false. eapply Hw. lia. - Qed. - -End intuitionistic_logic. *) End Requirements_Lowness_Correctness. @@ -281,7 +209,7 @@ Section Wall. destruct (@wall_convergence_test e) as [B [b Hb]]. set (M := max N b). destruct (He M) as [k [Hk Hk']]. eapply (@φ_spec χ e e k); first apply Hk'. - intros x Hx. unfold P, simple_extension.P. + intros x Hx. unfold P, simpleness.P. rewrite F_with_top. split. - intros (L & m & HL & HLs &HP). assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } @@ -421,7 +349,7 @@ Section Wall. apply H_; clear H_. set (M := max N b). destruct (He M) as [k [Hk Hk']]. eapply (@φ_spec χ e e k); first apply Hk'. - intros x Hx. unfold P, simple_extension.P. + intros x Hx. unfold P, simpleness.P. rewrite F_with_top. split. - intros (L & m & HL & HLs &HP). assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index c2fb29a..a1af487 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -1,10 +1,10 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions limit_computability simple. Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. Require Export SyntheticComputability.Shared.Pigeonhole. Require Export SyntheticComputability.Shared.ListAutomation. Require Import Arith Arith.Compare_dec Lia Coq.Program.Equality List. -Require Import SyntheticComputability.ReducibilityDegrees.priority_method. +From SyntheticComputability Require Import the_priority_method. Import SyntheticComputability.Axioms.EA.Assume_EA. Import ListNotations. @@ -706,70 +706,6 @@ Section Assume_EA. End Result. End Assume_WALL. -(* - Section Instance_of_Wall. - - Definition low_wall (e: nat) (l: list nat) (n: nat) := 42. - Lemma low_wall_spec: forall e, ~~ exists b, lim_to low_wall (low_wall e) b. - Proof. intro e. intros H_; apply H_. exists 42. exists 0; intuition. Qed. - - Definition Pw := P low_wall. - - Theorem P_simple_low_wall: simple Pw. - Proof. - eapply P_simple. apply low_wall_spec. - Qed. - - Definition effectively_simple P := - simple P /\ exists f, - forall e, (forall x, W e x -> (compl P) x) -> forall x, W e x -> x < (f e). - - Lemma attend_pick e k: attend low_wall e k -> exists x, x > 2*e /\ Pw x /\ W e x. - Proof. - intros [He H]. - (* edestruct (ext_pick_witness) as [x Hx]. - { destruct H. eapply e0. } - assert (P_ (S k) (Pf_ (S k))) by apply F_func_correctness. - inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. - exists e. assert (Pf_ k = l) as <-. - eapply F_uni. apply F_func_correctness. apply H3. - split; first easy. split; first easy. easy. - assert (x = a) as ->. eapply (@extend_uni simple_extendsion); cbn; eauto. - exists a. split. destruct H4, H0, H1, H4, H4. - assert (x=e) as HE. - { eapply least_unique; last eapply H. - enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } - lia. split. exists (Pf_ (S k)), (S k); split; eauto. now rewrite <- H2. - apply F_func_correctness. destruct H4, H0, H1, H4, H4, H4, H4. - assert (x=e) as HE. - { eapply least_unique; last eapply H. - enough (l=(Pf_ k)) as -> by easy. eapply F_uni; eauto. apply F_func_correctness. } - exists x0; congruence. - exfalso. eapply (H3 x); exists e; do 2 (split; eauto). - enough ((Pf_ k) = (Pf_ (S k))) as <- by easy. - assert (F_ simple_extendsion k (Pf_ k)) by apply F_func_correctness. - eapply F_uni; eauto. *) - Admitted. - - Theorem P_effectively_simple: effectively_simple Pw. - Proof. - split; first apply P_simple. apply low_wall_spec. - exists (fun e => 2 * e + 1). - intros e He x Hex. enough (~ 2 * e < x) by lia. - intros Hex'. - assert (W e #ₚ Pw). - { intros y Hy1 Hy2. now apply (He y). } - (* assert (forall k, (Pf_ low_wall k) # W[k] e). - { intros k y Hy1 [w [_ Hy2]]. eapply (H y). exists w; eauto. - exists (Pf_ k), k; split; eauto. apply F_func_correctness. } - enough (exists k, attend e k) as [k Hk]. - (* apply H1. intros [k Hk]. *) - eapply attend_pick in Hk. - destruct Hk as [y (Hx1&Hx2&Hx3)]. - admit. admit. *) - Admitted. - - End Instance_of_Wall. *) End Assume_EA. diff --git a/theories/PostsProblem/the_priority_method.v b/theories/PostsProblem/the_priority_method.v index d67b4a5..661ccb4 100644 --- a/theories/PostsProblem/the_priority_method.v +++ b/theories/PostsProblem/the_priority_method.v @@ -1,4 +1,4 @@ -From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions Limit simple. +From SyntheticComputability Require Import ArithmeticalHierarchySemantic reductions SemiDec TuringJump OracleComputability Definitions limit_computability simple. From stdpp Require Export list. Require Import SyntheticComputability.Synthetic.DecidabilityFacts. Require Export SyntheticComputability.Shared.FinitenessFacts. diff --git a/theories/PostsTheorem/PostsTheorem.v b/theories/PostsTheorem/PostsTheorem.v index f52f00f..87dd084 100644 --- a/theories/PostsTheorem/PostsTheorem.v +++ b/theories/PostsTheorem/PostsTheorem.v @@ -201,6 +201,13 @@ Section PostsTheorem. destruct seval as [ [ | [] ] | ]eqn:E; eauto. Qed. + Lemma Σ_semi_decidable_in_Π_forward {k} (p: (vec nat k) -> Prop) n (DN : LEM_Π n) : + isΣsem (S n) p -> exists (p': vec nat (S k) -> Prop), isΠsem n p' /\ oracle_semi_decidable p' p. + Proof. + apply Σ_semi_decidable_in_Π1. + assumption. + Qed. + Lemma Σ_semi_decidable_in_Π {k} (p: (vec nat k) -> Prop) n (DN : LEM_Σ n) : isΣsem (S n) p <-> exists (p': vec nat (S k) -> Prop), isΠsem n p' /\ oracle_semi_decidable p' p. Proof. @@ -211,6 +218,18 @@ Section PostsTheorem. Hint Resolve DNEimpl. + (* Lemma Σ_semi_decidable_in_Σ_forward {k} (p: (vec nat k) -> Prop) n (DN : LEM_Π n) : *) + (* isΣsem (S n) p -> exists (p': vec nat (S k) -> Prop), isΣsem n p' /\ oracle_semi_decidable p' p. *) + (* Proof. *) + (* intros H % Σ_semi_decidable_in_Π_forward; eauto. *) + (* destruct H as [p' [H S]]. eapply negΣinΠsem in H as H'. *) + (* 2: now eapply LEM_Σ_to_DNE_Σ. *) + (* eexists. split;[apply H'|]. *) + (* rewrite <- oracle_semi_decidable_complement_iff. eauto. *) + (* eapply DNEimpl; eauto. *) + (* now eapply LEM_Σ_to_DNE_Σ. *) + (* Qed. *) + Lemma Σ_semi_decidable_in_Σ {k} (p: (vec nat k) -> Prop) n (DN : LEM_Σ n) : isΣsem (S n) p <-> exists (p': vec nat (S k) -> Prop), isΣsem n p' /\ oracle_semi_decidable p' p. Proof. @@ -263,6 +282,15 @@ Section PostsTheorem. intros. apply red_m_impl_red_T. eapply jump_Σn_complete; eauto. Qed. + Lemma Σ_semi_decidable_jump_forward {k} (p: (vec nat k) -> Prop) n (DN : LEM_Σ n) : + isΣsem (S n) p -> oracle_semi_decidable (­{0}^(n)) p. + Proof. + intros [p' [Σp' Sp']]%Σ_semi_decidable_in_Σ. + eapply (Turing_transports_sdec Sp'). + eapply jump_Σn_complete_redT. + all: eauto. + Qed. + Lemma Σ_semi_decidable_jump {k} (p: (vec nat k) -> Prop) n (DN : LEM_Σ n) : isΣsem (S n) p <-> oracle_semi_decidable (­{0}^(n)) p. Proof. From 885692cefc52a7314f682758a98fc8b750df5cf6 Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Wed, 11 Dec 2024 10:45:56 +0100 Subject: [PATCH 41/54] =?UTF-8?q?=C2=AC=C2=AC(=C2=AC=C2=AC=CE=A3=E2=81=B0?= =?UTF-8?q?=E2=82=81)-LEM=20suffices?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- theories/PostsProblem/low_simple_predicates.v | 53 +++++++++++++++++-- 1 file changed, 49 insertions(+), 4 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index 455b176..c9c9869 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -118,22 +118,67 @@ Section LowFacts. End LowSimplePredicate. + Notation "(¬¬Σ⁰₁)-LEM" := ((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ~~ definite p)) (at level 0). + Theorem PostProblem_from_neg_negLPO : - ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (~~ (LEM_Σ 1) -> ¬ K ⪯ᴛ p). + ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM -> ¬ K ⪯ᴛ p). Proof. eexists. repeat split. - apply simple_undecidable. eapply P_simple. apply wall_convergence. - apply P_semi_decidable. - - intros L. intros G. apply L. clear L. intros L. revert G. - apply lowness. red. + - intros L. intros G. apply L. clear L. intros L. + assert (~~ definite K) as hK. + { + specialize (L 1 (fun v => K (Vector.hd v))). + intros h. apply L. + (* use Sigma01-ness of K *) + admit. + intros h1. apply h. + intros x. specialize (h1 (Vector.cons x Vector.nil)). exact h1. + } + apply hK. clear hK. intros hK. + assert (LEM_Σ 1). + { + intros n p Hs. + (* use many-one completeness of K *) + admit. + } + revert G. apply lowness. red. eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit_2; eauto. - Qed. + Admitted. End LowFacts. Check PostProblem_from_neg_negLPO. Print Assumptions PostProblem_from_neg_negLPO. +(* general proof that ¬¬(¬¬Σ⁰₁)-LEM <-> ¬¬(Σ⁰₁)-LEM *) +Section assume. + + Variable enumerable : (nat -> Prop) -> Prop. + + Variable K : nat -> Prop. + Variable eK : enumerable K. + Variable cK : forall p : nat -> Prop, enumerable p -> red_m p K. + + Goal ~~ (forall p : nat -> Prop, enumerable p -> forall x, p x \/ ~ p x) + <-> + ~~ (forall p : nat -> Prop, enumerable p -> ~~ forall x, p x \/ ~ p x). + Proof. + split. + - firstorder. + - intros nnLPO H. + apply nnLPO. clear nnLPO. intros nnLPO. + apply (nnLPO K eK). intros dK. + apply H. + intros p [f Hf] % cK x. + specialize (dK (f x)). + red in Hf. rewrite <- Hf in dK. + assumption. + Qed. + +End assume. + From 87a3fcf8416e2330b8f9ee367235aa1acc53dfc6 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Wed, 11 Dec 2024 12:34:25 +0100 Subject: [PATCH 42/54] fix proof --- theories/PostsProblem/low_simple_predicates.v | 33 +++++++++++++++---- 1 file changed, 27 insertions(+), 6 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index c9c9869..0df22be 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -23,6 +23,18 @@ Local Notation vec := Vector.t. essential property of low simple, i.e. Low simple as a solution to Post's Problem in Turing degree. **) +Section Facts. + Lemma m_red_complete {X Y} (P: X → Prop) (Q: Y → Prop): + semi_decidable P → Q ⪯ₘ P → semi_decidable Q. + Proof. intros [g H1] [f H2]; exists (fun x => g (f x)); firstorder. Qed. + + Lemma m_red_complete_definite {X Y} (P: X → Prop) (Q: Y → Prop): + definite P → Q ⪯ₘ P → definite Q. + Proof. intros H [f H2]; firstorder. Qed. +End Facts. + + + (* Definition of low *) Definition low (P: nat -> Prop) := P´ ⪯ᴛ K. @@ -120,6 +132,13 @@ Section LowFacts. Notation "(¬¬Σ⁰₁)-LEM" := ((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ~~ definite p)) (at level 0). + Lemma m_red_K_semi_decidable {n} (P: vec nat n → Prop): + semi_decidable P -> P ⪯ₘ K. + Proof. + intros H. unfold K. rewrite <- red_m_iff_semidec_jump_vec. + by apply semi_decidable_OracleSemiDecidable. eauto. + Qed. + Theorem PostProblem_from_neg_negLPO : ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM -> ¬ K ⪯ᴛ p). Proof. @@ -132,9 +151,10 @@ Section LowFacts. assert (~~ definite K) as hK. { specialize (L 1 (fun v => K (Vector.hd v))). - intros h. apply L. - (* use Sigma01-ness of K *) - admit. + intros h. apply L. + rewrite <- semi_dec_iff_Σ1. + eapply m_red_complete; first apply semi_dec_halting. + exists (fun v => Vector.hd v); done. intros h1. apply h. intros x. specialize (h1 (Vector.cons x Vector.nil)). exact h1. } @@ -142,13 +162,14 @@ Section LowFacts. assert (LEM_Σ 1). { intros n p Hs. - (* use many-one completeness of K *) - admit. + eapply m_red_complete_definite; first apply hK. + rewrite <- semi_dec_iff_Σ1 in Hs. + by eapply m_red_K_semi_decidable. } revert G. apply lowness. red. eapply limit_turing_red_K; eauto. exact 42. apply jump_P_limit_2; eauto. - Admitted. + Qed. End LowFacts. From 597fdbb2847baa0626cd0390a7cdea51e57c600a Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Wed, 11 Dec 2024 12:47:33 +0100 Subject: [PATCH 43/54] remove one double negation --- theories/PostsProblem/low_simple_predicates.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index 0df22be..54a9c64 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -176,7 +176,7 @@ End LowFacts. Check PostProblem_from_neg_negLPO. Print Assumptions PostProblem_from_neg_negLPO. -(* general proof that ¬¬(¬¬Σ⁰₁)-LEM <-> ¬¬(Σ⁰₁)-LEM *) +(* general proof that (¬¬Σ⁰₁)-LEM <-> ¬¬(Σ⁰₁)-LEM under many-one complete Σ⁰₁ predicate *) Section assume. Variable enumerable : (nat -> Prop) -> Prop. @@ -187,12 +187,11 @@ Section assume. Goal ~~ (forall p : nat -> Prop, enumerable p -> forall x, p x \/ ~ p x) <-> - ~~ (forall p : nat -> Prop, enumerable p -> ~~ forall x, p x \/ ~ p x). + (forall p : nat -> Prop, enumerable p -> ~~ forall x, p x \/ ~ p x). Proof. split. - firstorder. - intros nnLPO H. - apply nnLPO. clear nnLPO. intros nnLPO. apply (nnLPO K eK). intros dK. apply H. intros p [f Hf] % cK x. @@ -202,4 +201,3 @@ Section assume. Qed. End assume. - From e9817a0a65d273941d6db46a3262033060742bda Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Tue, 21 Jan 2025 11:02:25 +0100 Subject: [PATCH 44/54] generalise attend_at_most_once_bound --- theories/PostsProblem/simpleness.v | 121 +++++++++++++++++------------ 1 file changed, 73 insertions(+), 48 deletions(-) diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index a1af487..49efa7a 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -338,30 +338,37 @@ Section Assume_EA. apply (attend_always_act H1 Hm). Qed. - Lemma attend_at_most_once e: ~ ~ (exists s', forall s, s' < s -> ~ attend e s). + Lemma attend_at_most_once_gen e: + (pdec (exists k, attend e k)) -> + (exists s', forall s, s' < s -> ~ attend e s). Proof. - ccase (exists k, attend e k) as [[w Hw]|H]. - - intros H. eapply H. exists w. + intros [[w Hw]|H]. + - exists w. now eapply attend_always_not_attend. - - intros Hk. apply Hk. exists 0. + - exists 0. intros k' Hk' Ha. apply H. now exists k'. Qed. - Definition done e s := ∀ s', s < s' → ¬ attend e s'. - - Lemma attend_at_most_once_test: - LEM_Σ 1 → ∀ e, ∃ s, done e s. + Lemma attend_at_most_once e: ~ ~ (exists s', forall s, s' < s -> ~ attend e s). Proof. - intros Hlem e. - assert (pdec (exists k, attend e k)) as [[w Hw]|H]. - { eapply assume_Σ_1_lem. apply Hlem. eauto. } - - exists w. unfold done. by eapply attend_always_not_attend. - - exists 0. - intros k' Hk' Ha. - apply H. now exists k'. + intros H. + assert (~~ (pdec (exists k, attend e k))) as Hdec. + { unfold pdec. tauto. } + apply Hdec. clear Hdec. intros Hdec. + apply H, attend_at_most_once_gen. assumption. Qed. + Definition done e s := ∀ s', s < s' → ¬ attend e s'. + + (* Lemma attend_at_most_once_test: *) + (* LEM_Σ 1 → ∀ e, ∃ s, done e s. *) + (* Proof. *) + (* intros Hlem e. *) + (* apply attend_at_most_once_gen. *) + (* { eapply assume_Σ_1_lem. apply Hlem. eauto. } *) + (* Qed. *) + Lemma attend_uni e: unique (attend e). Proof. intros k1 k2 H1 H2. @@ -578,29 +585,17 @@ Section Assume_EA. specialize (IHe_ i x H1). lia. Qed. - Lemma attend_at_most_once_bound k: - ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). - Proof. - induction k. - - intros H. apply H. exists 42. intros ??. lia. - - intros H. apply IHk. intros [s Hs]; clear IHk. - specialize (@attend_at_most_once k) as Hk. - apply Hk. intros [sk Hsk]; clear Hk. - set (max sk s) as N. - eapply H. exists N. intros e He. - assert (e = k \/ e < k) as [->|Hek] by lia. - intros s' Hs'. eapply Hsk. lia. - intros s' Hs'. eapply Hs; lia. - Qed. - - - Lemma attend_at_most_once_bound_test: - LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k -> ∀ s', s < s' -> ~ attend e s'). + Lemma attend_at_most_once_bound_gen k: + (forall k', k' < k -> pdec (∃ k0 : nat, attend k' k0)) -> + exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). Proof. - intros Hlem. induction k. + intros Hle. + induction k. - exists 42. intros ??. lia. - destruct IHk as [s Hs]. - specialize (@attend_at_most_once_test Hlem k) as [sk Hsk]. + { intros. eapply Hle. lia. } + destruct (@attend_at_most_once_gen k) as [sk Hsk]. + { apply Hle. lia. } set (max sk s) as N. exists N. intros e He. assert (e = k \/ e < k) as [->|Hek] by lia. @@ -608,19 +603,49 @@ Section Assume_EA. intros s' Hs'. eapply Hs; lia. Qed. - Lemma non_finite_not_bounded e: - non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ - (forall n, forall i, i <= e -> wall i (P_func n) n < x). - Proof. - intro H. unfold non_finite in H. - intros He. rewrite non_finite_nat in H. - apply (@wall_of_wall e). intros [w Hw]. - pose (N := max (2*e + 1) w). specialize (H N). - apply H. intros [m [Hm1 [k Hmk]]]. - apply He. exists k, m. - repeat split. now exists k. - lia. intros n i Hie. specialize (Hw i n Hie). lia. - Qed. + Lemma finite_decs (P : nat -> Prop) k Q : + ((forall k', k' < k -> pdec (P k')) -> ~ Q) -> ~ Q. + Proof. + induction k. + - firstorder lia. + - intros H Hq. + assert (~~ pdec (P k)) as Hk. { cbv. tauto. } + apply Hk. clear Hk. intros Hk. + apply IHk. 2: assumption. + intros Ha. apply H. intros. + assert (k' = k \/ k' < k) as [-> | ] by lia. + assumption. + now apply Ha. + Qed. + + Lemma attend_at_most_once_bound k: + ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). + Proof. + eapply finite_decs. + intros H % (attend_at_most_once_bound_gen (k := k)). + tauto. + Qed. + + Lemma attend_at_most_once_bound_test: + LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k -> ∀ s', s < s' -> ~ attend e s'). + Proof. + intros Hlem k. apply attend_at_most_once_bound_gen. + intros. eapply assume_Σ_1_lem. apply Hlem. eauto. + Qed. + + Lemma non_finite_not_bounded e: + non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ + (forall n, forall i, i <= e -> wall i (P_func n) n < x). + Proof. + intro H. unfold non_finite in H. + intros He. rewrite non_finite_nat in H. + apply (@wall_of_wall e). intros [w Hw]. + pose (N := max (2*e + 1) w). specialize (H N). + apply H. intros [m [Hm1 [k Hmk]]]. + apply He. exists k, m. + repeat split. now exists k. + lia. intros n i Hie. specialize (Hw i n Hie). lia. + Qed. Lemma ext_pick_attend N e: e < N -> ext_pick (P_func N) N e -> From 89b0641d62d9e2ba630a6c303e5bd2dbd2522921 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 21 Jan 2025 14:42:29 +0100 Subject: [PATCH 45/54] polish --- theories/PostsProblem/low_simple_predicates.v | 2 +- theories/PostsProblem/lowness.v | 141 +++++++++--------- theories/PostsProblem/simpleness.v | 1 + 3 files changed, 71 insertions(+), 73 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index 068537a..7aed1e0 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -25,7 +25,7 @@ Post's Problem in Turing degree. **) Definition low (P: nat -> Prop) := P´ ⪯ᴛ K. Section LowFacts. - + Variable vec_to_nat : forall k, vec nat k -> nat. Variable nat_to_vec : forall k, nat -> vec nat k. Variable vec_nat_inv : forall k v, nat_to_vec k (vec_to_nat v) = v. diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index 9f54297..b76886e 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -25,7 +25,7 @@ Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. Global Instance dec_le: ∀ m n, dec (m ≤ n). Proof. intros n m; destruct (le_gt_dec n m); [by left|right; lia]. Qed. -Section Requirements_Lowness_Correctness. +Section requirements_verification. Variable P: nat → Prop. Variable f: nat → nat → bool. @@ -35,17 +35,15 @@ Section Requirements_Lowness_Correctness. Fact Φ_spec e x: Ξ e (char_rel P) x → ∞∀ n, Φ_ e x n ↓. Proof. intro H. unfold Φ_. apply (Φ_spec S_P H). Qed. - Definition Ω e n := Φ_ e e n. - - Section classic_logic. + Definition limit_decider e n: bool := Dec (Ω e n ↓). (** ** Requirements *) - Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. - Definition limit_decider e n: bool := Dec (Ω e n ↓). + Hypothesis step_ex_spec: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. + Hypothesis N_requirements: ∀ e, (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). - Section with_LEM_2. + (* Section with_LEM_2. Hypothesis LEM_Σ_2: ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). @@ -90,11 +88,9 @@ Section Requirements_Lowness_Correctness. enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. Qed. - End with_LEM_2. + End with_LEM_2. *) - Section with_LEM_1. - Hypothesis convergent: ∀ e, (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). Lemma Jump_limit_1 : limit_computable (P´). Proof. @@ -102,13 +98,13 @@ Section Requirements_Lowness_Correctness. - unfold J. split. intros [w Hw]%Φ_spec; exists w; intros??. apply Dec_auto. by eapply Hw. - intros [N HN]. eapply N_requirements. + intros [N HN]. eapply step_ex_spec. intros m. exists (S N + m); split; first lia. eapply Dec_true. eapply HN. lia. - unfold J. split; intros H. - destruct (convergent x) as [[k Hk]|h2]. + destruct (N_requirements x) as [[k Hk]|h2]. enough (Ξ x (char_rel P) x) by easy. - eapply N_requirements. intros m. exists (S k + m). + eapply step_ex_spec. intros m. exists (S k + m). split; first lia. eapply Hk. lia. destruct h2 as [w Hw]. exists w. intros. specialize (Hw n H0). unfold limit_decider. @@ -120,9 +116,9 @@ Section Requirements_Lowness_Correctness. enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. Qed. - End with_LEM_1. - End classic_logic. + End requirements_verification. + (* Section intuitionistic_logic. Hypothesis N_requirements: ∀ e, (∞∃ n, Ω e n ↓) → ¬¬ Ξ e (char_rel P) e. @@ -196,14 +192,14 @@ Section Requirements_Lowness_Correctness. End intuitionistic_logic. *) -End Requirements_Lowness_Correctness. - Section Wall. (** ** Construction *) - Instance wall_instance: Wall := λ e L n, φ (λ x, Dec (In x L)) e e n. + Instance wall_instance: Wall + := λ e L n, φ (λ x, Dec (In x L)) e e n. + Definition P := P wall. Definition P_func := P_func wall. Instance E: Extension := simple_extension wall. @@ -215,21 +211,21 @@ Section Wall. Definition P_Φ := (Φ_ χ). Definition P_Ω := (Ω χ). - Section Verification. + Section Construction. (** ** Verification *) Hypothesis Σ_1_lem: LEM_Σ 1. - Lemma attend_at_most_once_bound_constructive: + Lemma attention_bound: ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ attend wall e s'. Proof. by apply attend_at_most_once_bound_test. Qed. - Lemma eventally_wall: + Lemma eventally_greater_than_wall: ∀ e, (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). Proof. intros e. - destruct (@attend_at_most_once_bound_constructive (S e)) as [s Hs]. + destruct (@attention_bound (S e)) as [s Hs]. exists (S s). intros m Hm x [e_ [He_ He_']]. destruct (Dec (e_ < e)) as [E|E]. { exfalso. enough (attend wall e_ m). @@ -239,11 +235,46 @@ Section Wall. assert (e ≤ e_) by lia; clear E. destruct He_', H1, H1, H1, H1, H3. by eapply H5. - Qed. + Qed. + + Lemma N_requirements e : (∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓). + Proof. + destruct (@eventally_greater_than_wall e) as [N HN]. + assert (pdec (∃ k, N ≤ k ∧ P_Ω e k ↓)) as [[k [Hk HNk]]|H'] by (apply assume_Σ_1_lem; eauto). + - left. exists k. intros n Hm. + induction Hm; first done. + unfold P_Ω, Ω, Φ_ in *. + destruct (φ (χ m) e e m) eqn: E. + { eapply φ_spec0' in E. congruence. } + (* Check φ_spec2. *) + eapply (@φ_spec2 χ _ ); eauto. + intros x Hx; split. + + intros K%Dec_true. apply Dec_auto. + enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). + eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. + inv HE. + * assert (wall e (P_func m) m < a). + { eapply HN. lia. enough (P_func m = l) as ->. apply H2. + eapply F_uni; [apply F_func_correctness|apply H1]. } + assert (wall e (P_func m) m = S n). { unfold wall, wall_instance. + rewrite <-E. reflexivity. } + rewrite H3 in H. destruct (Dec (a = x)). + lia. apply Dec_auto. rewrite <- H0 in K. + destruct K; first done. + enough ((F_func (simple_extension wall) m) = l) as -> by done. + eapply F_uni; last apply H1; apply F_func_correctness. + * apply Dec_auto. + enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. + eapply F_uni; first apply F_func_correctness. + assumption. + - right. exists N. intros m Hm. + destruct (Dec (P_Ω e m ↓)); eauto. + Qed. Fact wall_convergence e : ∃ b : nat, lim_to wall (wall e) b. Proof. - destruct (@eventally_wall e) as [N HN]. + destruct (@eventally_greater_than_wall e) as [N HN]. assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. assert (pdec (∀ x, N ≤ x → wall e (P_func x) x = 0)) as [H'|H']. { apply assume_Π_1_lem. apply Σ_1_lem. intros x. eauto. } @@ -285,10 +316,10 @@ Section Wall. apply H. now exists x; split. Qed. - Lemma N_requirements: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. + Lemma step_ex_spec1: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. Proof. intros e He. - destruct (@eventally_wall e) as [N HN]. + destruct (@eventally_greater_than_wall e) as [N HN]. destruct (@wall_convergence e) as [B [b Hb]]. set (M := max N b). destruct (He M) as [k [Hk Hk']]. eapply (@φ_spec χ e e k); first apply Hk'. @@ -314,47 +345,13 @@ Section Wall. eapply F_func_correctness. Qed. - Lemma convergent e : (∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓). - Proof. - destruct (@eventally_wall e) as [N HN]. - assert (pdec (∃ k, N ≤ k ∧ P_Ω e k ↓)) as [[k [Hk HNk]]|H'] by (apply assume_Σ_1_lem; eauto). - - left. exists k. intros n Hm. - induction Hm; first done. - unfold P_Ω, Ω, Φ_ in *. - destruct (φ (χ m) e e m) eqn: E. - { eapply φ_spec0' in E. congruence. } - (* Check φ_spec2. *) - eapply (@φ_spec2 χ _ ); eauto. - intros x Hx; split. - + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). - eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. - inv HE. - * assert (wall e (P_func m) m < a). - { eapply HN. lia. enough (P_func m = l) as ->. apply H2. - eapply F_uni; [apply F_func_correctness|apply H1]. } - assert (wall e (P_func m) m = S n). { unfold wall, wall_instance. - rewrite <-E. reflexivity. } - rewrite H3 in H. destruct (Dec (a = x)). - lia. apply Dec_auto. rewrite <- H0 in K. - destruct K; first done. - enough ((F_func (simple_extension wall) m) = l) as -> by done. - eapply F_uni; last apply H1; apply F_func_correctness. - * apply Dec_auto. - enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. - eapply F_uni; first apply F_func_correctness. - assumption. - - right. exists N. intros m Hm. - destruct (Dec (P_Ω e m ↓)); eauto. - Qed. Fact P_simple: simple P. Proof. eapply P_simple. intro e. intros H. apply H. apply wall_convergence. Qed. - Hypothesis Σ_2_LEM: + (* Hypothesis Σ_2_LEM: ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). @@ -364,9 +361,9 @@ Section Wall. intros e He. eapply N_requirements; eauto. Qed. End Verification. - + *) - Lemma eventally_wall_db e: + Lemma eventally_greater_than_wall_classically e: ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). Proof. intros H_. eapply (@attend_at_most_once_bound wall (S e)). @@ -382,10 +379,10 @@ Section Wall. by eapply H5. Qed. - Fact wall_convergence_db e : ¬¬ ∃ b : nat, lim_to wall (wall e) b. + Fact wall_convergence_classically e : ¬¬ ∃ b : nat, lim_to wall (wall e) b. Proof. intros H_. - destruct (@eventally_wall_db e). intros [N HN]. + destruct (@eventally_greater_than_wall_classically e). intros [N HN]. assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. - apply H_; clear H_. exists 0. exists N. intros. by apply H'. @@ -424,7 +421,7 @@ Section Wall. apply H0. intros Hmn. apply H. now exists x; split. Qed. - +(* Lemma N_requirements_db: ∀ e, (∞∃ n, P_Ω e n ↓) → ¬ ¬ Ξ e (char_rel P) e. Proof. intros e He H_. @@ -453,12 +450,12 @@ Section Wall. - intros H%Dec_true. eapply F_with_top. exists (F_func _ k), k; split; eauto. eapply F_func_correctness. - Qed. + Qed. *) (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both eventally_wall and wall_convergence *) - Section with_LEM_2. + (* Section with_LEM_2. Hypothesis LEM_Σ_2: ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). @@ -470,7 +467,7 @@ Section Wall. intros e He. eapply DN, N_requirements_db; eauto. Qed. - End with_LEM_2. + End with_LEM_2. *) Section with_LEM_1. @@ -479,13 +476,13 @@ Section Wall. Fact jump_P_limit_2: limit_computable (P´). Proof. eapply Jump_limit_1; first apply F_with_χ. - - intros e He. eapply N_requirements; eauto. - - eapply convergent; eauto. + - intros e He. eapply step_ex_spec1; eauto. + - eapply N_requirements; eauto. Qed. End with_LEM_1. -End Wall. +End Construction. (* Check jump_P_limit_2. *) diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index 929902d..1168699 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -783,3 +783,4 @@ Section Assume_EA. End Assume_EA. + From 11b57b232c3ece60a9cb75acd29c2ab5b35194ed Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 21 Jan 2025 17:47:33 +0100 Subject: [PATCH 46/54] fix lowness --- theories/PostsProblem/lowness.v | 250 ++++++++++---------------------- 1 file changed, 74 insertions(+), 176 deletions(-) diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index 056d655..afa5289 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -37,69 +37,20 @@ Section requirements_verification. Hypothesis step_ex_spec: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. Hypothesis N_requirements: ∀ e, (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). - - (* Section with_LEM_2. - - Hypothesis LEM_Σ_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). - - Hypothesis LEM_Π_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∀ n, ∃ m, P n m) ∨ ¬ (∀ n, ∃ m, P n m). - - Lemma limit_case e: (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). - Proof. - assert (∀ m n, dec (m ≤ n → ¬ Ω e n ↓)) as H by eauto. - destruct (LEM_Σ_2 H); first by right. - left. apply Φ_spec, N_requirements. intros i. - assert (∀ n, dec (n ≤ i ∧ Ω e n ↓)) as H'' by eauto. - destruct (@LEM_Σ_2 (λ n _, i ≤ n ∧ Ω e n ↓) ) as [[w Hw]|h1]; first eauto. - exists w; apply Hw; exact 42. - assert (∀ n, i ≤ n → ~ Ω e n ↓). - { intros m Hm HM. eapply h1. exists m; eauto. } - exfalso. eapply H0. by exists i. - Qed. - - Lemma Jump_limit : limit_computable (P´). - Proof. - exists limit_decider; split; intros. - - unfold J. split. - intros [w Hw]%Φ_spec; exists w; intros??. - apply Dec_auto. by eapply Hw. - intros [N HN]. eapply N_requirements. - intros m. exists (Datatypes.S N + m)%nat; split; first lia. - eapply Dec_true. eapply HN. lia. - - unfold J. split; intros H. - destruct (limit_case x) as [[k Hk]|h2]. - enough (Ξ x (char_rel P) x) by easy. - eapply N_requirements. intros m. exists (Datatypes.S k + m). - split; first lia. eapply Hk. lia. - destruct h2 as [w Hw]. exists w. - intros. specialize (Hw n H0). unfold limit_decider. - destruct (Dec _); eauto. - destruct H as [w Hw]. - intros [k Hneg]%Φ_spec. - set (N := Datatypes.S (max w k)). - assert (Ω x N ↓). { eapply Hneg. lia. } - enough (¬ Ω x N ↓) by eauto. - eapply Dec_false. eapply Hw. lia. - Qed. - End with_LEM_2. *) - - - + Lemma Jump_limit_1 : limit_computable (P´). Proof. exists limit_decider; split; intros. - unfold J. split. intros [w Hw]%Φ_spec; exists w; intros??. apply Dec_auto. by eapply Hw. - intros [N HN]. eapply N_requirements. + intros [N HN]. eapply step_ex_spec. intros m. exists (Datatypes.S N + m); split; first lia. eapply Dec_true. eapply HN. lia. - unfold J. split; intros H. destruct (N_requirements x) as [[k Hk]|h2]. enough (Ξ x (char_rel P) x) by easy. - eapply N_requirements. intros m. exists (Datatypes.S k + m). + eapply step_ex_spec. intros m. exists (Datatypes.S k + m). split; first lia. eapply Hk. lia. destruct h2 as [w Hw]. exists w. intros. specialize (Hw n H0). unfold limit_decider. @@ -111,32 +62,28 @@ Section requirements_verification. enough (¬ Ω x N ↓) by eauto. eapply Dec_false. eapply Hw. lia. Qed. - - End classic_logic. - -Section Wall. +End requirements_verification. - (** ** Construction *) - Instance wall_instance: Wall - := λ e L n, φ (λ x, Dec (In x L)) e e n. +Section requirements_meet. - Definition P := P wall. - Definition P_func := P_func wall. - Instance E: Extension := simple_extension wall. + Variables wall: Wall. - Fact P_semi_decidable: semi_decidable P. - Proof. eapply P_semi_decidable. Qed. + Definition P := P wall. + Definition P_func := P_func wall. + Instance E: Extension := simple_extension wall. - Definition χ := χ (simple_extension wall). - Definition P_Φ := (Φ_ χ). - Definition P_Ω := (Ω χ). + Fact P_semi_decidable: semi_decidable P. + Proof. eapply P_semi_decidable. Qed. - Section Construction. + Definition χ := χ (simple_extension wall). + Definition P_Φ := (Φ_ χ). + Definition P_Ω := (Ω χ). - (** ** Verification *) + Section wall_greater_than_use. + Hypothesis wall_spec: forall e L n, φ (λ x, Dec (In x L)) e e n ≤ wall e L n. Hypothesis Σ_1_lem: LEM_Σ 1. Lemma attention_bound: @@ -179,9 +126,9 @@ Section Wall. * assert (wall e (P_func m) m < a). { eapply HN. lia. enough (P_func m = l) as ->. apply H2. eapply F_uni; [apply F_func_correctness|apply H1]. } - assert (wall e (P_func m) m = S n). { unfold wall, wall_instance. - rewrite <-E. reflexivity. } - rewrite H3 in H. destruct (Dec (a = x)). + assert (wall e (P_func m) m ≥ S n). { + rewrite <-E. apply wall_spec. } + destruct (Dec (a = x)). lia. apply Dec_auto. rewrite <- H0 in K. destruct K; first done. enough ((F_func (simple_extension wall) m) = l) as -> by done. @@ -194,20 +141,32 @@ Section Wall. destruct (Dec (P_Ω e m ↓)); eauto. Qed. + End wall_greater_than_use. + +End requirements_meet. + +Section concret_wall. + + (** ** Construction *) + + Definition wall: Wall := λ e L n, φ (λ x, Dec (In x L)) e e n. + Instance E_low: Extension := simple_extension wall. + Hypothesis Σ_1_lem: LEM_Σ 1. + Fact wall_convergence e : ∃ b : nat, lim_to wall (wall e) b. Proof. - destruct (@eventally_greater_than_wall e) as [N HN]. - assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. - assert (pdec (∀ x, N ≤ x → wall e (P_func x) x = 0)) as [H'|H']. + destruct (eventally_greater_than_wall wall Σ_1_lem e) as [N HN]. + assert (∀ m, dec (wall e (P_func wall m) m = 0)) as HD by eauto. + assert (pdec (∀ x, N ≤ x → wall e (P_func wall x) x = 0)) as [H'|H']. { apply assume_Π_1_lem. apply Σ_1_lem. intros x. eauto. } (* ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. *) - exists 0. exists N. intros. by apply H'. - - enough (∃ x, N ≤ x ∧ wall e (P_func x) x ≠ 0) as H''. + - enough (∃ x, N ≤ x ∧ wall e (P_func wall x) x ≠ 0) as H''. clear H'. destruct H'' as [x [H1 H2]]. - destruct (wall e (P_func x) x) as [|k] eqn: H; first done; clear H2. + destruct (wall e (P_func wall x) x) as [|k] eqn: H; first done; clear H2. exists (S k), x. intros t Ht. induction Ht; first done. - rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). + rewrite <- (@φ_spec1 (χ wall) _ _ _ _ IHHt). reflexivity. intros; split. + intros K%Dec_true. apply Dec_auto. @@ -215,10 +174,10 @@ Section Wall. eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. inv HE. - * assert (wall e (P_func m) m < a). - { eapply HN. lia. enough (P_func m = l) as ->. apply H5. + * assert (wall e (P_func wall m) m < a). + { eapply HN. lia. enough (P_func wall m = l) as ->. apply H5. eapply F_uni; [apply F_func_correctness|apply H4]. } - assert (wall e (P_func m) m = S k). { rewrite <-IHHt. reflexivity. } + assert (wall e (P_func wall m) m = S k). { rewrite <-IHHt. reflexivity. } rewrite H6 in H2. destruct (Dec (a = x0)). lia. apply Dec_auto. rewrite <- H3 in K. destruct K; first done. @@ -238,55 +197,38 @@ Section Wall. apply H. now exists x; split. Qed. - Lemma step_ex_spec1: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. + Lemma step_ex_spec: ∀ e, (∞∃ n, P_Ω wall e n ↓) → Ξ e (char_rel (P wall)) e. Proof. intros e He. - destruct (@eventally_greater_than_wall e) as [N HN]. + destruct (@eventally_greater_than_wall wall Σ_1_lem e) as [N HN]. destruct (@wall_convergence e) as [B [b Hb]]. set (M := max N b). destruct (He M) as [k [Hk Hk']]. - eapply (@φ_spec χ e e k); first apply Hk'. + eapply (@φ_spec (χ wall) e e k); first apply Hk'. intros x Hx. unfold P, simpleness.P. rewrite F_with_top. split. - intros (L & m & HL & HLs &HP). - assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } - assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } + assert (L = (P_func wall) m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } + assert (x::L = (P_func wall) (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. - enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. + enough (incl ((P_func wall) (S m)) ((P_func wall) k)). rewrite <-E' in H. eauto. eapply F_mono; last apply E_; apply F_func_correctness. assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). - assert (k ≤ m) as Ek by lia. enough (x ≤ φ (χ m) e e m). - exfalso. assert (φ (χ m) e e m < x). - apply HN. lia. enough (φ (χ m) e e m = φ (χ k) e e k) by congruence. - assert (φ (χ m) e e m = B). + assert (k ≤ m) as Ek by lia. enough (x ≤ φ (χ wall m) e e m). + exfalso. assert (φ (χ wall m) e e m < x). + apply HN. lia. enough (φ (χ wall m) e e m = φ (χ wall k) e e k) by congruence. + assert (φ (χ wall m) e e m = B). { rewrite <- (Hb m). reflexivity. lia. } - assert (φ (χ k) e e k = B). + assert (φ (χ wall k) e e k = B). { rewrite <- (Hb k). reflexivity. lia. } congruence. - intros H%Dec_true. - eapply F_with_top. exists (F_func _ k), k; split; eauto. + eapply F_with_top. + exists (F_func _ k), k; split; eauto. eapply F_func_correctness. Qed. - - Fact P_simple: simple P. - Proof. eapply P_simple. intro e. - intros H. apply H. apply wall_convergence. - Qed. - - (* Hypothesis Σ_2_LEM: - ∀ (P: nat → nat → Prop), - (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). - - Fact jump_P_limit_test: limit_computable (P´). - Proof. - eapply Jump_limit; last done. apply F_with_χ. - intros e He. eapply N_requirements; eauto. - Qed. - End Verification. - *) - Lemma eventally_greater_than_wall_classically e: - ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). + ¬¬ (∞∀ s, ∀ x, extendP (P_func wall s) s x → wall e (P_func wall s) s < x). Proof. intros H_. eapply (@attend_at_most_once_bound wall (S e)). intros [s Hs]. eapply H_; clear H_. @@ -305,16 +247,16 @@ Section Wall. Proof. intros H_. destruct (@eventally_greater_than_wall_classically e). intros [N HN]. - assert (∀ m, dec (wall e (P_func m) m = 0)) as HD by eauto. - ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. + assert (∀ m, dec (wall e (P_func wall m) m = 0)) as HD by eauto. + ccase (∀ x, N ≤ x → wall e (P_func wall x) x = 0) as [H'|H']. - apply H_; clear H_. exists 0. exists N. intros. by apply H'. - - enough (~~∃ x, N ≤ x ∧ wall e (P_func x) x ≠ 0) as H__. + - enough (~~∃ x, N ≤ x ∧ wall e (P_func wall x) x ≠ 0) as H__. apply H__. intros H''. clear H'. destruct H'' as [x [H1 H2]]. apply H_; clear H_. - destruct (wall e (P_func x) x) as [|k] eqn: H; first done; clear H2. + destruct (wall e (P_func wall x) x) as [|k] eqn: H; first done; clear H2. exists (S k), x. intros t Ht. induction Ht; first done. - rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). + rewrite <- (@φ_spec1 (χ wall) _ _ _ _ IHHt). reflexivity. intros; split. + intros K%Dec_true. apply Dec_auto. @@ -322,10 +264,10 @@ Section Wall. eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. inv HE. - * assert (wall e (P_func m) m < a). - { eapply HN. lia. enough (P_func m = l) as ->. apply H5. + * assert (wall e (P_func wall m) m < a). + { eapply HN. lia. enough (P_func wall m = l) as ->. apply H5. eapply F_uni; [apply F_func_correctness|apply H4]. } - assert (wall e (P_func m) m = S k). { rewrite <-IHHt. reflexivity. } + assert (wall e (P_func wall m) m = S k). { rewrite <-IHHt. reflexivity. } rewrite H6 in H2. destruct (Dec (a = x0)). lia. apply Dec_auto. rewrite <- H3 in K. destruct K; first done. @@ -343,70 +285,26 @@ Section Wall. apply H0. intros Hmn. apply H. now exists x; split. Qed. -(* - Lemma N_requirements_db: ∀ e, (∞∃ n, P_Ω e n ↓) → ¬ ¬ Ξ e (char_rel P) e. - Proof. - intros e He H_. - apply (@eventally_wall_db e). intros [N HN]. - apply (@wall_convergence_db e). intros [B [b Hb]]. - apply H_; clear H_. - set (M := max N b). destruct (He M) as [k [Hk Hk']]. - eapply (@φ_spec χ e e k); first apply Hk'. - intros x Hx. unfold P, simpleness.P. - rewrite F_with_top. split. - - intros (L & m & HL & HLs &HP). - assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } - assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } - apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. - enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. - eauto. eapply F_mono; last apply E_; apply F_func_correctness. - assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). - assert (k ≤ m) as Ek by lia. enough (x ≤ φ (χ m) e e m). - exfalso. assert (φ (χ m) e e m < x). - apply HN. lia. enough (φ (χ m) e e m = φ (χ k) e e k) by congruence. - assert (φ (χ m) e e m = B). - { rewrite <- (Hb m). reflexivity. lia. } - assert (φ (χ k) e e k = B). - { rewrite <- (Hb k). reflexivity. lia. } - congruence. - - intros H%Dec_true. - eapply F_with_top. exists (F_func _ k), k; split; eauto. - eapply F_func_correctness. - Qed. *) - - (* Should be enough if we use DNE_Σ_2 to drop out ¬¬ for both - eventally_wall and wall_convergence *) - - (* Section with_LEM_2. - - Hypothesis LEM_Σ_2: - ∀ (P: nat → nat → Prop), (∀ n m, dec (P n m)) → (∃ n, ∀ m, P n m) ∨ ¬ (∃ n, ∀ m, P n m). - Hypothesis DN: ∀ P, ¬ ¬ P → P. - Fact jump_P_limit: limit_computable (P´). - Proof. - eapply Jump_limit; last done. apply F_with_χ. - intros e He. eapply DN, N_requirements_db; eauto. - Qed. + Fact P_simple: simple (P wall). + Proof. eapply P_simple, wall_convergence_classically. Qed. - End with_LEM_2. *) +End concret_wall. - Section with_LEM_1. +Section result. - Hypothesis LEM_Σ_1: LEM_Σ 1. +Hypothesis LEM_Σ_1: LEM_Σ 1. - Fact jump_P_limit_2: limit_computable (P´). - Proof. - eapply Jump_limit_1; first apply F_with_χ. - - intros e He. eapply step_ex_spec1; eauto. - - eapply N_requirements; eauto. - Qed. +Fact jump_P_limit_2: limit_computable ((P wall)´). +Proof. + eapply Jump_limit_1; first apply F_with_χ. + - intros e He. eapply step_ex_spec; eauto. + - eapply N_requirements; eauto. +Qed. - End with_LEM_1. +End result. -End Construction. -(* Check jump_P_limit_2. *) From b6c341bb8a4e8593ca4dc971c8f014f31605c673 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 21 Jan 2025 22:10:34 +0100 Subject: [PATCH 47/54] refactoring --- theories/PostsProblem/low_simple_predicates.v | 17 +- theories/PostsProblem/lowness.v | 60 +- theories/PostsProblem/simpleness.v | 1163 ++++++++--------- 3 files changed, 604 insertions(+), 636 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index ee9090e..ebd427d 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -102,17 +102,6 @@ Section LowFacts. Hypothesis EA: forall P, semi_decidable P -> exists e, forall x, P x <-> exists n, η e n = Some x. - (* Used to eliminate ~~Φ *) - (* should be able to weaker to DNE_Σ_2 *) - Hypothesis DN: forall P, ~ ~ P -> P. - - (* Used to prove limit computable from N requirements *) - Hypothesis LEM_Σ_2: - forall (P: nat -> nat -> Prop), - (forall n m, dec (P n m)) -> - (exists n, forall m, P n m) \/ ~ (exists n, forall m, P n m). - - (* Used to prove Limit Lemma *) Hypothesis LEM_Σ_1: LEM_Σ 1. Hypothesis def_K: definite K. @@ -123,7 +112,7 @@ Section LowFacts. apply jump_P_limit; eauto. - eapply P_simple. intros. intros d. apply d. - apply wall_convergence_test. assumption. + apply wall_convergence. assumption. Qed. End LowSimplePredicate. @@ -143,7 +132,7 @@ Section LowFacts. eexists. repeat split. - apply simple_undecidable. - eapply P_simple. apply wall_convergence. + eapply P_simple. apply wall_convergence_classically. - apply P_semi_decidable. - intros L. intros G. apply L. clear L. intros L. assert (~~ definite K) as hK. @@ -166,7 +155,7 @@ Section LowFacts. } revert G. apply lowness. red. eapply limit_turing_red_K; eauto. exact 42. - apply jump_P_limit_2; eauto. + apply jump_P_limit; eauto. Qed. End LowFacts. diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index afa5289..5674e02 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -8,19 +8,19 @@ From SyntheticComputability Require Import the_priority_method. From SyntheticComputability Require Import simpleness. Definition inf_exists (P: nat → Prop) := ∀ n, ∃ m, n ≤ m ∧ P m. - Notation "'∞∃' x .. y , p" := - (inf_exists (λ x, .. (inf_exists (λ y, p)) ..)) - (at level 200, x binder, right associativity, - format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") - : type_scope. +Global Instance dec_le: ∀ m n, dec (m ≤ n). + Proof. intros n m; destruct (le_gt_dec n m); [by left|right; lia]. Qed. - Notation "f ↓" := (f = Some ()) (at level 30). +Notation "'∞∃' x .. y , p" := + (inf_exists (λ x, .. (inf_exists (λ y, p)) ..)) + (at level 200, x binder, right associativity, + format "'[' '∞∃' '/ ' x .. y , '/ ' p ']'") + : type_scope. - Global Instance dec_le: ∀ m n, dec (m ≤ n). - Proof. intros n m; destruct (le_gt_dec n m); [by left|right; lia]. Qed. +Notation "f ↓" := (f = Some ()) (at level 30). -Section requirements_verification. +Section Requirements_Verification. Variable P: nat → Prop. Variable f: nat → nat → bool. @@ -63,10 +63,9 @@ Section requirements_verification. eapply Dec_false. eapply Hw. lia. Qed. -End requirements_verification. +End Requirements_Verification. - -Section requirements_meet. +Section Requirements_Meet. Variables wall: Wall. @@ -87,8 +86,8 @@ Section requirements_meet. Hypothesis Σ_1_lem: LEM_Σ 1. Lemma attention_bound: - ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ attend wall e s'. - Proof. by apply attend_at_most_once_bound_test. Qed. + ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ recv_att wall e s'. + Proof. by apply recv_at_most_once_bound. Qed. Lemma eventally_greater_than_wall: ∀ e, (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). @@ -97,7 +96,7 @@ Section requirements_meet. destruct (@attention_bound (S e)) as [s Hs]. exists (S s). intros m Hm x [e_ [He_ He_']]. destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (attend wall e_ m). + { exfalso. enough (recv_att wall e_ m). unshelve eapply (Hs e_ _ m); eauto; lia. split; first lia. destruct He_' as [H _]. apply H. } @@ -143,9 +142,9 @@ Section requirements_meet. End wall_greater_than_use. -End requirements_meet. +End Requirements_Meet. -Section concret_wall. +Section Concret_Wall. (** ** Construction *) @@ -230,11 +229,11 @@ Section concret_wall. Lemma eventally_greater_than_wall_classically e: ¬¬ (∞∀ s, ∀ x, extendP (P_func wall s) s x → wall e (P_func wall s) s < x). Proof. - intros H_. eapply (@attend_at_most_once_bound wall (S e)). + intros H_. eapply (@recv_at_most_once_bound_classically wall (S e)). intros [s Hs]. eapply H_; clear H_. exists (S s). intros m Hm x [e_ [He_ He_']]. destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (attend wall e_ m). + { exfalso. enough (recv_att wall e_ m). unshelve eapply (Hs e_ _ m); eauto; lia. split; first lia. destruct He_' as [H _]. apply H. } @@ -289,23 +288,20 @@ Section concret_wall. Fact P_simple: simple (P wall). Proof. eapply P_simple, wall_convergence_classically. Qed. -End concret_wall. - -Section result. +End Concret_Wall. -Hypothesis LEM_Σ_1: LEM_Σ 1. - -Fact jump_P_limit_2: limit_computable ((P wall)´). -Proof. - eapply Jump_limit_1; first apply F_with_χ. - - intros e He. eapply step_ex_spec; eauto. - - eapply N_requirements; eauto. -Qed. - -End result. +Section Results. + Hypothesis LEM_Σ_1: LEM_Σ 1. + Fact jump_P_limit: limit_computable ((P wall)´). + Proof. + eapply Jump_limit_1; first apply F_with_χ. + - intros e He. eapply step_ex_spec; eauto. + - eapply N_requirements; eauto. + Qed. +End Results. diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index 33f4132..a25c246 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -149,599 +149,582 @@ Section Assume_EA. Qed. Section Assume_WALL. - Class Wall := wall: nat → list nat → nat → nat. - Variable wall_instance: Wall. - - Section Extension. - - (** ** Construction *) - - Definition ext_intersect_W L n e := L # W[n] e. - Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. - Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. - Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. - Definition ext_least_choice L n x := exists e, ext_choice L n e x. - - End Extension. - - Section Extension_Facts. - - #[export]Instance ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). - Proof. - apply BaseLists.list_forall_dec. - intro x. eapply dec_neg_dec, exists_bounded_dec; eauto. - Qed. - - #[export] Instance ext_wall L n e x: dec (forall i, i <= e -> wall i L n < x). - Proof. eapply forall_bounded_dec; eauto. Qed. - - #[export]Instance ext_has_wit_dec L n e x : dec (ext_has_wit L n e x). - Proof. apply and_dec; first apply exists_bounded_dec; eauto. Qed. - - #[export]Instance ext_has_wit_exists_dec L n e : dec (exists x, ext_has_wit L n e x). - Proof. - unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. - intro x; eapply W_bounded_dec; eauto. eauto. - Qed. - - #[export]Instance ext_pick_dec L n e : dec (ext_pick L n e). - Proof. - eapply and_dec; first apply ext_intersect_W_dec. - unfold ext_has_wit. - eapply bounded_dec; last apply W_bounded_bounded. - intros x. eapply exists_bounded_dec. - intro; apply W_dec. eauto. - Qed. - - #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). - Proof. - eapply exists_bounded_dec'. intro x. - eapply least_dec. intros y. - eapply ext_pick_dec. - Qed. - - Fact ext_least_choice_dec L n: - (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). - Proof. - unfold ext_least_choice. - destruct (ext_pick_exists_dec L n) as [H|H]. - - left. apply ewo_nat; first last. - destruct H as [e (h1 & [(h4 & h2) h3])]. - eapply least_ex in h2. destruct h2 as [k h6]. - exists k, e. split; first easy; split. - 2: { eapply h6. } - split; last apply h3. split; first apply h4. - destruct h6. now exists k. - eapply ext_has_wit_dec. - intro x. unfold ext_choice. eapply exists_bounded_dec'. - intros x'. eapply and_dec. - apply least_dec. eapply ext_pick_dec. - apply least_dec. eapply ext_has_wit_dec. - - right. intros x [e (h1 & h2 & _)]. apply H. - exists e. split; eauto. - Qed. - - Fact ext_least_choice_uni l x: unique (ext_least_choice l x). - Proof. - intros y1 y2 [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. - assert (x0 = x1) as ->. eapply least_unique; eauto. - eapply least_unique; eauto. - Qed. - - End Extension_Facts. - - Section Simple_Extension. - - Instance simple_extension: Extension. - Proof. - unshelve econstructor. - - exact ext_least_choice. - - apply ext_least_choice_dec. - - apply ext_least_choice_uni. - Defined. - - Definition P_ := F_ simple_extension. - Definition P_func := F_func simple_extension. - Definition P := F_with simple_extension. - - Definition non_finite e := ~ exhaustible (W e). - - Fact In_Pf k y: In y (P_func k) -> P y. - Proof. - intros H. exists (P_func k), k. - split; [easy| apply F_func_correctness]. - Qed. - - Definition lim_to (f: list nat -> nat -> nat) b := (exists n, forall m, n <= m -> f (P_func m) m = b). - - End Simple_Extension. - - Hypothesis wall_spec: forall e, ~~ exists b, lim_to (wall e) b. - - Section Requirements. - - (** ** Requirements *) - - Definition R_simple P e := non_finite e -> ~ (W e #ₚ P). - Definition attend e n := e < n /\ least (ext_pick (P_func n) n) e. - Definition act e n := ~ (P_func n) # W[n] e. - Definition pick_el e x := exists k, attend e k /\ ext_least_choice (P_func k) k x. - Definition finish e n := ∀ s, n < s → ¬ attend e s. - - #[export]Instance attend_dec e n: dec (attend e n). - Proof. - unfold attend. apply and_dec; first eauto. - eapply least_dec. intros y. - eapply ext_pick_dec. - Qed. - - End Requirements. - - Section Requirements_Facts. - - (** ** Verification *) - - Lemma ext_pick_witness L n e: - ext_pick L n e -> exists x, least (ext_has_wit L n e) x. - Proof. - intros [H1 H2]. - eapply least_ex. intro; eapply ext_has_wit_dec. - trivial. - Qed. - - Lemma W_incl e n m: - n <= m -> forall x, (W[n] e) x -> (W[m] e) x. - Proof. - intros H x [y [H1 H2]]. - exists y. split; lia + easy. - Qed. - - Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : - L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. - Proof. - intros H h1 h2 x Hx1 Hx2. - eapply (H x). now eapply h1. - now apply h2. - Qed. - - Lemma act_always_act e n: act e n -> forall m, n <= m -> act e m . - Proof. - intros H m Hm Hx. apply H. - eapply (intersect_mono Hx). - eapply F_mono; try eapply F_func_correctness; easy. - now eapply W_incl. - Qed. - - Lemma attend_act e k: attend e k -> act e (S k). - Proof. - intros [He H] Hcontr. - edestruct (ext_pick_witness) as [x Hx]. - { destruct H. eapply e0. } - apply (Hcontr x). - assert (P_ (S k) (P_func (S k))) by apply F_func_correctness. - inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. - exists e. assert (P_func k = l) as <-. - eapply F_uni. apply F_func_correctness. apply H3. - split; first easy. split; first easy. easy. - assert (x = a) as ->. eapply (@extend_uni simple_extension); cbn; eauto. - eauto. exfalso. eapply (H3 x). cbn. - exists e. enough ((P_func k) = (P_func (S k))) as <-. - split; first easy. easy. - assert (F_ simple_extension k (P_func k)) by apply F_func_correctness. - eapply F_uni; eauto. - firstorder. - Qed. - - Lemma act_not_attend e k: act e k -> ~ attend e k. - Proof. now intros h2 [_ [[h _] _]]. Qed. - - Lemma attend_always_act e k: attend e k -> forall m, k < m -> act e m. - Proof. - intros. eapply act_always_act. - apply attend_act. apply H. lia. - Qed. - - Lemma attend_always_not_attend e k: - attend e k -> forall m, k < m -> ~ attend e m. - Proof. - intros H1 m Hm. eapply act_not_attend. - apply (attend_always_act H1 Hm). - Qed. - - Lemma attend_at_most_once_gen e: - (pdec (exists k, attend e k)) -> - (exists s', forall s, s' < s -> ~ attend e s). - Proof. - intros [[w Hw]|H]. - - exists w. - now eapply attend_always_not_attend. - - exists 0. - intros k' Hk' Ha. - apply H. now exists k'. - Qed. - - Lemma attend_at_most_once e: ~ ~ (exists s', forall s, s' < s -> ~ attend e s). - Proof. - intros H. - assert (~~ (pdec (exists k, attend e k))) as Hdec. - { unfold pdec. tauto. } - apply Hdec. clear Hdec. intros Hdec. - apply H, attend_at_most_once_gen. assumption. - Qed. - - Definition done e s := ∀ s', s < s' → ¬ attend e s'. - - (* Lemma attend_at_most_once_test: *) - (* LEM_Σ 1 → ∀ e, ∃ s, done e s. *) - (* Proof. *) - (* intros Hlem e. *) - (* apply attend_at_most_once_gen. *) - (* { eapply assume_Σ_1_lem. apply Hlem. eauto. } *) - (* Qed. *) - - Lemma attend_uni e: unique (attend e). - Proof. - intros k1 k2 H1 H2. - specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H1 a b)) as H1'. - specialize (fun a b => @act_not_attend _ _ (@attend_always_act _ _ H2 a b)) as H2'. - enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. - intro Hk. eapply H1'. apply Hk. easy. - intro Hk. eapply H2'. apply Hk. easy. - Qed. - - Lemma pick_el_uni e: unique (pick_el e). - Proof. - intros x1 x2 [k [Ht Hk]] [k' [Ht' Hk']]. - assert (k=k') as <-. eapply attend_uni; eauto. - eapply (@extend_uni simple_extension); cbn; eauto. - Qed. - - End Requirements_Facts. - - Section Compl_P_non_finite. - - Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, pick_el e x /\ e < n. - Proof. - intros [[L [k [Hin Hk]]] Hn]. - dependent induction L. inv Hin. - destruct (Dec (a = x)) as [->|]. - - clear IHL Hin. - destruct (F_pick Hk) as [k' [Hk' [e He]]]. - exists e. split. - exists k'. split; unfold attend. - + destruct He; intuition eauto. enough (P_func k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. easy. - + exists e. unfold ext_choice. - enough (P_func k' = L) as -> by eauto. - eapply F_uni. apply F_func_correctness. - destruct He; intuition eauto. - + destruct He. destruct H0. destruct H1. destruct H1. - lia. - - destruct (F_pop Hk) as [m' Hm']. - eapply (IHL m'); eauto. firstorder. - Qed. - - Lemma P_extract_spec n L: - (forall x, In x L -> P x /\ x <= 2 * n) -> - forall x, In x L -> exists c, pick_el c x /\ c < n. - Proof. - intros. induction L. inv H0. - destruct H0 as [-> | Hln]; last apply IHL; eauto. - apply P_meet_spec. now apply H. - Qed. - - Lemma P_pullback_list n L: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> - exists (LC: list nat), NoDup LC /\ length LC = length L /\ - forall c, In c LC -> exists x, pick_el c x /\ In x L /\ c < n. - Proof. - intros HL H1. - induction L. - - exists []; intuition. - - remember (@P_extract_spec n (a::L) H1 a). - assert (In a (a::L)) as H0 by intuition. - apply e in H0 as [c0 [H0 E1]]. - assert (NoDup L) by (inversion HL; intuition). - apply IHL in H as [LC H]. - exists (c0::LC). intuition. - + constructor; last now exact H2. - intros In. inv HL. - apply H4 in In as [y (Hs & h2 & h3)]. - now rewrite (pick_el_uni H0 Hs) in H6. - + cbn. rewrite H. trivial. - + destruct H3 as [->|]. - * exists a; intuition. - * destruct (H4 c H3) as [y Hy]. - exists y; intuition. - + intros y In1. assert (In y (a::L)) by intuition. - now apply H1 in H2. - Qed. - - Definition PredListTo p : list nat -> nat -> Prop - := fun L b => forall x, In x L <-> p x /\ x <= b. - - Lemma NoDupBoundH {L} b: - NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). - Proof. - intros ND H x E. - constructor. - - intros H1 % H. lia. - - exact ND. - Qed. - - Lemma PredNoDupListTo_NNExist p: - forall b, ~~ exists L, PredListTo p L b /\ NoDup L. - Proof. - destruct (F_computable simple_extension) as [f Hf]. - induction b; intros H. - - ccase (p 0) as [H0 | H0]; apply H. - + exists [0]. split; try split. - * intros [E | E]; (try contradiction E). - rewrite <- E. intuition. - * intros E. assert (x = 0) by lia. - rewrite H1. intuition. - * constructor; intuition; constructor. - + exists nil. split; try split. - * contradiction. - * intros E. assert (x = 0) by lia. - rewrite H1 in E. firstorder. - * constructor. - - apply IHb. intros [L H1]. - ccase (p (1 + b)) as [H0 | H0]; apply H. - + exists ((1+ b) :: L). split; try split. - * intros [E | E]; try (rewrite <- E; intuition). - apply H1 in E. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** right. apply H1. intuition. - ** left. lia. - * apply (@NoDupBoundH _ b). - ** apply H1. - ** intros x H3 % H1. lia. - ** lia. - + exists L. split; try split. - * intros E % H1. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. - ** apply H1. intuition. - ** rewrite E in E1. firstorder. - * apply H1. - Qed. - - Lemma P_bounded L n: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. - Proof. - intros ND [LC H] % P_pullback_list; intuition. - rewrite <- H. - assert (incl LC (seq 0 n)). unfold incl. - - intros c [e [x Hx]] % H2. apply in_seq. lia. - - apply pigeonhole_length in H1. - + now rewrite seq_length in H1. - + intros. decide (x1 = x2); tauto. - + exact H0. - Qed. - - Lemma P_Listing: - forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). - Proof. - intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). - intros [L H1]. apply H. exists L; intuition. - apply P_bounded. - - exact H2. - - apply H0. - Qed. - - Lemma complToBound_compl p L b: - PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. - Proof. - intros H x. split. - - intros [H1 H1'] % in_filter_iff. - destruct Dec; cbn in H1'; try congruence. - enough (x <= b). - + firstorder. - + apply in_seq in H1. lia. - - intros [H1 H2]. eapply in_filter_iff. split. - + apply in_seq; lia. - + destruct Dec; cbn; try tauto. exfalso. firstorder. - Qed. - - Lemma compl_P_Listing: - forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L - /\ forall x, In x L -> ~ P x. - Proof. - intros n H. - apply (@P_Listing n). intros [L H1]. - apply H. exists (complToBound L (2*n)). repeat split. - - remember (complToBound_length L (2*n)). lia. - - apply complToBound_NoDup. - - intros x I % (@complToBound_compl P); intuition. - Qed. - - Lemma P_coinfinite : ~ exhaustible (compl P). - Proof. - eapply weakly_unbounded_non_finite. - intros n H. eapply compl_P_Listing with (n := n). - intros (l & ? & ? & H2). - eapply H. - exists (firstn n l). - repeat split. - - rewrite firstn_length. lia. - - now eapply firstn_NoDup. - - intros ? ? % firstn_In. now eapply H2. - Qed. - - End Compl_P_non_finite. - - Section Meet_Requirement. - - Lemma wall_of_wall' e: ~~ exists w, forall x, wall e (P_func x) x < w. - Proof. - intro H_. apply (@wall_spec e). intros [v [k Hk]]. - apply H_. - destruct (@extra_bounded (fun k => wall e (P_func k) k) k) as [w Hw]. - exists (S (max v w)). intros x. - destruct (Dec (x < k)). apply Hw in l. lia. - assert (wall e (P_func x) x = v). apply Hk. lia. lia. - Qed. - - Lemma wall_of_wall e: - ~~ exists w, forall i x, i <= e -> wall i (P_func x) x < w. - Proof. - intro H_. induction e. - - apply (@wall_of_wall' 0). intros [w Hw]. apply H_. exists w. - intros i x ?. inv H. trivial. - - apply IHe. intros [w IHe_]. - apply (@wall_of_wall' (S e)). intros [w' Hw']. - apply H_. exists (S (max w w')). intros i x H. - inv H. specialize (Hw' x). lia. - specialize (IHe_ i x H1). lia. - Qed. - - Lemma attend_at_most_once_bound_gen k: - (forall k', k' < k -> pdec (∃ k0 : nat, attend k' k0)) -> - exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). - Proof. - intros Hle. - induction k. - - exists 42. intros ??. lia. - - destruct IHk as [s Hs]. - { intros. eapply Hle. lia. } - destruct (@attend_at_most_once_gen k) as [sk Hsk]. - { apply Hle. lia. } - set (max sk s) as N. - exists N. intros e He. - assert (e = k \/ e < k) as [->|Hek] by lia. - intros s' Hs'. eapply Hsk. lia. - intros s' Hs'. eapply Hs; lia. - Qed. - - Lemma finite_decs (P : nat -> Prop) k Q : - ((forall k', k' < k -> pdec (P k')) -> ~ Q) -> ~ Q. - Proof. - induction k. - - firstorder lia. - - intros H Hq. - assert (~~ pdec (P k)) as Hk. { cbv. tauto. } - apply Hk. clear Hk. intros Hk. - apply IHk. 2: assumption. - intros Ha. apply H. intros. - assert (k' = k \/ k' < k) as [-> | ] by lia. - assumption. - now apply Ha. - Qed. - - Lemma attend_at_most_once_bound k: - ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ attend e s'). - Proof. - eapply finite_decs. - intros H % (attend_at_most_once_bound_gen (k := k)). - tauto. - Qed. - - Lemma attend_at_most_once_bound_test: - LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k -> ∀ s', s < s' -> ~ attend e s'). - Proof. - intros Hlem k. apply attend_at_most_once_bound_gen. - intros. eapply assume_Σ_1_lem. apply Hlem. eauto. - Qed. - - Lemma non_finite_not_bounded e: - non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ - (forall n, forall i, i <= e -> wall i (P_func n) n < x). - Proof. - intro H. unfold non_finite in H. - intros He. rewrite non_finite_nat in H. - apply (@wall_of_wall e). intros [w Hw]. - pose (N := max (2*e + 1) w). specialize (H N). - apply H. intros [m [Hm1 [k Hmk]]]. - apply He. exists k, m. - repeat split. now exists k. - lia. intros n i Hie. specialize (Hw i n Hie). lia. - Qed. - - Lemma ext_pick_attend N e: - e < N -> ext_pick (P_func N) N e -> - (exists w, w ≤ e /\ attend w N). - Proof. - intros HeN H1. - assert (exists w, ext_pick (P_func N) N w) by now exists e. - eapply least_ex in H; last eauto. - destruct H as [k Hk]. assert (k <= e). - { enough (~ k > e) by lia. intro Hkw. - destruct Hk. rewrite safe_char in H0. - specialize (H0 e H1). lia. } - exists k. do 2 (split; first lia). eapply Hk. - Qed. - - Lemma non_finite_attend e: - non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ attend e k) . - Proof. - intros H He. - eapply (non_finite_not_bounded H); intros (b & x & (Hx1 & Hx2 & Hx3)). - eapply (@attend_at_most_once_bound e); intros [s Hs]. - pose (N := S (max (max b s) e)). apply He. - destruct (Dec (ext_intersect_W (P_func N) N e)) as [He'|He']. - - exists N. right. - assert (ext_pick (P_func N) N e). - { split; first easy. exists x. split. - eapply W_incl; last apply Hx1. lia. - split; first lia. eapply Hx3. - } - split. lia. split. easy. - intros w HEw Hw. - eapply ext_pick_attend in Hw. - destruct Hw as [g [HEg Hg]]. - eapply Hs; last apply Hg; lia. lia. - - exists N. now left. - Qed. - - Lemma ext_intersect_W_intersect k e: - ~ (P_func k # W[k] e) -> W e #ₚ P -> False. - Proof. - unfold ext_intersect_W. - intros H1 H2. apply H1. - intros y Hy1 [w Hy2]. - eapply (H2 y). now exists w. - eapply (In_Pf Hy1). - Qed. - - Lemma P_meet_R_simple : forall e, R_simple P e. - Proof. - intros e He. intros He'. - eapply (non_finite_attend He). - intros [k [H|H]]. - - eapply ext_intersect_W_intersect; eauto. - - eapply attend_act in H. - eapply ext_intersect_W_intersect; eauto. - Qed. - - End Meet_Requirement. - - Section Result. - - Lemma P_semi_decidable : semi_decidable P. - Proof. - apply F_with_semi_decidable. - Qed. - - Theorem P_simple : simple P. - Proof. - unfold simple; repeat split. - - rewrite EA.enum_iff. now apply P_semi_decidable. - - apply P_coinfinite. - - intros [q (Hqenum & Hqinf & Hq)]. - rewrite EA.enum_iff in Hqenum. - destruct (W_spec Hqenum) as [c HWq]. - apply (@P_meet_R_simple c). - intros [l Hqe]; apply Hqinf. - exists l. intros x Hqx. apply (Hqe x). - now rewrite HWq in Hqx. - intros x HWcx HPx. unfold W in HWcx. - rewrite <- (HWq x) in HWcx. apply (Hq x HWcx HPx). - Qed. - - End Result. + Definition Wall := nat → list nat → nat → nat. + Variable wall: Wall. + + Section Extension. (** ** Construction *) + + Definition ext_intersect_W L n e := L # W[n] e. + Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. + Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. + Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. + Definition ext_least_choice L n x := exists e, ext_choice L n e x. + + End Extension. + + Section Extension_Facts. + + #[export]Instance ext_intersect_W_dec L n e: dec (ext_intersect_W L n e). + Proof. + apply BaseLists.list_forall_dec. + intro x. eapply dec_neg_dec, exists_bounded_dec; eauto. + Qed. + + #[export] Instance ext_wall L n e x: dec (forall i, i <= e -> wall i L n < x). + Proof. eapply forall_bounded_dec; eauto. Qed. + + #[export]Instance ext_has_wit_dec L n e x : dec (ext_has_wit L n e x). + Proof. apply and_dec; first apply exists_bounded_dec; eauto. Qed. + + #[export]Instance ext_has_wit_exists_dec L n e : dec (exists x, ext_has_wit L n e x). + Proof. + unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. + intro x; eapply W_bounded_dec; eauto. eauto. + Qed. + + #[export]Instance ext_pick_dec L n e : dec (ext_pick L n e). + Proof. + eapply and_dec; first apply ext_intersect_W_dec. + unfold ext_has_wit. + eapply bounded_dec; last apply W_bounded_bounded. + intros x. eapply exists_bounded_dec. + intro; apply W_dec. eauto. + Qed. + + #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). + Proof. + eapply exists_bounded_dec'. intro x. + eapply least_dec. intros y. + eapply ext_pick_dec. + Qed. + + Fact ext_least_choice_dec L n: + (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). + Proof. + unfold ext_least_choice. + destruct (ext_pick_exists_dec L n) as [H|H]. + - left. apply ewo_nat; first last. + destruct H as [e (h1 & [(h4 & h2) h3])]. + eapply least_ex in h2. destruct h2 as [k h6]. + exists k, e. split; first easy; split. + 2: { eapply h6. } + split; last apply h3. split; first apply h4. + destruct h6. now exists k. + eapply ext_has_wit_dec. + intro x. unfold ext_choice. eapply exists_bounded_dec'. + intros x'. eapply and_dec. + apply least_dec. eapply ext_pick_dec. + apply least_dec. eapply ext_has_wit_dec. + - right. intros x [e (h1 & h2 & _)]. apply H. + exists e. split; eauto. + Qed. + + Fact ext_least_choice_uni l x: unique (ext_least_choice l x). + Proof. + intros y1 y2 [? (h1 & h2 & h3)] [? (h1' & h2' & h3')]. + assert (x0 = x1) as ->. eapply least_unique; eauto. + eapply least_unique; eauto. + Qed. + + Instance simple_extension: Extension. + Proof. + unshelve econstructor. + - exact ext_least_choice. + - apply ext_least_choice_dec. + - apply ext_least_choice_uni. + Defined. + + Definition P_ := F_ simple_extension. + Definition P_func := F_func simple_extension. + Definition P := F_with simple_extension. + + Definition non_finite e := ~ exhaustible (W e). + + Fact In_Pf k y: In y (P_func k) -> P y. + Proof. + intros H. exists (P_func k), k. + split; [easy| apply F_func_correctness]. + Qed. + + Definition lim_to (f: list nat -> nat -> nat) b := + (exists n, forall m, n <= m -> f (P_func m) m = b). + + Lemma finite_decs (P : nat -> Prop) k Q : + ((forall k', k' < k -> pdec (P k')) -> ~ Q) -> ~ Q. + Proof. + induction k. + - firstorder lia. + - intros H Hq. + assert (~~ pdec (P k)) as Hk. { cbv. tauto. } + apply Hk. clear Hk. intros Hk. + apply IHk. 2: assumption. + intros Ha. apply H. intros. + assert (k' = k \/ k' < k) as [-> | ] by lia. + assumption. + now apply Ha. + Qed. + + End Extension_Facts. + + Hypothesis wall_spec: forall e, ~~ exists b, lim_to (wall e) b. + + Section Requirements. + + (** ** Requirements *) + + Definition P_requirements P e := non_finite e -> ~ (W e #ₚ P). + Definition recv_att e n := e < n /\ least (ext_pick (P_func n) n) e. + Definition act e n := ~ (P_func n) # W[n] e. + Definition act_by e x := exists k, recv_att e k /\ ext_least_choice (P_func k) k x. + Definition done e n := ∀ s, n < s → ¬ recv_att e s. + + #[export]Instance attend_dec e n: dec (recv_att e n). + Proof. + unfold recv_att. apply and_dec; first eauto. + eapply least_dec. intros y. + eapply ext_pick_dec. + Qed. + + End Requirements. + + Section Requirements_Facts. + + (** ** Verification *) + + Lemma ext_pick_witness L n e: + ext_pick L n e -> exists x, least (ext_has_wit L n e) x. + Proof. + intros [H1 H2]. + eapply least_ex. intro; eapply ext_has_wit_dec. + trivial. + Qed. + + Lemma W_incl e n m: + n <= m -> forall x, (W[n] e) x -> (W[m] e) x. + Proof. + intros H x [y [H1 H2]]. + exists y. split; lia + easy. + Qed. + + Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : + L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. + Proof. + intros H h1 h2 x Hx1 Hx2. + eapply (H x). now eapply h1. + now apply h2. + Qed. + + Lemma act_impl_always_act e n: act e n -> forall m, n <= m -> act e m . + Proof. + intros H m Hm Hx. apply H. + eapply (intersect_mono Hx). + eapply F_mono; try eapply F_func_correctness; easy. + now eapply W_incl. + Qed. + + Lemma recv_impl_next_act e k: recv_att e k -> act e (S k). + Proof. + intros [He H] Hcontr. + edestruct (ext_pick_witness) as [x Hx]. + { destruct H. eapply e0. } + apply (Hcontr x). + assert (P_ (S k) (P_func (S k))) by apply F_func_correctness. + inv H0. cbn in H4. assert (ext_least_choice l k x) as Hwitness. + exists e. assert (P_func k = l) as <-. + eapply F_uni. apply F_func_correctness. apply H3. + split; first easy. split; first easy. easy. + assert (x = a) as ->. eapply (@extend_uni simple_extension); cbn; eauto. + eauto. exfalso. eapply (H3 x). cbn. + exists e. enough ((P_func k) = (P_func (S k))) as <-. + split; first easy. easy. + assert (F_ simple_extension k (P_func k)) by apply F_func_correctness. + eapply F_uni; eauto. + firstorder. + Qed. + + Lemma act_impl_not_recv e k: act e k -> ~ recv_att e k. + Proof. now intros h2 [_ [[h _] _]]. Qed. + + Lemma recv_impl_next_always_act e k: recv_att e k -> forall m, k < m -> act e m. + Proof. + intros. eapply act_impl_always_act. + apply recv_impl_next_act. apply H. lia. + Qed. + + Lemma recv_impl_always_not_recv e k: + recv_att e k -> forall m, k < m -> ~ recv_att e m. + Proof. + intros H1 m Hm. eapply act_impl_not_recv. + apply (recv_impl_next_always_act H1 Hm). + Qed. + + Lemma recv_at_most_once_gen e: + (pdec (exists k, recv_att e k)) -> + (exists s', forall s, s' < s -> ~ recv_att e s). + Proof. + intros [[w Hw]|H]. + - exists w. + now eapply recv_impl_always_not_recv. + - exists 0. + intros k' Hk' Ha. + apply H. now exists k'. + Qed. + + Lemma recv_at_most_once e: ~ ~ (exists s', forall s, s' < s -> ~ recv_att e s). + Proof. + intros H. + assert (~~ (pdec (exists k, recv_att e k))) as Hdec. + { unfold pdec. tauto. } + apply Hdec. clear Hdec. intros Hdec. + apply H, recv_at_most_once_gen. assumption. + Qed. + + Lemma recv_at_most_once_bound_gen k: + (forall k', k' < k -> pdec (∃ k0 : nat, recv_att k' k0)) -> + exists s, (forall e, e < k -> forall s', s < s' -> ~ recv_att e s'). + Proof. + intros Hle. + induction k. + - exists 42. intros ??. lia. + - destruct IHk as [s Hs]. + { intros. eapply Hle. lia. } + destruct (@recv_at_most_once_gen k) as [sk Hsk]. + { apply Hle. lia. } + set (max sk s) as N. + exists N. intros e He. + assert (e = k \/ e < k) as [->|Hek] by lia. + intros s' Hs'. eapply Hsk. lia. + intros s' Hs'. eapply Hs; lia. + Qed. + + Lemma recv_at_most_once_bound_classically k: + ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ recv_att e s'). + Proof. + eapply finite_decs. + intros H % (recv_at_most_once_bound_gen (k := k)). + tauto. + Qed. + Lemma recv_at_most_once_bound: + LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k -> ∀ s', s < s' -> ~ recv_att e s'). + Proof. + intros Hlem k. apply recv_at_most_once_bound_gen. + intros. eapply assume_Σ_1_lem. apply Hlem. eauto. + Qed. + + Lemma attend_uni e: unique (recv_att e). + Proof. + intros k1 k2 H1 H2. + specialize (fun a b => @act_impl_not_recv _ _ (@recv_impl_next_always_act _ _ H1 a b)) as H1'. + specialize (fun a b => @act_impl_not_recv _ _ (@recv_impl_next_always_act _ _ H2 a b)) as H2'. + enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. + intro Hk. eapply H1'. apply Hk. easy. + intro Hk. eapply H2'. apply Hk. easy. + Qed. + + Lemma pick_el_uni e: unique (act_by e). + Proof. + intros x1 x2 [k [Ht Hk]] [k' [Ht' Hk']]. + assert (k=k') as <-. eapply attend_uni; eauto. + eapply (@extend_uni simple_extension); cbn; eauto. + Qed. + + End Requirements_Facts. + + Section Complement_Inf. + + Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, act_by e x /\ e < n. + Proof. + intros [[L [k [Hin Hk]]] Hn]. + dependent induction L. inv Hin. + destruct (Dec (a = x)) as [->|]. + - clear IHL Hin. + destruct (F_pick Hk) as [k' [Hk' [e He]]]. + exists e. split. + exists k'. split; unfold recv_att. + + destruct He; intuition eauto. enough (P_func k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. easy. + + exists e. unfold ext_choice. + enough (P_func k' = L) as -> by eauto. + eapply F_uni. apply F_func_correctness. + destruct He; intuition eauto. + + destruct He. destruct H0. destruct H1. destruct H1. + lia. + - destruct (F_pop Hk) as [m' Hm']. + eapply (IHL m'); eauto. firstorder. + Qed. + + Lemma P_extract_spec n L: + (forall x, In x L -> P x /\ x <= 2 * n) -> + forall x, In x L -> exists c, act_by c x /\ c < n. + Proof. + intros. induction L. inv H0. + destruct H0 as [-> | Hln]; last apply IHL; eauto. + apply P_meet_spec. now apply H. + Qed. + + Lemma P_pullback_list n L: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> + exists (LC: list nat), NoDup LC /\ length LC = length L /\ + forall c, In c LC -> exists x, act_by c x /\ In x L /\ c < n. + Proof. + intros HL H1. + induction L. + - exists []; intuition. + - remember (@P_extract_spec n (a::L) H1 a). + assert (In a (a::L)) as H0 by intuition. + apply e in H0 as [c0 [H0 E1]]. + assert (NoDup L) by (inversion HL; intuition). + apply IHL in H as [LC H]. + exists (c0::LC). intuition. + + constructor; last now exact H2. + intros In. inv HL. + apply H4 in In as [y (Hs & h2 & h3)]. + now rewrite (pick_el_uni H0 Hs) in H6. + + cbn. rewrite H. trivial. + + destruct H3 as [->|]. + * exists a; intuition. + * destruct (H4 c H3) as [y Hy]. + exists y; intuition. + + intros y In1. assert (In y (a::L)) by intuition. + now apply H1 in H2. + Qed. + + Definition PredListTo p : list nat -> nat -> Prop + := fun L b => forall x, In x L <-> p x /\ x <= b. + + Lemma NoDupBoundH {L} b: + NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). + Proof. + intros ND H x E. + constructor. + - intros H1 % H. lia. + - exact ND. + Qed. + + Lemma PredNoDupListTo_NNExist p: + forall b, ~~ exists L, PredListTo p L b /\ NoDup L. + Proof. + destruct (F_computable simple_extension) as [f Hf]. + induction b; intros H. + - ccase (p 0) as [H0 | H0]; apply H. + + exists [0]. split; try split. + * intros [E | E]; (try contradiction E). + rewrite <- E. intuition. + * intros E. assert (x = 0) by lia. + rewrite H1. intuition. + * constructor; intuition; constructor. + + exists nil. split; try split. + * contradiction. + * intros E. assert (x = 0) by lia. + rewrite H1 in E. firstorder. + * constructor. + - apply IHb. intros [L H1]. + ccase (p (1 + b)) as [H0 | H0]; apply H. + + exists ((1+ b) :: L). split; try split. + * intros [E | E]; try (rewrite <- E; intuition). + apply H1 in E. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** right. apply H1. intuition. + ** left. lia. + * apply (@NoDupBoundH _ b). + ** apply H1. + ** intros x H3 % H1. lia. + ** lia. + + exists L. split; try split. + * intros E % H1. intuition. + * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + ** apply H1. intuition. + ** rewrite E in E1. firstorder. + * apply H1. + Qed. + + Lemma P_bounded L n: + NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. + Proof. + intros ND [LC H] % P_pullback_list; intuition. + rewrite <- H. + assert (incl LC (seq 0 n)). unfold incl. + - intros c [e [x Hx]] % H2. apply in_seq. lia. + - apply pigeonhole_length in H1. + + now rewrite seq_length in H1. + + intros. decide (x1 = x2); tauto. + + exact H0. + Qed. + + Lemma P_Listing: + forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). + Proof. + intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). + intros [L H1]. apply H. exists L; intuition. + apply P_bounded. + - exact H2. + - apply H0. + Qed. + + Lemma complToBound_compl p L b: + PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. + Proof. + intros H x. split. + - intros [H1 H1'] % in_filter_iff. + destruct Dec; cbn in H1'; try congruence. + enough (x <= b). + + firstorder. + + apply in_seq in H1. lia. + - intros [H1 H2]. eapply in_filter_iff. split. + + apply in_seq; lia. + + destruct Dec; cbn; try tauto. exfalso. firstorder. + Qed. + + Lemma compl_P_Listing: + forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L + /\ forall x, In x L -> ~ P x. + Proof. + intros n H. + apply (@P_Listing n). intros [L H1]. + apply H. exists (complToBound L (2*n)). repeat split. + - remember (complToBound_length L (2*n)). lia. + - apply complToBound_NoDup. + - intros x I % (@complToBound_compl P); intuition. + Qed. + + Lemma P_coinfinite : ~ exhaustible (compl P). + Proof. + eapply weakly_unbounded_non_finite. + intros n H. eapply compl_P_Listing with (n := n). + intros (l & ? & ? & H2). + eapply H. + exists (firstn n l). + repeat split. + - rewrite firstn_length. lia. + - now eapply firstn_NoDup. + - intros ? ? % firstn_In. now eapply H2. + Qed. + + End Complement_Inf. + + Section Requirements_Meet. + + Lemma wall_bounded' e: ~~ exists w, forall x, wall e (P_func x) x < w. + Proof. + intro H_. apply (@wall_spec e). intros [v [k Hk]]. + apply H_. + destruct (@extra_bounded (fun k => wall e (P_func k) k) k) as [w Hw]. + exists (S (max v w)). intros x. + destruct (Dec (x < k)). apply Hw in l. lia. + assert (wall e (P_func x) x = v). apply Hk. lia. lia. + Qed. + + Lemma wall_bounded e: + ~~ exists w, forall i x, i <= e -> wall i (P_func x) x < w. + Proof. + intro H_. induction e. + - apply (@wall_bounded' 0). intros [w Hw]. apply H_. exists w. + intros i x ?. inv H. trivial. + - apply IHe. intros [w IHe_]. + apply (@wall_bounded' (S e)). intros [w' Hw']. + apply H_. exists (S (max w w')). intros i x H. + inv H. specialize (Hw' x). lia. + specialize (IHe_ i x H1). lia. + Qed. + + Lemma non_finite_not_bounded e: + non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ + (forall n, forall i, i <= e -> wall i (P_func n) n < x). + Proof. + intro H. unfold non_finite in H. + intros He. rewrite non_finite_nat in H. + apply (@wall_bounded e). intros [w Hw]. + pose (N := max (2*e + 1) w). specialize (H N). + apply H. intros [m [Hm1 [k Hmk]]]. + apply He. exists k, m. + repeat split. now exists k. + lia. intros n i Hie. specialize (Hw i n Hie). lia. + Qed. + + Lemma ext_pick_impl_recv N e: + e < N -> ext_pick (P_func N) N e -> + (exists w, w ≤ e /\ recv_att w N). + Proof. + intros HeN H1. + assert (exists w, ext_pick (P_func N) N w) by now exists e. + eapply least_ex in H; last eauto. + destruct H as [k Hk]. assert (k <= e). + { enough (~ k > e) by lia. intro Hkw. + destruct Hk. rewrite safe_char in H0. + specialize (H0 e H1). lia. } + exists k. do 2 (split; first lia). eapply Hk. + Qed. + + Lemma non_finite_attend e: + non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ recv_att e k) . + Proof. + intros H He. + eapply (non_finite_not_bounded H); intros (b & x & (Hx1 & Hx2 & Hx3)). + eapply (@recv_at_most_once_bound_classically e); intros [s Hs]. + pose (N := S (max (max b s) e)). apply He. + destruct (Dec (ext_intersect_W (P_func N) N e)) as [He'|He']. + - exists N. right. + assert (ext_pick (P_func N) N e). + { split; first easy. exists x. split. + eapply W_incl; last apply Hx1. lia. + split; first lia. eapply Hx3. + } + split. lia. split. easy. + intros w HEw Hw. + eapply ext_pick_impl_recv in Hw. + destruct Hw as [g [HEg Hg]]. + eapply Hs; last apply Hg; lia. lia. + - exists N. now left. + Qed. + + Lemma ext_intersect_W_intersect k e: + ~ (P_func k # W[k] e) -> W e #ₚ P -> False. + Proof. + unfold ext_intersect_W. + intros H1 H2. apply H1. + intros y Hy1 [w Hy2]. + eapply (H2 y). now exists w. + eapply (In_Pf Hy1). + Qed. + + Lemma P_requirements_meet : forall e, P_requirements P e. + Proof. + intros e He. intros He'. + eapply (non_finite_attend He). + intros [k [H|H]]. + - eapply ext_intersect_W_intersect; eauto. + - eapply recv_impl_next_act in H. + eapply ext_intersect_W_intersect; eauto. + Qed. + + End Requirements_Meet. + + Section Results. + + Lemma P_semi_decidable : semi_decidable P. + Proof. apply F_with_semi_decidable. Qed. + + Theorem P_simple : simple P. + Proof. + unfold simple; repeat split. + - rewrite EA.enum_iff. now apply P_semi_decidable. + - apply P_coinfinite. + - intros [q (Hqenum & Hqinf & Hq)]. + rewrite EA.enum_iff in Hqenum. + destruct (W_spec Hqenum) as [c HWq]. + apply (@P_requirements_meet c). + intros [l Hqe]; apply Hqinf. + exists l. intros x Hqx. apply (Hqe x). + now rewrite HWq in Hqx. + intros x HWcx HPx. unfold W in HWcx. + rewrite <- (HWq x) in HWcx. apply (Hq x HWcx HPx). + Qed. + + End Results. End Assume_WALL. End Assume_EA. + From affbbf76643fcbc8634e93fa39e41880df21a780 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Tue, 21 Jan 2025 22:13:11 +0100 Subject: [PATCH 48/54] polisj --- theories/PostsProblem/limit_computability.v | 1 - theories/PostsProblem/step_indexing.v | 15 --------------- 2 files changed, 16 deletions(-) diff --git a/theories/PostsProblem/limit_computability.v b/theories/PostsProblem/limit_computability.v index 0487cdd..42c89ba 100644 --- a/theories/PostsProblem/limit_computability.v +++ b/theories/PostsProblem/limit_computability.v @@ -23,7 +23,6 @@ Convention: **) - (* Definition of limit ciomputable *) Definition limit_computable {X} (P: X -> Prop) := diff --git a/theories/PostsProblem/step_indexing.v b/theories/PostsProblem/step_indexing.v index dac92cc..2a90ed7 100644 --- a/theories/PostsProblem/step_indexing.v +++ b/theories/PostsProblem/step_indexing.v @@ -359,21 +359,6 @@ Section Step_Eval. by exists depth; injection H as ->. Qed. - (* Lemma evalt_comp_step (τ: tree) f n m v qs ans q: - interrogation τ (λ x y, f x = y) qs ans → - length qs = n → - τ ans =! Ask q → - evalt_comp τ f n m = Some (Ask q) ∧ seval (τ (f q::ans)) m = Some v → - evalt_comp τ f (S n) m = Some v. - Proof. - intros HIn HE Hans [H1 H2]. - rewrite <- HE in *. - replace (length qs) with (length qs + 0) in H1 by lia. - eapply interrogation_plus_evalt_comp_1 in H1; last done. - induction HIn. - - simpl in *. admit. - - admit. - Admitted. *) Fact sub_tree (τ: tree) f use ans ans_: interrogation τ (λ x y, f x = y) use ans → ans_ `prefix_of` ans → From b3b0c1c0b1addecda6f3bb27b1eaa91cfc96e42e Mon Sep 17 00:00:00 2001 From: Haoyi Date: Wed, 22 Jan 2025 13:08:28 +0100 Subject: [PATCH 49/54] typo --- theories/PostsProblem/simpleness.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index a25c246..083b254 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -655,7 +655,7 @@ Section Assume_EA. exists k. do 2 (split; first lia). eapply Hk. Qed. - Lemma non_finite_attend e: + Lemma non_finite_recv e: non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ recv_att e k) . Proof. intros H He. @@ -690,7 +690,7 @@ Section Assume_EA. Lemma P_requirements_meet : forall e, P_requirements P e. Proof. intros e He. intros He'. - eapply (non_finite_attend He). + eapply (non_finite_recv He). intros [k [H|H]]. - eapply ext_intersect_W_intersect; eauto. - eapply recv_impl_next_act in H. From d67c4672cc3ea18d2c220f02c51b5602c80877f8 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Wed, 22 Jan 2025 15:29:22 +0100 Subject: [PATCH 50/54] using unicode --- theories/PostsProblem/lowness.v | 2 +- theories/PostsProblem/simpleness.v | 165 +++++++++++++++-------------- 2 files changed, 84 insertions(+), 83 deletions(-) diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index 5674e02..2d1cef2 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -86,7 +86,7 @@ Section Requirements_Meet. Hypothesis Σ_1_lem: LEM_Σ 1. Lemma attention_bound: - ∀ k, ∃ s, ∀ e, e < k -> ∀ s', s < s' -> ~ recv_att wall e s'. + ∀ k, ∃ s, ∀ e, e < k → ∀ s', s < s' → ~ recv_att wall e s'. Proof. by apply recv_at_most_once_bound. Qed. Lemma eventally_greater_than_wall: diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index 083b254..b5513d6 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -14,10 +14,10 @@ Import ListNotations. Section ComplToBound. Definition complToBound L b : list nat - := filter (fun x => Dec (~ In x L)) (seq 0 (S b)). + := filter (fun x => Dec (¬ In x L)) (seq 0 (S b)). Lemma complToBound_Bound L b : - forall x, In x (complToBound L b) -> x <= b. + ∀ x, In x (complToBound L b) → x <= b. Proof. intros x [H % in_seq ?] % in_filter_iff. lia. Qed. @@ -29,7 +29,7 @@ Section ComplToBound. - destruct f; cbn; lia. Qed. Lemma filter_NoDup {X} f (l : list X) : - NoDup l -> NoDup (filter f l). + NoDup l → NoDup (filter f l). Proof. induction 1; cbn. - econstructor. @@ -58,12 +58,12 @@ Section ComplToBound. Proof. eapply filter_NoDup, seq_NoDup. Qed. - Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) -> In x l. + Lemma firstn_In {X} (l : list X) n x : In x (firstn n l) → In x l. Proof. induction n in x, l |- *; destruct l; cbn; firstorder. Qed. - Lemma firstn_NoDup {X} (l : list X) n : NoDup l -> NoDup (firstn n l). + Lemma firstn_NoDup {X} (l : list X) n : NoDup l → NoDup (firstn n l). Proof. induction 1 in n |- *; destruct n; cbn; try now econstructor. econstructor; eauto. @@ -85,16 +85,16 @@ Section Assume_EA. Qed. Definition W_ n e x := φ n e = Some x. - Definition W e x := exists n, W_ e n x. + Definition W e x := ∃ n, W_ e n x. - Lemma W_spec: forall P, semi_decidable P -> exists e, forall x, P x <-> W e x. + Lemma W_spec: ∀ P, semi_decidable P → ∃ e, ∀ x, P x <-> W e x. Proof. intros P [e He]%EA. exists e; intros x; now rewrite He. Qed. - Notation "'W[' s ']' e" := (fun x => exists n, n <= s /\ W_ e n x) (at level 30). + Notation "'W[' s ']' e" := (fun x => ∃ n, n <= s ∧ W_ e n x) (at level 30). Section EA_dec. - Lemma W_dec e: forall x n, dec (W_ e n x). + Lemma W_dec e: ∀ x n, dec (W_ e n x). Proof. intros x n. destruct (φ e n) eqn: E. @@ -103,7 +103,7 @@ Section Assume_EA. - right. intros H. congruence. Qed. - Lemma W_bounded_dec e : forall x s, dec ((W[s] e) x). + Lemma W_bounded_dec e : ∀ x s, dec ((W[s] e) x). Proof. intros x s. cbn. eapply exists_bounded_dec. intro; apply W_dec. @@ -134,12 +134,12 @@ Section Assume_EA. End EA_dec. - Definition disj_list_pred {X} (A: list X) (B: X -> Prop) := forall x, In x A -> B x -> False. - Definition disj_pred {X} (A B: X -> Prop) := forall x, A x -> B x -> False. + Definition disj_list_pred {X} (A: list X) (B: X → Prop) := ∀ x, In x A → B x → False. + Definition disj_pred {X} (A B: X → Prop) := ∀ x, A x → B x → False. Notation "A # B" := (disj_list_pred A B) (at level 30). Notation "A #ₚ B" := (disj_pred A B) (at level 30). - Lemma extra_bounded f m: Σ b, forall n, n < m -> f n < b. + Lemma extra_bounded f m: Σ b, ∀ n, n < m → f n < b. Proof. induction m. - exists 42. intros. inv H. @@ -155,10 +155,10 @@ Section Assume_EA. Section Extension. (** ** Construction *) Definition ext_intersect_W L n e := L # W[n] e. - Definition ext_has_wit L n e x := (W[n] e) x /\ 2 * e < x /\ forall i, i <= e -> wall i L n < x. - Definition ext_pick L n e := ext_intersect_W L n e /\ exists x, ext_has_wit L n e x. - Definition ext_choice L n e x := e < n /\ least (ext_pick L n) e /\ least (ext_has_wit L n e) x. - Definition ext_least_choice L n x := exists e, ext_choice L n e x. + Definition ext_has_wit L n e x := (W[n] e) x ∧ 2 * e < x ∧ ∀ i, i <= e → wall i L n < x. + Definition ext_pick L n e := ext_intersect_W L n e ∧ ∃ x, ext_has_wit L n e x. + Definition ext_choice L n e x := e < n ∧ least (ext_pick L n) e ∧ least (ext_has_wit L n e) x. + Definition ext_least_choice L n x := ∃ e, ext_choice L n e x. End Extension. @@ -170,13 +170,13 @@ Section Assume_EA. intro x. eapply dec_neg_dec, exists_bounded_dec; eauto. Qed. - #[export] Instance ext_wall L n e x: dec (forall i, i <= e -> wall i L n < x). + #[export] Instance ext_wall L n e x: dec (∀ i, i <= e → wall i L n < x). Proof. eapply forall_bounded_dec; eauto. Qed. #[export]Instance ext_has_wit_dec L n e x : dec (ext_has_wit L n e x). Proof. apply and_dec; first apply exists_bounded_dec; eauto. Qed. - #[export]Instance ext_has_wit_exists_dec L n e : dec (exists x, ext_has_wit L n e x). + #[export]Instance ext_has_wit_exists_dec L n e : dec (∃ x, ext_has_wit L n e x). Proof. unfold ext_has_wit. eapply bounded_dec; last eapply W_bounded_bounded. intro x; eapply W_bounded_dec; eauto. eauto. @@ -191,7 +191,7 @@ Section Assume_EA. intro; apply W_dec. eauto. Qed. - #[export]Instance ext_pick_exists_dec L n: dec (exists e, e < n /\ least (ext_pick L n) e). + #[export]Instance ext_pick_exists_dec L n: dec (∃ e, e < n ∧ least (ext_pick L n) e). Proof. eapply exists_bounded_dec'. intro x. eapply least_dec. intros y. @@ -199,7 +199,7 @@ Section Assume_EA. Qed. Fact ext_least_choice_dec L n: - (Σ x : nat, ext_least_choice L n x) + (forall x : nat, ~ ext_least_choice L n x). + (Σ x : nat, ext_least_choice L n x) + (∀ x : nat, ¬ ext_least_choice L n x). Proof. unfold ext_least_choice. destruct (ext_pick_exists_dec L n) as [H|H]. @@ -238,44 +238,44 @@ Section Assume_EA. Definition P_func := F_func simple_extension. Definition P := F_with simple_extension. - Definition non_finite e := ~ exhaustible (W e). + Definition non_finite e := ¬ exhaustible (W e). - Fact In_Pf k y: In y (P_func k) -> P y. + Fact In_Pf k y: In y (P_func k) → P y. Proof. intros H. exists (P_func k), k. split; [easy| apply F_func_correctness]. Qed. - Definition lim_to (f: list nat -> nat -> nat) b := - (exists n, forall m, n <= m -> f (P_func m) m = b). + Definition lim_to (f: list nat → nat → nat) b := + (∃ n, ∀ m, n <= m → f (P_func m) m = b). - Lemma finite_decs (P : nat -> Prop) k Q : - ((forall k', k' < k -> pdec (P k')) -> ~ Q) -> ~ Q. + Lemma finite_decs (P : nat → Prop) k Q : + ((∀ k', k' < k → pdec (P k')) → ¬ Q) → ¬ Q. Proof. induction k. - firstorder lia. - intros H Hq. - assert (~~ pdec (P k)) as Hk. { cbv. tauto. } + assert (¬¬ pdec (P k)) as Hk. { cbv. tauto. } apply Hk. clear Hk. intros Hk. apply IHk. 2: assumption. intros Ha. apply H. intros. - assert (k' = k \/ k' < k) as [-> | ] by lia. + assert (k' = k ∨ k' < k) as [-> | ] by lia. assumption. now apply Ha. Qed. End Extension_Facts. - Hypothesis wall_spec: forall e, ~~ exists b, lim_to (wall e) b. + Hypothesis wall_spec: ∀ e, ¬¬ ∃ b, lim_to (wall e) b. Section Requirements. (** ** Requirements *) - Definition P_requirements P e := non_finite e -> ~ (W e #ₚ P). - Definition recv_att e n := e < n /\ least (ext_pick (P_func n) n) e. - Definition act e n := ~ (P_func n) # W[n] e. - Definition act_by e x := exists k, recv_att e k /\ ext_least_choice (P_func k) k x. + Definition P_requirements P e := non_finite e → ¬ (W e #ₚ P). + Definition recv_att e n := e < n ∧ least (ext_pick (P_func n) n) e. + Definition act e n := ¬ (P_func n) # W[n] e. + Definition act_by e x := ∃ k, recv_att e k ∧ ext_least_choice (P_func k) k x. Definition done e n := ∀ s, n < s → ¬ recv_att e s. #[export]Instance attend_dec e n: dec (recv_att e n). @@ -292,7 +292,7 @@ Section Assume_EA. (** ** Verification *) Lemma ext_pick_witness L n e: - ext_pick L n e -> exists x, least (ext_has_wit L n e) x. + ext_pick L n e → ∃ x, least (ext_has_wit L n e) x. Proof. intros [H1 H2]. eapply least_ex. intro; eapply ext_has_wit_dec. @@ -300,21 +300,21 @@ Section Assume_EA. Qed. Lemma W_incl e n m: - n <= m -> forall x, (W[n] e) x -> (W[m] e) x. + n <= m → ∀ x, (W[n] e) x → (W[m] e) x. Proof. intros H x [y [H1 H2]]. exists y. split; lia + easy. Qed. - Lemma intersect_mono {X} (L L': list X) (P P': X -> Prop) : - L' # P' -> incl L L' -> (forall x, P x -> P' x) -> L # P. + Lemma intersect_mono {X} (L L': list X) (P P': X → Prop) : + L' # P' → incl L L' → (∀ x, P x → P' x) → L # P. Proof. intros H h1 h2 x Hx1 Hx2. eapply (H x). now eapply h1. now apply h2. Qed. - Lemma act_impl_always_act e n: act e n -> forall m, n <= m -> act e m . + Lemma act_impl_always_act e n: act e n → ∀ m, n <= m → act e m . Proof. intros H m Hm Hx. apply H. eapply (intersect_mono Hx). @@ -322,7 +322,7 @@ Section Assume_EA. now eapply W_incl. Qed. - Lemma recv_impl_next_act e k: recv_att e k -> act e (S k). + Lemma recv_impl_next_act e k: recv_att e k → act e (S k). Proof. intros [He H] Hcontr. edestruct (ext_pick_witness) as [x Hx]. @@ -342,25 +342,25 @@ Section Assume_EA. firstorder. Qed. - Lemma act_impl_not_recv e k: act e k -> ~ recv_att e k. + Lemma act_impl_not_recv e k: act e k → ¬ recv_att e k. Proof. now intros h2 [_ [[h _] _]]. Qed. - Lemma recv_impl_next_always_act e k: recv_att e k -> forall m, k < m -> act e m. + Lemma recv_impl_next_always_act e k: recv_att e k → ∀ m, k < m → act e m. Proof. intros. eapply act_impl_always_act. apply recv_impl_next_act. apply H. lia. Qed. Lemma recv_impl_always_not_recv e k: - recv_att e k -> forall m, k < m -> ~ recv_att e m. + recv_att e k → ∀ m, k < m → ¬ recv_att e m. Proof. intros H1 m Hm. eapply act_impl_not_recv. apply (recv_impl_next_always_act H1 Hm). Qed. Lemma recv_at_most_once_gen e: - (pdec (exists k, recv_att e k)) -> - (exists s', forall s, s' < s -> ~ recv_att e s). + (pdec (∃ k, recv_att e k)) -> + (∃ s', ∀ s, s' < s → ¬ recv_att e s). Proof. intros [[w Hw]|H]. - exists w. @@ -370,18 +370,18 @@ Section Assume_EA. apply H. now exists k'. Qed. - Lemma recv_at_most_once e: ~ ~ (exists s', forall s, s' < s -> ~ recv_att e s). + Lemma recv_at_most_once e: ¬ ¬ (∃ s', ∀ s, s' < s → ¬ recv_att e s). Proof. intros H. - assert (~~ (pdec (exists k, recv_att e k))) as Hdec. + assert (¬¬ (pdec (∃ k, recv_att e k))) as Hdec. { unfold pdec. tauto. } apply Hdec. clear Hdec. intros Hdec. apply H, recv_at_most_once_gen. assumption. Qed. Lemma recv_at_most_once_bound_gen k: - (forall k', k' < k -> pdec (∃ k0 : nat, recv_att k' k0)) -> - exists s, (forall e, e < k -> forall s', s < s' -> ~ recv_att e s'). + (∀ k', k' < k → pdec (∃ k0 : nat, recv_att k' k0)) -> + ∃ s, (∀ e, e < k → ∀ s', s < s' → ¬ recv_att e s'). Proof. intros Hle. induction k. @@ -392,20 +392,20 @@ Section Assume_EA. { apply Hle. lia. } set (max sk s) as N. exists N. intros e He. - assert (e = k \/ e < k) as [->|Hek] by lia. + assert (e = k ∨ e < k) as [->|Hek] by lia. intros s' Hs'. eapply Hsk. lia. intros s' Hs'. eapply Hs; lia. Qed. Lemma recv_at_most_once_bound_classically k: - ~ ~ exists s, (forall e, e < k -> forall s', s < s' -> ~ recv_att e s'). + ¬ ¬ ∃ s, (∀ e, e < k → ∀ s', s < s' → ¬ recv_att e s'). Proof. eapply finite_decs. intros H % (recv_at_most_once_bound_gen (k := k)). tauto. Qed. Lemma recv_at_most_once_bound: - LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k -> ∀ s', s < s' -> ~ recv_att e s'). + LEM_Σ 1 → ∀ k, ∃ s, (∀ e, e < k → ∀ s', s < s' → ¬ recv_att e s'). Proof. intros Hlem k. apply recv_at_most_once_bound_gen. intros. eapply assume_Σ_1_lem. apply Hlem. eauto. @@ -416,7 +416,7 @@ Section Assume_EA. intros k1 k2 H1 H2. specialize (fun a b => @act_impl_not_recv _ _ (@recv_impl_next_always_act _ _ H1 a b)) as H1'. specialize (fun a b => @act_impl_not_recv _ _ (@recv_impl_next_always_act _ _ H2 a b)) as H2'. - enough (~ k1 < k2 /\ ~ k2 < k1) by lia; split. + enough (¬ k1 < k2 ∧ ¬ k2 < k1) by lia; split. intro Hk. eapply H1'. apply Hk. easy. intro Hk. eapply H2'. apply Hk. easy. Qed. @@ -432,7 +432,7 @@ Section Assume_EA. Section Complement_Inf. - Lemma P_meet_spec x n : P x /\ x <= 2*n -> exists e, act_by e x /\ e < n. + Lemma P_meet_spec x n : P x ∧ x <= 2*n → ∃ e, act_by e x ∧ e < n. Proof. intros [[L [k [Hin Hk]]] Hn]. dependent induction L. inv Hin. @@ -454,8 +454,8 @@ Section Assume_EA. Qed. Lemma P_extract_spec n L: - (forall x, In x L -> P x /\ x <= 2 * n) -> - forall x, In x L -> exists c, act_by c x /\ c < n. + (∀ x, In x L → P x ∧ x <= 2 * n) → + ∀ x, In x L → ∃ c, act_by c x ∧ c < n. Proof. intros. induction L. inv H0. destruct H0 as [-> | Hln]; last apply IHL; eauto. @@ -463,9 +463,9 @@ Section Assume_EA. Qed. Lemma P_pullback_list n L: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> - exists (LC: list nat), NoDup LC /\ length LC = length L /\ - forall c, In c LC -> exists x, act_by c x /\ In x L /\ c < n. + NoDup L → (∀ x, In x L → P x ∧ x <= 2 * n) → + ∃ (LC: list nat), NoDup LC ∧ length LC = length L ∧ + ∀ c, In c LC → ∃ x, act_by c x ∧ In x L ∧ c < n. Proof. intros HL H1. induction L. @@ -489,11 +489,11 @@ Section Assume_EA. now apply H1 in H2. Qed. - Definition PredListTo p : list nat -> nat -> Prop - := fun L b => forall x, In x L <-> p x /\ x <= b. + Definition PredListTo p : list nat → nat → Prop + := fun L b => ∀ x, In x L <-> p x ∧ x <= b. Lemma NoDupBoundH {L} b: - NoDup L -> (forall x, In x L -> x <= b) -> forall x, x > b -> NoDup (x::L). + NoDup L → (∀ x, In x L → x <= b) → ∀ x, x > b → NoDup (x::L). Proof. intros ND H x E. constructor. @@ -502,7 +502,7 @@ Section Assume_EA. Qed. Lemma PredNoDupListTo_NNExist p: - forall b, ~~ exists L, PredListTo p L b /\ NoDup L. + ∀ b, ¬¬ ∃ L, PredListTo p L b ∧ NoDup L. Proof. destruct (F_computable simple_extension) as [f Hf]. induction b; intros H. @@ -523,7 +523,7 @@ Section Assume_EA. + exists ((1+ b) :: L). split; try split. * intros [E | E]; try (rewrite <- E; intuition). apply H1 in E. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + * intros [E1 E2]. assert (x <= b ∨ x = 1 + b) as [E | E] by lia. ** right. apply H1. intuition. ** left. lia. * apply (@NoDupBoundH _ b). @@ -532,14 +532,14 @@ Section Assume_EA. ** lia. + exists L. split; try split. * intros E % H1. intuition. - * intros [E1 E2]. assert (x <= b \/ x = 1 + b) as [E | E] by lia. + * intros [E1 E2]. assert (x <= b ∨ x = 1 + b) as [E | E] by lia. ** apply H1. intuition. ** rewrite E in E1. firstorder. * apply H1. Qed. Lemma P_bounded L n: - NoDup L -> (forall x, In x L -> P x /\ x <= 2 * n) -> length L <= n. + NoDup L → (∀ x, In x L → P x ∧ x <= 2 * n) → length L <= n. Proof. intros ND [LC H] % P_pullback_list; intuition. rewrite <- H. @@ -552,7 +552,7 @@ Section Assume_EA. Qed. Lemma P_Listing: - forall n, ~~ exists L, NoDup L /\ length L <= n /\ PredListTo P L (2*n). + ∀ n, ¬¬ ∃ L, NoDup L ∧ length L <= n ∧ PredListTo P L (2*n). Proof. intros n H. apply (@PredNoDupListTo_NNExist P (2*n)). intros [L H1]. apply H. exists L; intuition. @@ -562,7 +562,7 @@ Section Assume_EA. Qed. Lemma complToBound_compl p L b: - PredListTo p L b -> PredListTo (compl p) (complToBound L b) b. + PredListTo p L b → PredListTo (compl p) (complToBound L b) b. Proof. intros H x. split. - intros [H1 H1'] % in_filter_iff. @@ -576,8 +576,8 @@ Section Assume_EA. Qed. Lemma compl_P_Listing: - forall (n: nat) , ~~ exists L, length L >= n /\ NoDup L - /\ forall x, In x L -> ~ P x. + ∀ (n: nat) , ¬¬ ∃ L, length L >= n ∧ NoDup L + ∧ ∀ x, In x L → ¬ P x. Proof. intros n H. apply (@P_Listing n). intros [L H1]. @@ -587,7 +587,7 @@ Section Assume_EA. - intros x I % (@complToBound_compl P); intuition. Qed. - Lemma P_coinfinite : ~ exhaustible (compl P). + Lemma P_coinfinite : ¬ exhaustible (compl P). Proof. eapply weakly_unbounded_non_finite. intros n H. eapply compl_P_Listing with (n := n). @@ -604,7 +604,7 @@ Section Assume_EA. Section Requirements_Meet. - Lemma wall_bounded' e: ~~ exists w, forall x, wall e (P_func x) x < w. + Lemma wall_bounded' e: ¬¬ ∃ w, ∀ x, wall e (P_func x) x < w. Proof. intro H_. apply (@wall_spec e). intros [v [k Hk]]. apply H_. @@ -615,7 +615,7 @@ Section Assume_EA. Qed. Lemma wall_bounded e: - ~~ exists w, forall i x, i <= e -> wall i (P_func x) x < w. + ¬¬ ∃ w, ∀ i x, i <= e → wall i (P_func x) x < w. Proof. intro H_. induction e. - apply (@wall_bounded' 0). intros [w Hw]. apply H_. exists w. @@ -628,8 +628,8 @@ Section Assume_EA. Qed. Lemma non_finite_not_bounded e: - non_finite e -> ~~ exists k, exists x, (W[k] e) x /\ 2 * e < x /\ - (forall n, forall i, i <= e -> wall i (P_func n) n < x). + non_finite e → ¬¬ ∃ k, ∃ x, (W[k] e) x ∧ 2 * e < x ∧ + (∀ n, ∀ i, i <= e → wall i (P_func n) n < x). Proof. intro H. unfold non_finite in H. intros He. rewrite non_finite_nat in H. @@ -642,21 +642,21 @@ Section Assume_EA. Qed. Lemma ext_pick_impl_recv N e: - e < N -> ext_pick (P_func N) N e -> - (exists w, w ≤ e /\ recv_att w N). + e < N → ext_pick (P_func N) N e → + (∃ w, w ≤ e ∧ recv_att w N). Proof. intros HeN H1. assert (exists w, ext_pick (P_func N) N w) by now exists e. eapply least_ex in H; last eauto. destruct H as [k Hk]. assert (k <= e). - { enough (~ k > e) by lia. intro Hkw. + { enough (¬ k > e) by lia. intro Hkw. destruct Hk. rewrite safe_char in H0. specialize (H0 e H1). lia. } exists k. do 2 (split; first lia). eapply Hk. Qed. Lemma non_finite_recv e: - non_finite e -> ~ ~ (exists k, ~ ext_intersect_W (P_func k) k e \/ recv_att e k) . + non_finite e → ¬ ¬ (∃ k, ¬ ext_intersect_W (P_func k) k e ∨ recv_att e k) . Proof. intros H He. eapply (non_finite_not_bounded H); intros (b & x & (Hx1 & Hx2 & Hx3)). @@ -678,7 +678,7 @@ Section Assume_EA. Qed. Lemma ext_intersect_W_intersect k e: - ~ (P_func k # W[k] e) -> W e #ₚ P -> False. + ¬ (P_func k # W[k] e) → W e #ₚ P → False. Proof. unfold ext_intersect_W. intros H1 H2. apply H1. @@ -687,9 +687,9 @@ Section Assume_EA. eapply (In_Pf Hy1). Qed. - Lemma P_requirements_meet : forall e, P_requirements P e. + Lemma P_requirements_meet : ∀ e, P_requirements P e. Proof. - intros e He. intros He'. + intros e He He'. eapply (non_finite_recv He). intros [k [H|H]]. - eapply ext_intersect_W_intersect; eauto. @@ -728,3 +728,4 @@ End Assume_EA. + From 11614eea41da1c35b23e1a299f2c26db9ac6b23e Mon Sep 17 00:00:00 2001 From: Haoyi Date: Wed, 22 Jan 2025 15:48:00 +0100 Subject: [PATCH 51/54] unicode --- theories/PostsProblem/limit_computability.v | 164 ++++++++++---------- theories/PostsProblem/simpleness.v | 12 +- 2 files changed, 88 insertions(+), 88 deletions(-) diff --git a/theories/PostsProblem/limit_computability.v b/theories/PostsProblem/limit_computability.v index 42c89ba..7b8313c 100644 --- a/theories/PostsProblem/limit_computability.v +++ b/theories/PostsProblem/limit_computability.v @@ -1,6 +1,6 @@ From SyntheticComputability Require Import ArithmeticalHierarchySemantic PostsTheorem reductions SemiDec TuringJump OracleComputability Definitions partial Pigeonhole. -Require Import Vectors.VectorDef Arith.Compare_dec Lia. +Require Import stdpp.list Vectors.VectorDef Arith.Compare_dec Lia. Import Vector.VectorNotations. Local Notation vec := Vector.t. @@ -25,27 +25,27 @@ Convention: (* Definition of limit ciomputable *) - Definition limit_computable {X} (P: X -> Prop) := - exists f: X -> nat -> bool, forall x, - (P x <-> exists N, forall n, n >= N -> f x n = true) /\ - (~ P x <-> exists N, forall n, n >= N -> f x n = false). + Definition limit_computable {X} (P: X → Prop) := + ∃ f: X → nat → bool, ∀ x, + (P x ↔ ∃ N, ∀ n, n ≥ N → f x n = true) ∧ + (¬ P x ↔ ∃ N, ∀ n, n ≥ N → f x n = false). - Definition char_rel_limit_computable {X} (P: X -> bool -> Prop) := - exists f: X -> nat -> bool, forall x y, P x y <-> exists N, forall n, n >= N -> f x n = y. + Definition char_rel_limit_computable {X} (P: X → bool → Prop) := + ∃ f: X → nat → bool, ∀ x y, P x y ↔ ∃ N, ∀ n, n ≥ N → f x n = y. - Definition char_rel_limit_computable' {X} (P: X -> bool -> Prop) := - exists f: X -> nat -> bool, forall x y, P x y -> exists N, forall n, n >= N -> f x n = y. + Definition char_rel_limit_computable' {X} (P: X → bool → Prop) := + ∃ f: X → nat → bool, ∀ x y, P x y → ∃ N, ∀ n, n ≥ N → f x n = y. - Lemma char_rel_limit_equiv {X} (P: X -> Prop): - char_rel_limit_computable (char_rel P) <-> limit_computable P. + Lemma char_rel_limit_equiv {X} (P: X → Prop): + char_rel_limit_computable (char_rel P) ↔ limit_computable P. Proof. split; intros [f Hf]; exists f; intros x. - split; firstorder. - intros []; destruct (Hf x) as [h1 h2]; eauto. Qed. - Lemma char_rel_limit_equiv' {X} (P: X -> Prop): - definite P -> char_rel_limit_computable (char_rel P) <-> char_rel_limit_computable' (char_rel P) . + Lemma char_rel_limit_equiv' {X} (P: X → Prop): + definite P → char_rel_limit_computable (char_rel P) ↔ char_rel_limit_computable' (char_rel P) . Proof. intros HP; split. - intros [f Hf]. exists f; intros. @@ -76,28 +76,28 @@ Convention: Section LimitLemma1. (* Limit computable predicate P is reduciable to K *) - Variable vec_to_nat : forall k, vec nat k -> nat. - Variable nat_to_vec : forall k, nat -> vec nat k. - Variable vec_nat_inv : forall k v, nat_to_vec k (vec_to_nat v) = v. - Variable nat_vec_inv : forall k n, vec_to_nat (nat_to_vec k n) = n. + Variable vec_to_nat : ∀ k, vec nat k → nat. + Variable nat_to_vec : ∀ k, nat → vec nat k. + Variable vec_nat_inv : ∀ k v, nat_to_vec k (vec_to_nat v) = v. + Variable nat_vec_inv : ∀ k n, vec_to_nat (nat_to_vec k n) = n. - Variable list_vec_to_nat : forall k, list (vec nat k) -> nat. - Variable nat_to_list_vec : forall k, nat -> list (vec nat k). - Variable list_vec_nat_inv : forall k v, nat_to_list_vec k (list_vec_to_nat v) = v. - Variable nat_list_vec_inv : forall k n, list_vec_to_nat (nat_to_list_vec k n) = n. + Variable list_vec_to_nat : ∀ k, list (vec nat k) → nat. + Variable nat_to_list_vec : ∀ k, nat → list (vec nat k). + Variable list_vec_nat_inv : ∀ k v, nat_to_list_vec k (list_vec_to_nat v) = v. + Variable nat_list_vec_inv : ∀ k n, list_vec_to_nat (nat_to_list_vec k n) = n. - Variable nat_to_list_bool : nat -> list bool. - Variable list_bool_to_nat : list bool -> nat. - Variable list_bool_nat_inv : forall l, nat_to_list_bool (list_bool_to_nat l) = l. - Variable nat_list_bool_inv : forall n, list_bool_to_nat (nat_to_list_bool n) = n. + Variable nat_to_list_bool : nat → list bool. + Variable list_bool_to_nat : list bool → nat. + Variable list_bool_nat_inv : ∀ l, nat_to_list_bool (list_bool_to_nat l) = l. + Variable nat_list_bool_inv : ∀ n, list_bool_to_nat (nat_to_list_bool n) = n. Section def_K. Hypothesis LEM_Σ_1: LEM_Σ 1. - Lemma semi_dec_def {X} (p: X -> Prop): - semi_decidable p -> definite p. + Lemma semi_dec_def {X} (p: X → Prop): + semi_decidable p → definite p. Proof. intros [f Hf]. unfold semi_decider in Hf. destruct level1 as (_&H2&_). @@ -128,18 +128,18 @@ Section LimitLemma1. (* Extensionality of Σ2, i.e. P t iff ∃ x. ∀ y. f(x, y, t) = true *) - Lemma char_Σ2 {k: nat} (P: vec nat k -> Prop) : - (exists f: nat -> nat -> vec nat k -> bool, forall x, P x <-> (exists n, forall m, f n m x = true)) -> + Lemma char_Σ2 {k: nat} (P: vec nat k → Prop) : + (∃ f: nat → nat → vec nat k → bool, ∀ x, P x ↔ (∃ n, ∀ m, f n m x = true)) → isΣsem 2 P. Proof. intros [f H]. - eapply isΣsemS_ with (p := fun v => forall y, f (hd v) y (tl v) = true). + eapply isΣsemS_ with (p := fun v => ∀ y, f (hd v) y (tl v) = true). eapply isΠsemS_ with (p := fun v => f (hd (tl v)) (hd v) (tl (tl v)) = true). eapply isΣsem0. all: easy. Qed. - Lemma limit_Σ2 {k: nat} (P: vec nat k -> Prop) : - limit_computable P -> isΣsem 2 P /\ isΣsem 2 (compl P). + Lemma limit_Σ2 {k: nat} (P: vec nat k → Prop) : + limit_computable P → isΣsem 2 P ∧ isΣsem 2 (compl P). Proof. intros [f H]; split; eapply char_Σ2. - exists (fun N n x => if lt_dec n N then true else f x n). @@ -153,10 +153,10 @@ Section LimitLemma1. destruct (lt_dec n N); auto; [lia|destruct (f w n); easy]. Qed. - Lemma limit_semi_dec_K {k: nat} (P: vec nat k -> Prop) : - LEM_Σ 1 -> - limit_computable P -> - OracleSemiDecidable K P /\ + Lemma limit_semi_dec_K {k: nat} (P: vec nat k → Prop) : + LEM_Σ 1 → + limit_computable P → + OracleSemiDecidable K P ∧ OracleSemiDecidable K (compl P). Proof. intros LEM H%limit_Σ2. @@ -164,10 +164,10 @@ Section LimitLemma1. all: eauto. Qed. - Lemma limit_turing_red_K' {k: nat} (P: vec nat k -> Prop) : - LEM_Σ 1 -> - definite K -> - limit_computable P -> + Lemma limit_turing_red_K' {k: nat} (P: vec nat k → Prop) : + LEM_Σ 1 → + definite K → + limit_computable P → P ⪯ᴛ K. Proof. intros LEM D H % (limit_semi_dec_K LEM); destruct H as [h1 h2]. @@ -175,22 +175,22 @@ Section LimitLemma1. apply Dec.nat_eq_dec. Qed. - Fact elim_vec (P: nat -> Prop): + Fact elim_vec (P: nat → Prop): P ⪯ₘ (fun x: vec nat 1 => P (hd x)) . Proof. exists (fun x => [x]). now intros x. Qed. (** ** The Limit Lemma 1 *) - Lemma limit_turing_red_K {k: nat} (P: nat -> Prop) : - LEM_Σ 1 -> - limit_computable P -> + Lemma limit_turing_red_K {k: nat} (P: nat → Prop) : + LEM_Σ 1 → + limit_computable P → P ⪯ᴛ K. Proof. intros Hc [h Hh]. specialize (def_K Hc) as Hk. eapply Turing_transitive; last eapply (@limit_turing_red_K' 1); eauto. eapply red_m_impl_red_T. apply elim_vec. - exists (fun v n => h (hd v) n). + exists (fun v n => h (hd v) n). intros x; split; destruct (Hh (hd x)) as [Hh1 Hh2]; eauto. Qed. @@ -204,7 +204,7 @@ Section Σ1Approximation. Lemma semi_dec_halting : semi_decidable K. Proof. eapply OracleSemiDecidable_semi_decidable with (q := ­{0}). - - exists (fun n => match n with | O => true | _ => false end); intros [|n]; easy. + - exists (λ n, match n with | O => true | _ => false end); intros [|n]; easy. - eapply semidecidable_J. Qed. @@ -212,16 +212,16 @@ Section Σ1Approximation. (* Stabilizing the semi decider allows the semi decider to be used as a Σ1 approximation *) - Definition stable (f: nat -> bool) := forall n m, n <= m -> f n = true -> f m = true. + Definition stable (f: nat → bool) := ∀ n m, n ≤ m → f n = true → f m = true. - Fixpoint stabilize_step {X} (f: X -> nat -> bool) x n := + Fixpoint stabilize_step {X} (f: X → nat → bool) x n := match n with | O => false | S n => if f x n then true else stabilize_step f x n end. - Lemma stabilize {X} (P: X -> Prop) : - semi_decidable P -> exists f, semi_decider f P /\ forall x, stable (f x). + Lemma stabilize {X} (P: X → Prop) : + semi_decidable P → ∃ f, semi_decider f P ∧ ∀ x, stable (f x). Proof. intros [f Hf]. exists (fun x n => stabilize_step f x n); split. @@ -239,25 +239,25 @@ Section Σ1Approximation. Qed. (* The Σ1 approximation output correct answers for arbitray list of questions *) - Definition approximation_list {A} (P: A -> Prop) (f: A -> bool) L := - forall i, List.In i L -> P i <-> f i = true. + Definition approximation_list {A} (P: A → Prop) (f: A → bool) L := + ∀ i, List.In i L → P i ↔ f i = true. - Definition approximation_Σ1 {A} (P: A -> Prop) := - exists P_ : nat -> A -> bool, - forall L, exists c, forall c', c' >= c -> approximation_list P (P_ c') L. + Definition approximation_Σ1 {A} (P: A → Prop) := + ∃ P_ : nat → A → bool, + ∀ L, ∃ c, ∀ c', c' ≥ c → approximation_list P (P_ c') L. - Definition approximation_Σ1_strong {A} (P: A -> Prop) := - exists P_ : nat -> A -> bool, - (forall L, exists c, forall c', c' >= c -> approximation_list P (P_ c') L) /\ - (forall tau q a, @interrogation _ _ _ bool tau (char_rel P) q a -> exists n, forall m, m >= n -> interrogation tau (fun q a => P_ m q = a) q a). + Definition approximation_Σ1_strong {A} (P: A → Prop) := + ∃ P_ : nat → A → bool, + (∀ L, ∃ c, ∀ c', c' ≥ c → approximation_list P (P_ c') L) ∧ + (∀ tau q a, @interrogation _ _ _ bool tau (char_rel P) q a → ∃ n, ∀ m, m ≥ n → interrogation tau (fun q a => P_ m q = a) q a). - Definition approximation_Σ1_weak {A} (P: A -> Prop) := - exists P_ : nat -> A -> bool, - (forall tau q a, @interrogation _ _ _ bool tau (char_rel P) q a -> exists n, forall m, m >= n -> interrogation tau (fun q a => P_ m q = a) q a). + Definition approximation_Σ1_weak {A} (P: A → Prop) := + ∃ P_ : nat → A → bool, + (∀ tau q a, @interrogation _ _ _ bool tau (char_rel P) q a → ∃ n, ∀ m, m ≥ n → interrogation tau (λ q a, P_ m q = a) q a). - Lemma semi_dec_approximation_Σ1 {X} (P: X -> Prop) : - definite P -> - semi_decidable P -> approximation_Σ1 P. + Lemma semi_dec_approximation_Σ1 {X} (P: X → Prop) : + definite P → + semi_decidable P → approximation_Σ1 P. Proof. intros defP semiP; unfold approximation_Σ1, approximation_list. destruct (stabilize semiP) as [h [Hh HS]]. @@ -277,9 +277,9 @@ Section Σ1Approximation. rewrite Hc; eauto. Qed. - Lemma semi_dec_approximation_Σ1_strong {X} (P: X -> Prop) : - definite P -> - semi_decidable P -> approximation_Σ1_strong P. + Lemma semi_dec_approximation_Σ1_strong {X} (P: X → Prop) : + definite P → + semi_decidable P → approximation_Σ1_strong P. Proof. intros defP semiP. destruct (semi_dec_approximation_Σ1 defP semiP) as [P_ HP_]. @@ -296,10 +296,10 @@ Section Σ1Approximation. firstorder. Qed. - Lemma approximation_Σ1_halting : definite K -> approximation_Σ1 K. + Lemma approximation_Σ1_halting : definite K → approximation_Σ1 K. Proof. now intros H; apply semi_dec_approximation_Σ1, semi_dec_halting. Qed. - Lemma approximation_Σ1_halting_strong: definite K -> approximation_Σ1_strong K. + Lemma approximation_Σ1_halting_strong: definite K → approximation_Σ1_strong K. Proof. now intros H; apply semi_dec_approximation_Σ1_strong, semi_dec_halting. Qed. @@ -312,19 +312,19 @@ Section LimitLemma2. Section Construction. - Variable f : nat -> nat -> bool. - Variable tau : nat -> tree nat bool bool. - Hypothesis Hf: forall L, exists c, forall c', c' >= c -> approximation_list K (f c') L. + Variable f : nat → nat → bool. + Variable tau : nat → tree nat bool bool. + Hypothesis Hf: ∀ L, ∃ c, ∀ c', c' ≥ c → approximation_list K (f c') L. Definition K_ n := fun i o => f n i = o. Definition char_K_ n := fun i => ret (f n i). - Lemma dec_K_ n : decidable (fun i => f n i = true). + Lemma dec_K_ n : decidable (λ i, f n i = true). Proof. exists (f n). easy. Qed. - Lemma pcomputes_K_ n: pcomputes (char_K_ n) (fun i o => f n i = o). + Lemma pcomputes_K_ n: pcomputes (char_K_ n) (λ i o, f n i = o). Proof. intros i o; split; intro H. now apply ret_hasvalue_inv. @@ -335,10 +335,10 @@ Section LimitLemma2. (** ** The Limit Lemma 2 *) - Theorem turing_red_K_lim (P: nat -> Prop) : - P ⪯ᴛ K -> - definite K -> - definite P -> + Theorem turing_red_K_lim (P: nat → Prop) : + P ⪯ᴛ K → + definite K → + definite P → limit_computable P. Proof. intros [F [H HF]] defK defP. @@ -348,7 +348,7 @@ Section LimitLemma2. pose (char_K_ n := char_K_ k_ n). pose (K_ n := K_ k_ n). pose (Phi x n := evalt_comp (tau x) (k_ n) n n). - assert (forall x y, char_rel P x y -> exists N : nat, forall n : nat, n >= N -> (evalt_comp (tau x) (k_ n)) n n = Some (inr y)) as HL. + assert (∀ x y, char_rel P x y → ∃ N : nat, ∀ n : nat, n ≥ N → (evalt_comp (tau x) (k_ n)) n n = Some (inr y)) as HL. { intros x y H. rewrite HF in H. @@ -364,9 +364,9 @@ Section LimitLemma2. eapply Out. exists L. intros. now apply Hlimt. } - assert (exists f, forall x y, char_rel P x y -> exists N : nat, forall n : nat, n >= N -> f x n = y) as [f HL']. + assert (∃ f, ∀ x y, char_rel P x y → ∃ N : nat, ∀ n : nat, n ≥ N → f x n = y) as [f HL']. { - exists (fun x n => match (Phi x n) with + exists (λ x n, match (Phi x n) with | Some (inr y) => y | _ => false end). intros x y Hxy%HL. destruct (Hxy) as [N HN]. diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index b5513d6..100972c 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -14,7 +14,7 @@ Import ListNotations. Section ComplToBound. Definition complToBound L b : list nat - := filter (fun x => Dec (¬ In x L)) (seq 0 (S b)). + := filter (λ x, Dec (¬ In x L)) (seq 0 (S b)). Lemma complToBound_Bound L b : ∀ x, In x (complToBound L b) → x <= b. @@ -22,7 +22,7 @@ Section ComplToBound. intros x [H % in_seq ?] % in_filter_iff. lia. Qed. Lemma filter_length {X} f (l : list X) : - length l = length (filter f l) + length (filter (fun x => (negb (f x))) l). + length l = length (filter f l) + length (filter (λ x, (negb (f x))) l). Proof. induction l; cbn. - reflexivity. @@ -87,10 +87,10 @@ Section Assume_EA. Definition W_ n e x := φ n e = Some x. Definition W e x := ∃ n, W_ e n x. - Lemma W_spec: ∀ P, semi_decidable P → ∃ e, ∀ x, P x <-> W e x. + Lemma W_spec: ∀ P, semi_decidable P → ∃ e, ∀ x, P x ↔ W e x. Proof. intros P [e He]%EA. exists e; intros x; now rewrite He. Qed. - Notation "'W[' s ']' e" := (fun x => ∃ n, n <= s ∧ W_ e n x) (at level 30). + Notation "'W[' s ']' e" := (λ x, ∃ n, n <= s ∧ W_ e n x) (at level 30). Section EA_dec. @@ -380,7 +380,7 @@ Section Assume_EA. Qed. Lemma recv_at_most_once_bound_gen k: - (∀ k', k' < k → pdec (∃ k0 : nat, recv_att k' k0)) -> + (∀ k', k' < k → pdec (∃ k0 : nat, recv_att k' k0)) → ∃ s, (∀ e, e < k → ∀ s', s < s' → ¬ recv_att e s'). Proof. intros Hle. @@ -490,7 +490,7 @@ Section Assume_EA. Qed. Definition PredListTo p : list nat → nat → Prop - := fun L b => ∀ x, In x L <-> p x ∧ x <= b. + := λ L b, ∀ x, In x L ↔ p x ∧ x <= b. Lemma NoDupBoundH {L} b: NoDup L → (∀ x, In x L → x <= b) → ∀ x, x > b → NoDup (x::L). From 3d9ff2db175276a64e283da514979c0ee4271745 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Thu, 23 Jan 2025 14:48:14 +0100 Subject: [PATCH 52/54] fixed code --- theories/PostsProblem/low_simple_predicates.v | 82 +++-- theories/PostsProblem/lowness.v | 347 ++++++++++-------- theories/PostsProblem/simpleness.v | 2 +- 3 files changed, 233 insertions(+), 198 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index ebd427d..663ea9d 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -34,35 +34,35 @@ End Facts. (* Definition of low *) - Definition low (P: nat -> Prop) := P´ ⪯ᴛ K. + Definition low (P: nat → Prop) := P´ ⪯ᴛ K. Section LowFacts. - Variable vec_to_nat : forall k, vec nat k -> nat. - Variable nat_to_vec : forall k, nat -> vec nat k. - Variable vec_nat_inv : forall k v, nat_to_vec k (vec_to_nat v) = v. - Variable nat_vec_inv : forall k n, vec_to_nat (nat_to_vec k n) = n. - - Variable list_vec_to_nat : forall k, list (vec nat k) -> nat. - Variable nat_to_list_vec : forall k, nat -> list (vec nat k). - Variable list_vec_nat_inv : forall k v, nat_to_list_vec k (list_vec_to_nat v) = v. - Variable nat_list_vec_inv : forall k n, list_vec_to_nat (nat_to_list_vec k n) = n. - - Variable nat_to_list_bool : nat -> list bool. - Variable list_bool_to_nat : list bool -> nat. - Variable list_bool_nat_inv : forall l, nat_to_list_bool (list_bool_to_nat l) = l. - Variable nat_list_bool_inv : forall n, list_bool_to_nat (nat_to_list_bool n) = n. - - Lemma lowness (P: nat -> Prop) : - low P -> ~ K ⪯ᴛ P. + Variable vec_to_nat : ∀ k, vec nat k → nat. + Variable nat_to_vec : ∀ k, nat → vec nat k. + Variable vec_nat_inv : ∀ k v, nat_to_vec k (vec_to_nat v) = v. + Variable nat_vec_inv : ∀ k n, vec_to_nat (nat_to_vec k n) = n. + + Variable list_vec_to_nat : ∀ k, list (vec nat k) → nat. + Variable nat_to_list_vec : ∀ k, nat → list (vec nat k). + Variable list_vec_nat_inv : ∀ k v, nat_to_list_vec k (list_vec_to_nat v) = v. + Variable nat_list_vec_inv : ∀ k n, list_vec_to_nat (nat_to_list_vec k n) = n. + + Variable nat_to_list_bool : nat → list bool. + Variable list_bool_to_nat : list bool → nat. + Variable list_bool_nat_inv : ∀ l, nat_to_list_bool (list_bool_to_nat l) = l. + Variable nat_list_bool_inv : ∀ n, list_bool_to_nat (nat_to_list_bool n) = n. + + Lemma lowness (P: nat → Prop) : + low P → ¬ K ⪯ᴛ P. Proof. intros H IH. eapply not_turing_red_J with (Q := P). eapply Turing_transitive; [apply H| easy]. Qed. - Lemma DN_lowness (P: nat -> Prop) : - ~ ~ low P -> ~ K ⪯ᴛ P. + Lemma DN_lowness (P: nat → Prop) : + ¬¬ low P → ¬ K ⪯ᴛ P. Proof. intros H_ IH. apply H_. intros H. @@ -70,23 +70,23 @@ Section LowFacts. eapply Turing_transitive; [apply H| easy]. Qed. - Lemma limit_jump_lowness (A: nat -> Prop) : - LEM_Σ 1 -> - definite K -> - limit_computable (A´) -> ~ K ⪯ᴛ A. + Lemma limit_jump_lowness (A: nat → Prop) : + LEM_Σ 1 → + definite K → + limit_computable (A´) → ¬ K ⪯ᴛ A. Proof. intros LEM defK H IH. apply lowness with (P := A); [|apply IH]. eapply limit_turing_red_K; eauto. exact 42. Qed. - Definition low_simple P := low P /\ simple P. + Definition low_simple P := low P ∧ simple P. - Definition sol_Post's_problem (P: nat -> Prop) := - (~ decidable P) /\ (enumerable P) /\ ~ (K ⪯ᴛ P). + Definition sol_Post's_problem (P: nat → Prop) := + (¬ decidable P) ∧ (enumerable P) ∧ ¬ (K ⪯ᴛ P). Fact low_simple_correct: - forall P, low_simple P -> sol_Post's_problem P. + ∀ P, low_simple P → sol_Post's_problem P. Proof. intros P [H1 H2]; split; [now apply simple_undecidable|]. split; [destruct H2 as [H2 _]; eauto| now apply lowness]. @@ -98,14 +98,14 @@ Section LowFacts. Section LowSimplePredicate. - Variable η: nat -> nat -> option nat. + Variable η: nat → nat → option nat. Hypothesis EA: - forall P, semi_decidable P -> exists e, forall x, P x <-> exists n, η e n = Some x. + ∀ P, semi_decidable P → exists e, ∀ x, P x ↔ ∃ n, η e n = Some x. Hypothesis LEM_Σ_1: LEM_Σ 1. Hypothesis def_K: definite K. - Theorem a_sol_Post's_problem: exists P, sol_Post's_problem P. + Theorem a_sol_Post's_problem: ∃ P, sol_Post's_problem P. Proof. eexists. eapply low_simple_correct; split. - eapply limit_turing_red_K; eauto. exact 42. @@ -117,17 +117,18 @@ Section LowFacts. End LowSimplePredicate. - Notation "(¬¬Σ⁰₁)-LEM" := ((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ~~ definite p)) (at level 0). + Notation "(¬¬Σ⁰₁)-LEM" := + ((∀ (k : nat) (p : vec nat k → Prop), isΣsem 1 p → ¬¬ definite p)) (at level 0). Lemma m_red_K_semi_decidable {n} (P: vec nat n → Prop): - semi_decidable P -> P ⪯ₘ K. + semi_decidable P → P ⪯ₘ K. Proof. intros H. unfold K. rewrite <- red_m_iff_semidec_jump_vec. by apply semi_decidable_OracleSemiDecidable. eauto. Qed. Theorem PostProblem_from_neg_negLPO : - ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM -> ¬ K ⪯ᴛ p). + ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM → ¬ K ⪯ᴛ p). Proof. eexists. repeat split. @@ -166,15 +167,16 @@ Print Assumptions PostProblem_from_neg_negLPO. (* general proof that (¬¬Σ⁰₁)-LEM <-> ¬¬(Σ⁰₁)-LEM under many-one complete Σ⁰₁ predicate *) Section assume. - Variable enumerable : (nat -> Prop) -> Prop. + Variable enumerable : (nat → Prop) → Prop. - Variable K : nat -> Prop. + Variable K : nat → Prop. Variable eK : enumerable K. - Variable cK : forall p : nat -> Prop, enumerable p -> red_m p K. + Variable cK : ∀ p : nat → Prop, enumerable p → red_m p K. - Goal ~~ (forall p : nat -> Prop, enumerable p -> forall x, p x \/ ~ p x) - <-> - (forall p : nat -> Prop, enumerable p -> ~~ forall x, p x \/ ~ p x). + Goal + ¬¬ (∀ p : nat → Prop, enumerable p → ∀ x, p x ∨ ¬ p x) + ↔ + (∀ p : nat → Prop, enumerable p → ¬¬ ∀ x, p x ∨ ¬ p x). Proof. split. - firstorder. diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index 2d1cef2..daf2715 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -38,7 +38,7 @@ Section Requirements_Verification. Hypothesis step_ex_spec: ∀ e, (∞∃ n, Ω e n ↓) → Ξ e (char_rel P) e. Hypothesis N_requirements: ∀ e, (∞∀ n, Ω e n ↓) ∨ (∞∀ n, ¬ Ω e n ↓). - Lemma Jump_limit_1 : limit_computable (P´). + Lemma Jump_limit : limit_computable (P´). Proof. exists limit_decider; split; intros. - unfold J. split. @@ -83,28 +83,10 @@ Section Requirements_Meet. Section wall_greater_than_use. Hypothesis wall_spec: forall e L n, φ (λ x, Dec (In x L)) e e n ≤ wall e L n. - Hypothesis Σ_1_lem: LEM_Σ 1. - Lemma attention_bound: - ∀ k, ∃ s, ∀ e, e < k → ∀ s', s < s' → ~ recv_att wall e s'. - Proof. by apply recv_at_most_once_bound. Qed. - Lemma eventally_greater_than_wall: - ∀ e, (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). - Proof. - intros e. - destruct (@attention_bound (S e)) as [s Hs]. - exists (S s). intros m Hm x [e_ [He_ He_']]. - destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (recv_att wall e_ m). - unshelve eapply (Hs e_ _ m); eauto; lia. - split; first lia. destruct He_' as [H _]. - apply H. } - assert (e ≤ e_) by lia; clear E. - destruct He_', H1, H1, H1, H1, H3. - by eapply H5. - Qed. +(* Lemma N_requirements e : (∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓). Proof. destruct (@eventally_greater_than_wall e) as [N HN]. @@ -138,7 +120,189 @@ Section Requirements_Meet. assumption. - right. exists N. intros m Hm. destruct (Dec (P_Ω e m ↓)); eauto. - Qed. + Qed. *) + + + Definition use := λ e n, φ (χ n) e e n. + + Lemma eventally_greater_than_use_gen e: + (∃ n, (∀ e', e' < e → ∀ n', n < n' → ¬ (recv_att wall) e' n')) → + (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). + Proof. + intros Hc. destruct Hc as [s Hs]. + exists (S s). intros m Hm x [e_ [He_ He_']]. + destruct (Dec (e_ < e)) as [E|E]. + { exfalso. enough (recv_att wall e_ m). + unshelve eapply (Hs e_ _ m); eauto; lia. + split; first lia. destruct He_' as [H _]. + apply H. } + assert (e ≤ e_) by lia; clear E. + destruct He_', H1, H1, H1, H1, H3. + specialize(H5 e H). + enough (use e m <= wall e (P_func m) m) by lia. + apply wall_spec. + Qed. + + Lemma eventally_greater_than_use_classically e: + ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). + Proof. + intros H_. eapply (@recv_at_most_once_bound_classically wall e). + intros H'. apply H_. + by apply eventally_greater_than_use_gen. + Qed. + + + Fact wall_convergence_gen e : + (∃ N, (∀ s, N ≤ s → ∀ x, extendP (P_func s) s x → use e s < x) + ∧ pdec (∃ x, N ≤ x ∧ use e x ≠ 0)) -> + (∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b). + Proof. + intros (N & HN & Hp). + assert (∀ m, dec (use e m = 0)) as HD by eauto. + destruct Hp as [H'|H']. + - enough (∃ x, N ≤ x ∧ use e x ≠ 0) as H''. + clear H'. destruct H'' as [x [H1 H2]]. + destruct (use e x) as [|k] eqn: H; first done; clear H2. + exists (S k), x. intros t Ht. induction Ht; first done. + rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). + reflexivity. + intros; split. + + intros K%Dec_true. apply Dec_auto. + enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). + eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. + inv HE. + * assert (use e m < a). + { eapply HN. lia. enough (P_func m = l) as ->. apply H5. + eapply F_uni; [apply F_func_correctness|apply H4]. } + assert (wall e (P_func m) m ≥ S k). { rewrite <-IHHt. + unfold use. unfold χ, the_priority_method.χ. + specialize (wall_spec e (F_func (simple_extension wall) m) m). + eapply wall_spec. } + destruct (Dec (a = x0)). + lia. apply Dec_auto. rewrite <- H3 in K. + destruct K; first done. + enough ((F_func (simple_extension wall) m) = l) as -> by done. + eapply F_uni; last apply H4; apply F_func_correctness. + * apply Dec_auto. + enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. + eapply F_uni; first apply F_func_correctness. + assumption. + + apply H'. + - exists 0. exists N. intros. + destruct (use e m) eqn: E; first done. + exfalso. apply H'. exists m. split; first done. + congruence. + Qed. + + + Fact wall_convergence_classically e : ¬¬ ∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b. + Proof. + intros H. eapply (eventally_greater_than_use_classically). intros [N HN]. + assert (¬¬ pdec (∃ x, N ≤ x ∧ use e x ≠ 0)). + { ccase (∃ x, N ≤ x ∧ use e x ≠ 0) as [H_|H'_]; eauto. + - intros Hp. apply Hp. left. done. + - intros Hp. apply Hp. right. done. } + apply H0. intros H0'. + apply H, wall_convergence_gen. + exists N; eauto. + Qed. + + Hypothesis Σ_1_lem: LEM_Σ 1. + + Lemma attention_bound: + ∀ k, ∃ s, ∀ e, e < k → ∀ s', s < s' → ~ recv_att wall e s'. + Proof. by apply recv_at_most_once_bound. Qed. + + Lemma eventally_greater_than_use e: + (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). + Proof. apply eventally_greater_than_use_gen, attention_bound. Qed. + + Fact wall_convergence e : ∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b. + Proof. + apply wall_convergence_gen. + destruct (eventally_greater_than_use e) as [N HN]. + exists N. split; first done. + destruct level1 as (_&h2&_). + rewrite h2 in Σ_1_lem. + unfold principles.LPO in *. + destruct (@Σ_1_lem (λ x, Dec (N ≤ x ∧ use e x ≠ 0))) as [[N' HN']|]. + left. exists N'. eauto. + right. intros [x Hx]. apply H. exists x. eauto. + Qed. + + Corollary N_requirements e : ((∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓)). + Proof. + destruct (wall_convergence e) as (b&n&Hn). + destruct b. + - right. exists n. intros. + enough(P_Ω e m = None) by congruence. + unfold P_Ω, Ω. + apply φ_spec0'. + by apply Hn. + - left. exists n. intros. + unfold P_Ω, Ω. + apply φ_spec0_1. + unfold use in Hn. + by rewrite (Hn m H). + Qed. + + Lemma eventally_greater_than_wall e: + (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). + Proof. + destruct (@attention_bound (S e)) as [s Hs]. + exists (S s). intros m Hm x [e_ [He_ He_']]. + destruct (Dec (e_ < e)) as [E|E]. + { exfalso. enough (recv_att wall e_ m). + unshelve eapply (Hs e_ _ m); eauto; lia. + split; first lia. destruct He_' as [H _]. + apply H. } + assert (e ≤ e_) by lia; clear E. + destruct He_', H1, H1, H1, H1, H3. + by eapply H5. + Qed. + + Lemma step_ex_spec: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. + Proof. + intros e He. + destruct (eventally_greater_than_wall e) as [N HN]. + destruct (wall_convergence e) as [B [b Hb]]. + set (M := max N b). destruct (He M) as [k [Hk Hk']]. + eapply (@φ_spec χ e e k); first apply Hk'. + intros x Hx. unfold P, simpleness.P. + rewrite F_with_top. split. + - intros (L & m & HL & HLs &HP). + assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } + assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } + apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. + enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. + eauto. eapply F_mono; last apply E_; apply F_func_correctness. + assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). + assert (k ≤ m) as Ek by lia. + enough (x ≤ φ (χ m) e e m). + exfalso. assert (φ (χ m) e e m < x). + assert(φ (χ m) e e m ≤ wall e (P_func m) m). + { specialize (wall_spec e (F_func (simple_extension wall) m) m). + unfold χ, the_priority_method.χ. + by rewrite wall_spec. } + lia. + enough (φ (χ m) e e m = φ (χ k) e e k) by lia. + assert (φ (χ m) e e m = B). + { rewrite <- (Hb m). reflexivity. lia. } + assert (φ (χ k) e e k = B). + { rewrite <- (Hb k). reflexivity. lia. } + lia. + enough (φ (χ m) e e m = φ (χ k) e e k) by lia. + assert (φ (χ m) e e m = B). + { rewrite <- (Hb m). reflexivity. lia. } + assert (φ (χ k) e e k = B). + { rewrite <- (Hb k). reflexivity. lia. } + lia. + - intros H%Dec_true. + eapply F_with_top. + exists (F_func _ k), k; split; eauto. + eapply F_func_correctness. + Qed. End wall_greater_than_use. @@ -152,141 +316,9 @@ Section Concret_Wall. Instance E_low: Extension := simple_extension wall. Hypothesis Σ_1_lem: LEM_Σ 1. - Fact wall_convergence e : ∃ b : nat, lim_to wall (wall e) b. - Proof. - destruct (eventally_greater_than_wall wall Σ_1_lem e) as [N HN]. - assert (∀ m, dec (wall e (P_func wall m) m = 0)) as HD by eauto. - assert (pdec (∀ x, N ≤ x → wall e (P_func wall x) x = 0)) as [H'|H']. - { apply assume_Π_1_lem. apply Σ_1_lem. intros x. eauto. } - - (* ccase (∀ x, N ≤ x → wall e (P_func x) x = 0) as [H'|H']. *) - - exists 0. exists N. intros. by apply H'. - - enough (∃ x, N ≤ x ∧ wall e (P_func wall x) x ≠ 0) as H''. - clear H'. destruct H'' as [x [H1 H2]]. - destruct (wall e (P_func wall x) x) as [|k] eqn: H; first done; clear H2. - exists (S k), x. intros t Ht. induction Ht; first done. - rewrite <- (@φ_spec1 (χ wall) _ _ _ _ IHHt). - reflexivity. - intros; split. - + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). - eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. - inv HE. - * assert (wall e (P_func wall m) m < a). - { eapply HN. lia. enough (P_func wall m = l) as ->. apply H5. - eapply F_uni; [apply F_func_correctness|apply H4]. } - assert (wall e (P_func wall m) m = S k). { rewrite <-IHHt. reflexivity. } - rewrite H6 in H2. destruct (Dec (a = x0)). - lia. apply Dec_auto. rewrite <- H3 in K. - destruct K; first done. - enough ((F_func (simple_extension wall) m) = l) as -> by done. - eapply F_uni; last apply H4; apply F_func_correctness. - * apply Dec_auto. - enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. - eapply F_uni; first apply F_func_correctness. - assumption. - + eapply assume_Σ_1_dne. apply Σ_1_lem. - { intro x. eauto. } - intro H. apply H'. intros x Hx. - assert (∀ n m: nat, ~~n=m → n=m). - { intros n m Hnm. destruct (Dec (n=m)); eauto. - exfalso. by apply Hnm. } - apply H0. intros Hmn. - apply H. now exists x; split. - Qed. - - Lemma step_ex_spec: ∀ e, (∞∃ n, P_Ω wall e n ↓) → Ξ e (char_rel (P wall)) e. - Proof. - intros e He. - destruct (@eventally_greater_than_wall wall Σ_1_lem e) as [N HN]. - destruct (@wall_convergence e) as [B [b Hb]]. - set (M := max N b). destruct (He M) as [k [Hk Hk']]. - eapply (@φ_spec (χ wall) e e k); first apply Hk'. - intros x Hx. unfold P, simpleness.P. - rewrite F_with_top. split. - - intros (L & m & HL & HLs &HP). - assert (L = (P_func wall) m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } - assert (x::L = (P_func wall) (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } - apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. - enough (incl ((P_func wall) (S m)) ((P_func wall) k)). rewrite <-E' in H. - eauto. eapply F_mono; last apply E_; apply F_func_correctness. - assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). - assert (k ≤ m) as Ek by lia. enough (x ≤ φ (χ wall m) e e m). - exfalso. assert (φ (χ wall m) e e m < x). - apply HN. lia. enough (φ (χ wall m) e e m = φ (χ wall k) e e k) by congruence. - assert (φ (χ wall m) e e m = B). - { rewrite <- (Hb m). reflexivity. lia. } - assert (φ (χ wall k) e e k = B). - { rewrite <- (Hb k). reflexivity. lia. } - congruence. - - intros H%Dec_true. - eapply F_with_top. - exists (F_func _ k), k; split; eauto. - eapply F_func_correctness. - Qed. - - Lemma eventally_greater_than_wall_classically e: - ¬¬ (∞∀ s, ∀ x, extendP (P_func wall s) s x → wall e (P_func wall s) s < x). - Proof. - intros H_. eapply (@recv_at_most_once_bound_classically wall (S e)). - intros [s Hs]. eapply H_; clear H_. - exists (S s). intros m Hm x [e_ [He_ He_']]. - destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (recv_att wall e_ m). - unshelve eapply (Hs e_ _ m); eauto; lia. - split; first lia. destruct He_' as [H _]. - apply H. } - assert (e ≤ e_) by lia; clear E. - destruct He_', H1, H1, H1, H1, H3. - by eapply H5. - Qed. - - Fact wall_convergence_classically e : ¬¬ ∃ b : nat, lim_to wall (wall e) b. - Proof. - intros H_. - destruct (@eventally_greater_than_wall_classically e). intros [N HN]. - assert (∀ m, dec (wall e (P_func wall m) m = 0)) as HD by eauto. - ccase (∀ x, N ≤ x → wall e (P_func wall x) x = 0) as [H'|H']. - - apply H_; clear H_. exists 0. exists N. intros. by apply H'. - - enough (~~∃ x, N ≤ x ∧ wall e (P_func wall x) x ≠ 0) as H__. - apply H__. intros H''. - clear H'. destruct H'' as [x [H1 H2]]. - apply H_; clear H_. - destruct (wall e (P_func wall x) x) as [|k] eqn: H; first done; clear H2. - exists (S k), x. intros t Ht. induction Ht; first done. - rewrite <- (@φ_spec1 (χ wall) _ _ _ _ IHHt). - reflexivity. - intros; split. - + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). - eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. - inv HE. - * assert (wall e (P_func wall m) m < a). - { eapply HN. lia. enough (P_func wall m = l) as ->. apply H5. - eapply F_uni; [apply F_func_correctness|apply H4]. } - assert (wall e (P_func wall m) m = S k). { rewrite <-IHHt. reflexivity. } - rewrite H6 in H2. destruct (Dec (a = x0)). - lia. apply Dec_auto. rewrite <- H3 in K. - destruct K; first done. - enough ((F_func (simple_extension wall) m) = l) as -> by done. - eapply F_uni; last apply H4; apply F_func_correctness. - * apply Dec_auto. - enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. - eapply F_uni; first apply F_func_correctness. - assumption. - + intro H. - apply H'. intros x Hx. - assert (∀ n m: nat, ~~n=m → n=m). - { intros n m Hnm. destruct (Dec (n=m)); eauto. - exfalso. by apply Hnm. } - apply H0. intros Hmn. - apply H. now exists x; split. - Qed. - Fact P_simple: simple (P wall). - Proof. eapply P_simple, wall_convergence_classically. Qed. + Proof. eapply P_simple, wall_convergence_classically. + unfold wall. done. Qed. End Concret_Wall. @@ -296,8 +328,9 @@ Section Results. Fact jump_P_limit: limit_computable ((P wall)´). Proof. - eapply Jump_limit_1; first apply F_with_χ. - - intros e He. eapply step_ex_spec; eauto. + eapply Jump_limit; first apply F_with_χ. + - intros e He. + eapply step_ex_spec; eauto. - eapply N_requirements; eauto. Qed. diff --git a/theories/PostsProblem/simpleness.v b/theories/PostsProblem/simpleness.v index 100972c..b586d87 100644 --- a/theories/PostsProblem/simpleness.v +++ b/theories/PostsProblem/simpleness.v @@ -638,7 +638,7 @@ Section Assume_EA. apply H. intros [m [Hm1 [k Hmk]]]. apply He. exists k, m. repeat split. now exists k. - lia. intros n i Hie. specialize (Hw i n Hie). lia. + lia. intros n i Hie. specialize (Hw i n Hie). lia. Qed. Lemma ext_pick_impl_recv N e: From e79e1823a0742a858c8c695a427ee43abb0b56cd Mon Sep 17 00:00:00 2001 From: Haoyi Date: Thu, 23 Jan 2025 19:39:45 +0100 Subject: [PATCH 53/54] updated --- theories/PostsProblem/lowness.v | 359 ++++++++++++++------------------ 1 file changed, 160 insertions(+), 199 deletions(-) diff --git a/theories/PostsProblem/lowness.v b/theories/PostsProblem/lowness.v index daf2715..76e18eb 100644 --- a/theories/PostsProblem/lowness.v +++ b/theories/PostsProblem/lowness.v @@ -84,225 +84,186 @@ Section Requirements_Meet. Hypothesis wall_spec: forall e L n, φ (λ x, Dec (In x L)) e e n ≤ wall e L n. + Definition use := λ e n, φ (χ n) e e n. + + Lemma eventally_greater_than_use_gen e: + (∃ n, (∀ e', e' < e → ∀ n', n < n' → ¬ (recv_att wall) e' n')) → + (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). + Proof. + intros Hc. destruct Hc as [s Hs]. + exists (S s). intros m Hm x [e_ [He_ He_']]. + destruct (Dec (e_ < e)) as [E|E]. + { exfalso. enough (recv_att wall e_ m). + unshelve eapply (Hs e_ _ m); eauto; lia. + split; first lia. destruct He_' as [H _]. + apply H. } + assert (e ≤ e_) by lia; clear E. + destruct He_', H1, H1, H1, H1, H3. + specialize(H5 e H). + enough (use e m <= wall e (P_func m) m) by lia. + apply wall_spec. + Qed. + + Lemma eventally_greater_than_use_classically e: + ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). + Proof. + intros H_. eapply (@recv_at_most_once_bound_classically wall e). + intros H'. apply H_. + by apply eventally_greater_than_use_gen. + Qed. -(* - Lemma N_requirements e : (∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓). + Fact wall_convergence_gen e : + (∃ N, (∀ s, N ≤ s → ∀ x, extendP (P_func s) s x → use e s < x) + ∧ pdec (∃ x, N ≤ x ∧ use e x ≠ 0)) -> + (∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b). Proof. - destruct (@eventally_greater_than_wall e) as [N HN]. - assert (pdec (∃ k, N ≤ k ∧ P_Ω e k ↓)) as [[k [Hk HNk]]|H'] by (apply assume_Σ_1_lem; eauto). - - left. exists k. intros n Hm. - induction Hm; first done. - unfold P_Ω, Ω, Φ_ in *. - destruct (φ (χ m) e e m) eqn: E. - { eapply φ_spec0' in E. congruence. } - (* Check φ_spec2. *) - eapply (@φ_spec2 χ _ ); eauto. - intros x Hx; split. + intros (N & HN & Hp). + assert (∀ m, dec (use e m = 0)) as HD by eauto. + destruct Hp as [H'|H']. + - enough (∃ x, N ≤ x ∧ use e x ≠ 0) as H''. + clear H'. destruct H'' as [x [H1 H2]]. + destruct (use e x) as [|k] eqn: H; first done; clear H2. + exists (S k), x. intros t Ht. induction Ht; first done. + rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). + reflexivity. + intros; split. + intros K%Dec_true. apply Dec_auto. enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. inv HE. - * assert (wall e (P_func m) m < a). - { eapply HN. lia. enough (P_func m = l) as ->. apply H2. - eapply F_uni; [apply F_func_correctness|apply H1]. } - assert (wall e (P_func m) m ≥ S n). { - rewrite <-E. apply wall_spec. } - destruct (Dec (a = x)). - lia. apply Dec_auto. rewrite <- H0 in K. + * assert (use e m < a). + { eapply HN. lia. enough (P_func m = l) as ->. apply H5. + eapply F_uni; [apply F_func_correctness|apply H4]. } + assert (wall e (P_func m) m ≥ S k). { rewrite <-IHHt. + unfold use. unfold χ, the_priority_method.χ. + specialize (wall_spec e (F_func (simple_extension wall) m) m). + eapply wall_spec. } + destruct (Dec (a = x0)). + lia. apply Dec_auto. rewrite <- H3 in K. destruct K; first done. enough ((F_func (simple_extension wall) m) = l) as -> by done. - eapply F_uni; last apply H1; apply F_func_correctness. + eapply F_uni; last apply H4; apply F_func_correctness. * apply Dec_auto. enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. eapply F_uni; first apply F_func_correctness. - assumption. - - right. exists N. intros m Hm. - destruct (Dec (P_Ω e m ↓)); eauto. - Qed. *) - - - Definition use := λ e n, φ (χ n) e e n. - - Lemma eventally_greater_than_use_gen e: - (∃ n, (∀ e', e' < e → ∀ n', n < n' → ¬ (recv_att wall) e' n')) → - (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). - Proof. - intros Hc. destruct Hc as [s Hs]. - exists (S s). intros m Hm x [e_ [He_ He_']]. - destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (recv_att wall e_ m). - unshelve eapply (Hs e_ _ m); eauto; lia. - split; first lia. destruct He_' as [H _]. - apply H. } - assert (e ≤ e_) by lia; clear E. - destruct He_', H1, H1, H1, H1, H3. - specialize(H5 e H). - enough (use e m <= wall e (P_func m) m) by lia. - apply wall_spec. - Qed. - - Lemma eventally_greater_than_use_classically e: - ¬¬ (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). - Proof. - intros H_. eapply (@recv_at_most_once_bound_classically wall e). - intros H'. apply H_. - by apply eventally_greater_than_use_gen. - Qed. - - - Fact wall_convergence_gen e : - (∃ N, (∀ s, N ≤ s → ∀ x, extendP (P_func s) s x → use e s < x) - ∧ pdec (∃ x, N ≤ x ∧ use e x ≠ 0)) -> - (∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b). - Proof. - intros (N & HN & Hp). - assert (∀ m, dec (use e m = 0)) as HD by eauto. - destruct Hp as [H'|H']. - - enough (∃ x, N ≤ x ∧ use e x ≠ 0) as H''. - clear H'. destruct H'' as [x [H1 H2]]. - destruct (use e x) as [|k] eqn: H; first done; clear H2. - exists (S k), x. intros t Ht. induction Ht; first done. - rewrite <- (@φ_spec1 χ _ _ _ _ IHHt). - reflexivity. - intros; split. - + intros K%Dec_true. apply Dec_auto. - enough (incl (F_func (simple_extension wall) m) (F_func (simple_extension wall) (S m))). - eauto. eapply F_mono; [apply F_func_correctness|apply F_func_correctness|lia]. - + intros K%Dec_true. specialize (F_func_correctness (simple_extension wall) (S m)) as HE. - inv HE. - * assert (use e m < a). - { eapply HN. lia. enough (P_func m = l) as ->. apply H5. - eapply F_uni; [apply F_func_correctness|apply H4]. } - assert (wall e (P_func m) m ≥ S k). { rewrite <-IHHt. - unfold use. unfold χ, the_priority_method.χ. - specialize (wall_spec e (F_func (simple_extension wall) m) m). - eapply wall_spec. } - destruct (Dec (a = x0)). - lia. apply Dec_auto. rewrite <- H3 in K. - destruct K; first done. - enough ((F_func (simple_extension wall) m) = l) as -> by done. - eapply F_uni; last apply H4; apply F_func_correctness. - * apply Dec_auto. - enough (F_func (simple_extension wall) m = F_func (simple_extension wall) (S m)) as -> by eauto. - eapply F_uni; first apply F_func_correctness. - assumption. - + apply H'. - - exists 0. exists N. intros. - destruct (use e m) eqn: E; first done. - exfalso. apply H'. exists m. split; first done. - congruence. - Qed. - - - Fact wall_convergence_classically e : ¬¬ ∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b. - Proof. - intros H. eapply (eventally_greater_than_use_classically). intros [N HN]. - assert (¬¬ pdec (∃ x, N ≤ x ∧ use e x ≠ 0)). - { ccase (∃ x, N ≤ x ∧ use e x ≠ 0) as [H_|H'_]; eauto. - - intros Hp. apply Hp. left. done. - - intros Hp. apply Hp. right. done. } - apply H0. intros H0'. - apply H, wall_convergence_gen. - exists N; eauto. - Qed. + assumption. + + apply H'. + - exists 0. exists N. intros. + destruct (use e m) eqn: E; first done. + exfalso. apply H'. exists m. split; first done. + congruence. + Qed. - Hypothesis Σ_1_lem: LEM_Σ 1. - Lemma attention_bound: - ∀ k, ∃ s, ∀ e, e < k → ∀ s', s < s' → ~ recv_att wall e s'. - Proof. by apply recv_at_most_once_bound. Qed. + Fact wall_convergence_classically e : ¬¬ ∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b. + Proof. + intros H. eapply (eventally_greater_than_use_classically). intros [N HN]. + assert (¬¬ pdec (∃ x, N ≤ x ∧ use e x ≠ 0)). + { ccase (∃ x, N ≤ x ∧ use e x ≠ 0) as [H_|H'_]; eauto. + - intros Hp. apply Hp. left. done. + - intros Hp. apply Hp. right. done. } + apply H0. intros H0'. + apply H, wall_convergence_gen. + exists N; eauto. + Qed. - Lemma eventally_greater_than_use e: - (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). - Proof. apply eventally_greater_than_use_gen, attention_bound. Qed. + Hypothesis Σ_1_lem: LEM_Σ 1. - Fact wall_convergence e : ∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b. - Proof. - apply wall_convergence_gen. - destruct (eventally_greater_than_use e) as [N HN]. - exists N. split; first done. - destruct level1 as (_&h2&_). - rewrite h2 in Σ_1_lem. - unfold principles.LPO in *. - destruct (@Σ_1_lem (λ x, Dec (N ≤ x ∧ use e x ≠ 0))) as [[N' HN']|]. - left. exists N'. eauto. - right. intros [x Hx]. apply H. exists x. eauto. - Qed. + Lemma attention_bound: + ∀ k, ∃ s, ∀ e, e < k → ∀ s', s < s' → ~ recv_att wall e s'. + Proof. by apply recv_at_most_once_bound. Qed. - Corollary N_requirements e : ((∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓)). - Proof. - destruct (wall_convergence e) as (b&n&Hn). - destruct b. - - right. exists n. intros. - enough(P_Ω e m = None) by congruence. - unfold P_Ω, Ω. - apply φ_spec0'. - by apply Hn. - - left. exists n. intros. - unfold P_Ω, Ω. - apply φ_spec0_1. - unfold use in Hn. - by rewrite (Hn m H). - Qed. + Lemma eventally_greater_than_use e: + (∞∀ s, ∀ x, extendP (P_func s) s x → use e s < x). + Proof. apply eventally_greater_than_use_gen, attention_bound. Qed. - Lemma eventally_greater_than_wall e: - (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). - Proof. - destruct (@attention_bound (S e)) as [s Hs]. - exists (S s). intros m Hm x [e_ [He_ He_']]. - destruct (Dec (e_ < e)) as [E|E]. - { exfalso. enough (recv_att wall e_ m). - unshelve eapply (Hs e_ _ m); eauto; lia. - split; first lia. destruct He_' as [H _]. - apply H. } - assert (e ≤ e_) by lia; clear E. - destruct He_', H1, H1, H1, H1, H3. - by eapply H5. - Qed. - - Lemma step_ex_spec: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. - Proof. - intros e He. - destruct (eventally_greater_than_wall e) as [N HN]. - destruct (wall_convergence e) as [B [b Hb]]. - set (M := max N b). destruct (He M) as [k [Hk Hk']]. - eapply (@φ_spec χ e e k); first apply Hk'. - intros x Hx. unfold P, simpleness.P. - rewrite F_with_top. split. - - intros (L & m & HL & HLs &HP). - assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } - assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } - apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. - enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. - eauto. eapply F_mono; last apply E_; apply F_func_correctness. - assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). - assert (k ≤ m) as Ek by lia. - enough (x ≤ φ (χ m) e e m). - exfalso. assert (φ (χ m) e e m < x). - assert(φ (χ m) e e m ≤ wall e (P_func m) m). - { specialize (wall_spec e (F_func (simple_extension wall) m) m). - unfold χ, the_priority_method.χ. - by rewrite wall_spec. } - lia. - enough (φ (χ m) e e m = φ (χ k) e e k) by lia. - assert (φ (χ m) e e m = B). - { rewrite <- (Hb m). reflexivity. lia. } - assert (φ (χ k) e e k = B). - { rewrite <- (Hb k). reflexivity. lia. } + Fact wall_convergence e : ∃ b n: nat, ∀ m : nat, n ≤ m → use e m = b. + Proof. + apply wall_convergence_gen. + destruct (eventally_greater_than_use e) as [N HN]. + exists N. split; first done. + destruct level1 as (_&h2&_). + rewrite h2 in Σ_1_lem. + unfold principles.LPO in *. + destruct (@Σ_1_lem (λ x, Dec (N ≤ x ∧ use e x ≠ 0))) as [[N' HN']|]. + left. exists N'. eauto. + right. intros [x Hx]. apply H. exists x. eauto. + Qed. + + Corollary N_requirements e : ((∞∀ n, P_Ω e n ↓) ∨ (∞∀ n, ¬ P_Ω e n ↓)). + Proof. + destruct (wall_convergence e) as (b&n&Hn). + destruct b. + - right. exists n. intros. + enough(P_Ω e m = None) by congruence. + unfold P_Ω, Ω. + apply φ_spec0'. + by apply Hn. + - left. exists n. intros. + unfold P_Ω, Ω. + apply φ_spec0_1. + unfold use in Hn. + by rewrite (Hn m H). + Qed. + + Lemma eventally_greater_than_wall e: + (∞∀ s, ∀ x, extendP (P_func s) s x → wall e (P_func s) s < x). + Proof. + destruct (@attention_bound (S e)) as [s Hs]. + exists (S s). intros m Hm x [e_ [He_ He_']]. + destruct (Dec (e_ < e)) as [E|E]. + { exfalso. enough (recv_att wall e_ m). + unshelve eapply (Hs e_ _ m); eauto; lia. + split; first lia. destruct He_' as [H _]. + apply H. } + assert (e ≤ e_) by lia; clear E. + destruct He_', H1, H1, H1, H1, H3. + by eapply H5. + Qed. + + Lemma step_ex_spec: ∀ e, (∞∃ n, P_Ω e n ↓) → Ξ e (char_rel P) e. + Proof. + intros e He. + destruct (eventally_greater_than_wall e) as [N HN]. + destruct (wall_convergence e) as [B [b Hb]]. + set (M := max N b). destruct (He M) as [k [Hk Hk']]. + eapply (@φ_spec χ e e k); first apply Hk'. + intros x Hx. unfold P, simpleness.P. + rewrite F_with_top. split. + - intros (L & m & HL & HLs &HP). + assert (L = P_func m) as E. { eapply F_uni. apply HL. apply F_func_correctness. } + assert (x::L = P_func (S m)) as E'. { eapply F_uni. apply HLs. apply F_func_correctness. } + apply Dec_auto. destruct (Dec (S m ≤ k)) as [E_|E_]. + enough (incl (P_func (S m)) (P_func k)). rewrite <-E' in H. + eauto. eapply F_mono; last apply E_; apply F_func_correctness. + assert (N ≤ m) as Em by lia. rewrite E in HP. specialize (HN m Em x HP). + assert (k ≤ m) as Ek by lia. + enough (x ≤ φ (χ m) e e m). + exfalso. assert (φ (χ m) e e m < x). + assert(φ (χ m) e e m ≤ wall e (P_func m) m). + { specialize (wall_spec e (F_func (simple_extension wall) m) m). + unfold χ, the_priority_method.χ. + by rewrite wall_spec. } + lia. + enough (φ (χ m) e e m = φ (χ k) e e k) by lia. + assert (φ (χ m) e e m = B). + { rewrite <- (Hb m). reflexivity. lia. } + assert (φ (χ k) e e k = B). + { rewrite <- (Hb k). reflexivity. lia. } + lia. + enough (φ (χ m) e e m = φ (χ k) e e k) by lia. + assert (φ (χ m) e e m = B). + { rewrite <- (Hb m). reflexivity. lia. } + assert (φ (χ k) e e k = B). + { rewrite <- (Hb k). reflexivity. lia. } lia. - enough (φ (χ m) e e m = φ (χ k) e e k) by lia. - assert (φ (χ m) e e m = B). - { rewrite <- (Hb m). reflexivity. lia. } - assert (φ (χ k) e e k = B). - { rewrite <- (Hb k). reflexivity. lia. } - lia. - - intros H%Dec_true. - eapply F_with_top. - exists (F_func _ k), k; split; eauto. - eapply F_func_correctness. - Qed. + - intros H%Dec_true. + eapply F_with_top. + exists (F_func _ k), k; split; eauto. + eapply F_func_correctness. + Qed. End wall_greater_than_use. From 2dd4e442d180667704c7b3028e46a84814000fd2 Mon Sep 17 00:00:00 2001 From: Haoyi Date: Thu, 23 Jan 2025 19:44:06 +0100 Subject: [PATCH 54/54] fix error --- theories/PostsProblem/low_simple_predicates.v | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/theories/PostsProblem/low_simple_predicates.v b/theories/PostsProblem/low_simple_predicates.v index 663ea9d..4c40409 100644 --- a/theories/PostsProblem/low_simple_predicates.v +++ b/theories/PostsProblem/low_simple_predicates.v @@ -112,7 +112,8 @@ Section LowFacts. apply jump_P_limit; eauto. - eapply P_simple. intros. intros d. apply d. - apply wall_convergence. assumption. + apply wall_convergence. by unfold wall. + assumption. Qed. End LowSimplePredicate. @@ -130,10 +131,13 @@ Section LowFacts. Theorem PostProblem_from_neg_negLPO : ∃ p: nat → Prop, ¬ decidable p ∧ semi_decidable p ∧ (¬¬ (¬¬Σ⁰₁)-LEM → ¬ K ⪯ᴛ p). Proof. - eexists. + exists (P wall). repeat split. - apply simple_undecidable. - eapply P_simple. apply wall_convergence_classically. + eapply P_simple. intro e. + unfold lim_to. cbn. + apply wall_convergence_classically. + by unfold wall. - apply P_semi_decidable. - intros L. intros G. apply L. clear L. intros L. assert (~~ definite K) as hK.