From ab680b957ff8c4ec3ee11212cf0b258fe77126d5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Fri, 30 Aug 2024 12:00:44 +0200 Subject: [PATCH 01/12] Add a weight-preserving simplification during PropQuant computation --- theories/iSL/Order.v | 2 + theories/iSL/PropQuantifiers.v | 84 +++++++++++++++++++++++++--------- 2 files changed, 64 insertions(+), 22 deletions(-) diff --git a/theories/iSL/Order.v b/theories/iSL/Order.v index 1070bcf..2d9a5be 100644 --- a/theories/iSL/Order.v +++ b/theories/iSL/Order.v @@ -298,6 +298,8 @@ Ltac get_diff_form g := match g with | ?Γ ∖{[?φ]} => φ | _ (?Γ ∖{[?φ]}) => φ | _ (rm ?φ _) => φ +| (rm ?φ _) => φ +| ?Γ ++ _ => get_diff_form Γ | ?Γ • _ => get_diff_form Γ end. diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index 65de444..f91c7a6 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -32,11 +32,39 @@ Section PropQuantDefinition. Local Notation "Δ '•' φ" := (cons φ Δ). -(* solves the obligations of the following programs *) -Obligation Tactic := intros; order_tac. - Notation "□⁻¹ Γ" := (map open_box Γ) (at level 75). +(** TODO: Experimental: weight_preserving simplifications *) + +Definition isimp (φ : form) : list form := [φ]. +Opaque isimp. + +Lemma env_order_refl_isimp φ: env_order_refl (isimp φ) [φ]. +Proof. +Admitted. + + +Lemma isim_vars {φ ψ x}: occurs_in x ψ -> ψ ∈ isimp φ -> occurs_in x φ. +Proof. +Admitted. + +Lemma isimp_complete_L (Γ : env) φ θ: Γ ⊎ list_to_set_disj (isimp φ) ⊢ θ -> Γ ⊎ {[φ]} ⊢ θ. +Proof. Admitted. + +Ltac isimp_order := match goal with +| |- ?Δ ++ isimp ?φ ≺ _ => eapply env_order_le_lt_trans; +[apply env_order_refl_disj_union_compat ; [ isimp_order |apply (env_order_refl_isimp φ)]| +repeat rewrite <- Permutation_cons_append] +| |- env_order_refl (_ ++ isimp _) _ => +apply env_order_refl_disj_union_compat; [ isimp_order|apply env_order_refl_isimp] +| |- _ • _ ≺ _ => eapply env_order_le_lt_trans; [repeat apply env_order_eq_add; isimp_order| +repeat rewrite <- Permutation_cons_append] +| _ => right; reflexivity +end. + +(* solves the obligations of the following programs *) +Obligation Tactic := try solve[intros; order_tac; try isimp_order; order_tac]. + (** First, the implementation of the rules for calculating E. The names of the rules refer to the table in Pitts' paper. *) (** note the use of "lazy" conjunctions, disjunctions and implications *) @@ -52,23 +80,23 @@ match θ with if decide (p = q) then ⊤ (* default *) else q (* E1 modified *) (* E2 *) -| δ₁ ∧ δ₂ => E ((Δ'•δ₁)•δ₂) _ +| δ₁ ∧ δ₂ => E ((Δ' ++ isimp δ₁) ++ isimp δ₂) _ (* E3 *) -| δ₁ ∨ δ₂ => E (Δ'•δ₁) _ ⊻ E (Δ' •δ₂) _ +| δ₁ ∨ δ₂ => E (Δ'•δ₁) _ ⊻ E (Δ' ++ isimp δ₂) _ | Var q → δ => - if decide (Var q ∈ Δ) then E (Δ'•δ) _ (* E5 modified *) + if decide (Var q ∈ Δ) then E (Δ' ++ isimp δ) _ (* E5 modified *) else if decide (p = q) then ⊤ - else q ⇢ E (Δ'•δ) _ (* E4 *) + else q ⇢ E (Δ' ++ isimp δ) _ (* E4 *) (* E6 *) -| (δ₁ ∧ δ₂)→ δ₃ => E (Δ'•(δ₁ → (δ₂ → δ₃))) _ +| (δ₁ ∧ δ₂)→ δ₃ => E (Δ' ++ isimp (δ₁ → (δ₂ → δ₃))) _ (* E7 *) -| (δ₁ ∨ δ₂)→ δ₃ => E (Δ'•(δ₁ → δ₃)•(δ₂ → δ₃)) _ +| (δ₁ ∨ δ₂)→ δ₃ => E ((Δ' ++ isimp (δ₁ → δ₃)) ++ isimp (δ₂ → δ₃)) _ (* E8 modified *) | ((δ₁→ δ₂)→ δ₃) => - (E (Δ'•(δ₂ → δ₃) • δ₁) _⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) ⇢ E (Δ'•δ₃) _ + (E ((Δ' ++ isimp (δ₂ → δ₃)) ++ isimp δ₁) _⇢ A ((Δ' ++ isimp (δ₂ → δ₃)) ++ isimp δ₁, δ₂) _) ⇢ E (Δ' ++ isimp δ₃) _ | Bot → _ => ⊤ -| □ φ => □(E ((□⁻¹ Δ') • φ ) _) (* very redundant ; identical for each box *) -| (□δ1 → δ2) => (□(E((□⁻¹ Δ') • δ2 • □δ1) _ ⇢ A((□⁻¹ Δ') • δ2 • □δ1, δ1) _)) ⇢ E(Δ' • δ2) _ +| □ φ => □(E ((□⁻¹ Δ') ++ isimp φ ) _) (* very redundant ; identical for each box *) +| (□δ1 → δ2) => (□(E(((□⁻¹ Δ') ++ isimp δ2) ++ isimp (□δ1)) _ ⇢ A(((□⁻¹ Δ') ++ isimp δ2) ++ isimp (□δ1), δ1) _)) ⇢ E(Δ' ++ isimp δ2) _ end. @@ -89,29 +117,31 @@ match θ with else ⊥ else ⊥ (* A1 modified : A (Δ', ϕ) can be removed *) (* A2 *) -| δ₁ ∧ δ₂ => A ((Δ'•δ₁)•δ₂, ϕ) _ +| δ₁ ∧ δ₂ => A ((Δ' ++ isimp δ₁) ++ isimp δ₂, ϕ) _ (* A3 *) | δ₁ ∨ δ₂ => - (E (Δ'•δ₁) _ ⇢ A (Δ'•δ₁, ϕ) _) - ⊼ (E (Δ'•δ₂) _ ⇢ A (Δ'•δ₂, ϕ) _) + (E (Δ' ++ isimp δ₁) _ ⇢ A (Δ' ++ isimp δ₁, ϕ) _) + ⊼ (E (Δ' ++ isimp δ₂) _ ⇢ A (Δ' ++ isimp δ₂, ϕ) _) | Var q → δ => if decide (Var q ∈ Δ) then A (Δ'•δ, ϕ) _ (* A5 modified *) else if decide (p = q) then ⊥ - else q ⊼ A (Δ'•δ, ϕ) _ (* A4 *) + else q ⊼ A (Δ' ++ isimp δ, ϕ) _ (* A4 *) (* A6 *) | (δ₁ ∧ δ₂)→ δ₃ => - A (Δ'•(δ₁ → (δ₂ → δ₃)), ϕ) _ + A (Δ' ++ isimp (δ₁ → (δ₂ → δ₃)), ϕ) _ (* A7 *) | (δ₁ ∨ δ₂)→ δ₃ => - A ((Δ'•(δ₁ → δ₃))•(δ₂ → δ₃), ϕ) _ + A ((Δ' ++ isimp (δ₁ → δ₃)) ++ isimp (δ₂ → δ₃), ϕ) _ (* A8 modified*) | ((δ₁→ δ₂)→ δ₃) => - (E (Δ'•(δ₂ → δ₃) • δ₁) _ ⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) - ⊼ A (Δ'•δ₃, ϕ) _ + (E ((Δ' ++ isimp (δ₂ → δ₃)) ++ isimp δ₁) _ ⇢ A ((Δ' ++ isimp (δ₂ → δ₃) )++ isimp δ₁, δ₂) _) + ⊼ A (Δ' ++ isimp δ₃, ϕ) _ | Bot => ⊥ | Bot → _ => ⊥ | □δ => ⊥ -| ((□δ1) → δ2) => (□(E((□⁻¹ Δ')• δ2 • □δ1) _ ⇢ A((□⁻¹ Δ')• δ2 • □δ1, δ1) _)) ∧ A(Δ' • δ2, ϕ) _ +| ((□δ1) → δ2) => (□(E(((□⁻¹ Δ') ++ isimp δ2 ) ++ isimp (□δ1)) _ + ⇢ A(((□⁻¹ Δ') ++ isimp δ2 ) ++ isimp (□δ1), δ1) _)) + ∧ A(Δ' ++ isimp δ2, ϕ) _ (* using ⊼ here breaks congruence *) end. @@ -137,6 +167,7 @@ match ϕ with | Bot => ⊥ | □δ => □((E ((□⁻¹ Δ) • □δ, δ) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) end. +(* TODO: use isimp here too? *) Obligation Tactic := intros; order_tac. Program Fixpoint EA (pe : list form * form) {wf pointed_env_order pe} := @@ -262,6 +293,8 @@ try match goal with | H : occurs_in _ (?a ⇢ ?b) |- _ => apply occurs_in_make_impl in H | H : occurs_in _ (?a ⊻ ?b) |- _ => apply occurs_in_make_disj in H | H : occurs_in _ (?a ⊼ ?b) |- _ => apply occurs_in_make_conj in H +|H1 : ?x0 ∈ (?Δ ++ _), H2 : occurs_in ?x ?x0 |- _ => repeat rewrite elem_of_app in H1 +|H1 : ?x0 ∈ (isimp _), H2 : occurs_in ?x ?x0 |- _ => apply (isim_vars H2) in H1 |H1 : ?x0 ∈ (⊗ ?Δ), H2 : occurs_in ?x ?x0 |- _ => apply (occurs_in_open_boxes _ _ _ H2) in H1 |H1 : ?x0 ∈ (map open_box ?Δ), H2 : occurs_in ?x ?x0 |- _ => @@ -378,7 +411,14 @@ rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (equiv_disj_union_compat + subst. case_decide; subst; auto with proof. + exch 0... - constructor 2. -- simpl; exch 0. apply AndL. exch 1; exch 0. do 2 l_tac... +- simpl; exch 0. apply AndL. +(* requires a specific lemma using isimp_complete_L, equiv_assoc and list_to_set_disj_app *) +assert((list_to_set_disj (((rm (θ1 ∧ θ2) Δ) ++ isimp θ1) ++ isimp θ2) + • (EA0 ((rm (θ1 ∧ θ2) Δ ++ isimp θ1) ++ isimp θ2, ϕ) + (a_rule_env_obligation_1 Δ ϕ (θ1 ∧ θ2) Hin θ1 θ2 eq_refl)).2) ⊢ ϕ). + +auto with proof. + exch 1; exch 0. do 2 l_tac... - apply make_conj_sound_L. exch 0. apply OrL; exch 0. + apply AndL. apply make_impl_sound_L. exch 0. apply make_impl_sound_L. From 15dc02f17929a476906094e3980f5b2b81b1fac9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Tue, 3 Sep 2024 13:57:54 +0200 Subject: [PATCH 02/12] Decision procedure for iSL --- theories/iSL/DecisionProcedure.v | 321 +++++++++++++++++++++++++++++++ 1 file changed, 321 insertions(+) create mode 100644 theories/iSL/DecisionProcedure.v diff --git a/theories/iSL/DecisionProcedure.v b/theories/iSL/DecisionProcedure.v new file mode 100644 index 0000000..26e94e0 --- /dev/null +++ b/theories/iSL/DecisionProcedure.v @@ -0,0 +1,321 @@ +Require Import ISL.Sequents ISL.SequentProps ISL.Order. + +Global Instance proper_rm : Proper ((=) ==> (≡ₚ) ==> (≡ₚ)) rm. +Proof. +intros x y Heq. subst y. +induction 1; simpl; trivial. +- case form_eq_dec; auto with *. +- case form_eq_dec; auto with * ; + case form_eq_dec; auto with *. intros. apply Permutation_swap. +- now rewrite IHPermutation1. +Qed. + +Definition exists_dec {A : Type} (P : A -> bool) (l : list A): {x & (In x l) /\ P x} + {forall x, In x l -> ¬ P x}. +Proof. +induction l as [|x l]. +- right. tauto. +- case_eq (P x); intro Heq. + + left. exists x. split; auto with *. + + destruct IHl as [(y & Hin & Hy)|Hf]. + * left. exists y. split; auto with *. + * right. simpl. intros z [Hz|Hz]; subst; try rewrite Heq; auto with *. +Defined. + +Proposition Provable_dec Γ φ : + {_ : list_to_set_disj Γ ⊢ φ & True} + {forall H : list_to_set_disj Γ ⊢ φ, False}. +Proof. +(* duplicate *) +Ltac l_tac := repeat rewrite list_to_set_disj_open_boxes; + rewrite (proper_Provable _ _ (list_to_set_disj_env_add _ _) _ _ eq_refl) +|| rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_env_add _ _)) _ _ eq_refl) +|| rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (equiv_disj_union_compat_r (list_to_set_disj_env_add _ _))) _ _ eq_refl) +|| rewrite (proper_Provable _ _ (equiv_disj_union_compat_r(equiv_disj_union_compat_r (equiv_disj_union_compat_r (list_to_set_disj_env_add _ _)))) _ _ eq_refl). +remember (Γ, φ) as pe. +replace Γ with pe.1 by now inversion Heqpe. +replace φ with pe.2 by now inversion Heqpe. clear Heqpe Γ φ. +revert pe. +refine (@well_founded_induction _ _ wf_pointed_order _ _). +intros (Γ& φ) Hind; simpl. +assert(Hind' := λ Γ' φ', Hind(Γ', φ')). simpl in Hind'. clear Hind. rename Hind' into Hind. + +case (decide (⊥ ∈ Γ)); intro Hbot. +{ left. eexists; trivial. apply elem_of_list_to_set_disj in Hbot. exhibit Hbot 0. apply ExFalso. } +assert(HAndR : {φ1 & {φ2 & φ = (And φ1 φ2)}} + {∀ φ1 φ2, φ ≠ (And φ1 φ2)}) by (destruct φ; eauto). +destruct HAndR as [(φ1 & φ2 & Heq) | HAndR]. +{ subst. + destruct (Hind Γ φ1) as [(Hp1&_) | H1]. order_tac. + - destruct (Hind Γ φ2) as [(Hp2&_) | H2]. order_tac. + + left. eexists; trivial. apply AndR; assumption. + + right. intro Hp. apply AndR_rev in Hp. tauto. + - right. intro Hp. apply AndR_rev in Hp. tauto. +} +assert(Hvar : {p & φ = Var p & Var p ∈ Γ} + {∀ p, φ = Var p -> Var p ∈ Γ -> False}). { + destruct φ. 2-6: right; auto with *. + case (decide (Var v ∈ Γ)); intro Hin. + - left. exists v; trivial. + - right; auto with *. } +destruct Hvar as [[p Heq Hp]|Hvar]. +{ subst. left. eexists; trivial. apply elem_of_list_to_set_disj in Hp. exhibit Hp 0. apply Atom. } +assert(HAndL : {ψ1 & {ψ2 & (And ψ1 ψ2) ∈ Γ}} + {∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False}). { + pose (fA := (fun θ => match θ with |And _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 3: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +} +destruct HAndL as [(ψ1 & ψ2 & Hin)|HAndL]. +{ destruct (Hind (ψ1 :: ψ2 :: rm (And ψ1 ψ2) Γ) φ) as [[Hp' _] | Hf]. order_tac. + - left. eexists; trivial. apply elem_of_list_to_set_disj in Hin. + exhibit Hin 0. + rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_rm _ _)) _ _ eq_refl). + apply AndL. peapply Hp'. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add (ψ2 :: rm (And ψ1 ψ2) Γ) ψ1)) 0. + rw (symmetry (list_to_set_disj_env_add (rm (And ψ1 ψ2) Γ) ψ2)) 1. + exch 0. apply AndL_rev. + rw (symmetry (list_to_set_disj_rm Γ(And ψ1 ψ2))) 1. + apply elem_of_list_to_set_disj in Hin. + pose (difference_singleton (list_to_set_disj Γ) (And ψ1 ψ2)). + peapply Hf'. +} +assert(HImpR : {φ1 & {φ2 & φ = (Implies φ1 φ2)}} + {∀ φ1 φ2, φ ≠ (Implies φ1 φ2)}) by (destruct φ; eauto). +destruct HImpR as [(φ1 & φ2 & Heq) | HImpR]. +{ subst. + destruct (Hind (φ1 :: Γ) φ2) as [(Hp1&_) | H1]. order_tac. + - left. eexists; trivial. apply ImpR. peapply Hp1. + - right. intro Hf. apply H1. apply ImpR_rev in Hf. peapply Hf. +} +assert(HOrL : {ψ1 & {ψ2 & (Or ψ1 ψ2) ∈ Γ}} + {∀ ψ1 ψ2, (Or ψ1 ψ2) ∈ Γ -> False}). { + pose (fA := (fun θ => match θ with |Or _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 4: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +} +destruct HOrL as [(ψ1 & ψ2 & Hin)|HOrL]. +{ apply elem_of_list_to_set_disj in Hin. + destruct (Hind (ψ1 :: rm (Or ψ1 ψ2) Γ) φ) as [[Hp1 _] | Hf]. order_tac. + - destruct (Hind (ψ2 :: rm (Or ψ1 ψ2) Γ) φ) as [[Hp2 _] | Hf]. order_tac. + + left. eexists; trivial. exhibit Hin 0. + rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_rm _ _)) _ _ eq_refl). + apply OrL. peapply Hp1. peapply Hp2. + + right; intro Hf'. assert(Hf'' :list_to_set_disj (rm (Or ψ1 ψ2) Γ) • Or ψ1 ψ2 ⊢ φ). { + rw (symmetry (list_to_set_disj_rm Γ(Or ψ1 ψ2))) 1. + pose (difference_singleton (list_to_set_disj Γ) (Or ψ1 ψ2)). peapply Hf'. + } + apply OrL_rev in Hf''. apply Hf. peapply Hf''. + - right; intro Hf'. assert(Hf'' :list_to_set_disj (rm (Or ψ1 ψ2) Γ) • Or ψ1 ψ2 ⊢ φ). { + rw (symmetry (list_to_set_disj_rm Γ(Or ψ1 ψ2))) 1. + pose (difference_singleton (list_to_set_disj Γ) (Or ψ1 ψ2)). peapply Hf'. + } + apply OrL_rev in Hf''. apply Hf. peapply Hf''.1. +} +assert(HImpLVar : {p & {ψ & Var p ∈ Γ /\ (Implies (Var p) ψ) ∈ Γ}} + + {∀ p ψ, Var p ∈ Γ -> (Implies (Var p) ψ) ∈ Γ -> False}). { + pose (fIp :=λ p θ, match θ with | Implies (Var q) _ => if decide (p = q) then true else false | _ => false end). + pose (fp:= (fun θ => match θ with |Var p => if (exists_dec (fIp p) Γ) then true else false | _ => false end)). + destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fp. destruct θ. 2-6: auto with *. + case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. + unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. + destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. + subst. apply elem_of_list_In in Hinψ, Hin. + do 2 eexists. split; eauto. + - right. intros p ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. + simpl in Hp. case exists_dec as [|Hf'] in Hp. auto with *. + apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. +} +destruct HImpLVar as [[p [ψ [Hinp Hinψ]]]|HImpLVar]. +{ apply elem_of_list_to_set_disj in Hinp. + apply elem_of_list_to_set_disj in Hinψ. + assert(Hinp' : Var p ∈ (list_to_set_disj Γ ∖ {[Implies p ψ]} : env)) + by (apply in_difference; [discriminate| assumption]). + destruct (Hind (ψ :: rm (Implies (Var p) ψ) Γ) φ) as [[Hp _]|Hf]. order_tac. + - left. eexists; trivial. exhibit Hinψ 0. + exhibit Hinp' 1. apply ImpLVar. + rw (symmetry (difference_singleton (list_to_set_disj Γ ∖ {[Implies p ψ]}) (Var p) Hinp')) 1. + rw (list_to_set_disj_rm Γ(Implies p ψ)) 1. l_tac. exact Hp. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add (rm (Implies p ψ) Γ) ψ)) 0. + rw (symmetry (list_to_set_disj_rm Γ(Implies p ψ))) 1. + exhibit Hinp' 1. apply ImpLVar_rev. + rw (symmetry (difference_singleton _ _ Hinp')) 1. + rw (symmetry (difference_singleton _ _ Hinψ)) 0. + exact Hf'. +} +assert(HImpLAnd : {φ1 & {φ2 & {φ3 & (Implies (And φ1 φ2) φ3) ∈ Γ}}} + + {∀ φ1 φ2 φ3, (Implies (And φ1 φ2) φ3) ∈ Γ -> False}). { + pose (fII := (fun θ => match θ with |Implies (And _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-2,4-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +} +destruct HImpLAnd as [(φ1&φ2&φ3&Hin)|HImpLAnd]. +{ apply elem_of_list_to_set_disj in Hin. + destruct (Hind (Implies φ1 (Implies φ2 φ3) :: rm (Implies (And φ1 φ2) φ3) Γ) φ) as [[Hp _]|Hf]. order_tac. + - left. eexists; trivial. exhibit Hin 0. apply ImpLAnd. + rw (list_to_set_disj_rm Γ(Implies (And φ1 φ2) φ3)) 1. l_tac. exact Hp. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add (rm (Implies (And φ1 φ2) φ3) Γ) (Implies φ1 (Implies φ2 φ3)))) 0. + rw (symmetry (list_to_set_disj_rm Γ(Implies (And φ1 φ2) φ3))) 1. + apply ImpLAnd_rev. + rw (symmetry (difference_singleton _ _ Hin)) 0. exact Hf'. +} +assert(HImpLOr : {φ1 & {φ2 & {φ3 & (Implies (Or φ1 φ2) φ3) ∈ Γ}}} + + {∀ φ1 φ2 φ3, (Implies (Or φ1 φ2) φ3) ∈ Γ -> False}). { + pose (fII := (fun θ => match θ with |Implies (Or _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-3, 5-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +} +destruct HImpLOr as [(φ1&φ2&φ3&Hin)|HImpLOr]. +{ apply elem_of_list_to_set_disj in Hin. + destruct (Hind (Implies φ2 φ3 :: Implies φ1 φ3 :: rm (Implies (Or φ1 φ2) φ3) Γ) φ) as [[Hp _]|Hf]. order_tac. + - left. eexists; trivial. exhibit Hin 0. apply ImpLOr. + rw (list_to_set_disj_rm Γ(Implies (Or φ1 φ2) φ3)) 2. do 2 l_tac. exact Hp. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add ( Implies φ1 φ3 :: rm (Implies (Or φ1 φ2) φ3) Γ) (Implies φ2 φ3))) 0. + rw (symmetry (list_to_set_disj_env_add (rm (Implies (Or φ1 φ2) φ3) Γ) (Implies φ1 φ3))) 1. + rw (symmetry (list_to_set_disj_rm Γ(Implies (Or φ1 φ2) φ3))) 2. + apply ImpLOr_rev. + rw (symmetry (difference_singleton _ _ Hin)) 0. exact Hf'. +} +(* non invertible right rules *) +assert(HOrR1 : {φ1 & {φ2 & {Hp : (list_to_set_disj Γ ⊢ φ1) & φ = (Or φ1 φ2)}}} + + {∀ φ1 φ2, ∀ (H : list_to_set_disj Γ ⊢ φ1), φ = (Or φ1 φ2) -> False}). +{ + destruct φ. 4: { destruct (Hind Γ φ1)as [[Hp _]|Hf]. order_tac. + - left. do 3 eexists; eauto. + - right. intros ? ? Hp Heq. inversion Heq. subst. tauto. + } + all: right; auto with *. +} +destruct HOrR1 as [(φ1 & φ2 & Hp & Heq)| HOrR1]. +{ subst. left. eexists; trivial. apply OrR1, Hp. } +assert(HOrR2 : {φ1 & {φ2 & {Hp : (list_to_set_disj Γ ⊢ φ2) & φ = (Or φ1 φ2)}}} + + {∀ φ1 φ2, ∀ (H : list_to_set_disj Γ ⊢ φ2), φ = (Or φ1 φ2) -> False}). +{ + destruct φ. 4: { destruct (Hind Γ φ2)as [[Hp _]|Hf]. order_tac. + - left. do 3 eexists; eauto. + - right. intros ? ? Hp Heq. inversion Heq. subst. tauto. + } + all: right; auto with *. +} +destruct HOrR2 as [(φ1 & φ2 & Hp & Heq)| HOrR2 ]. +{ subst. left. eexists; trivial. apply OrR2, Hp. } +assert(HBoxR : {φ' & {Hp : (⊗ (list_to_set_disj Γ) • □ φ' ⊢ φ') & φ = (□ φ')}} + + {∀ φ', ∀ (H : ⊗ (list_to_set_disj Γ) • □ φ' ⊢ φ'), φ = (□ φ') -> False}). +{ + destruct φ. 6: { destruct (Hind ((□ φ) :: map open_box Γ) φ)as [[Hp _]|Hf]. order_tac. + - left. do 2 eexists; eauto. l_tac. exact Hp. + - right. intros ? Hp Heq. inversion Heq. subst. apply Hf. + rw (symmetry (list_to_set_disj_env_add (map open_box Γ) (□ φ'))) 0. + rewrite <- list_to_set_disj_open_boxes. exact Hp. + } + all: right; auto with *. +} +destruct HBoxR as [(φ' & Hp & Heq)| HBoxR ]. +{ left. subst. eexists; trivial. apply BoxR, Hp. } +assert(Hempty: ∀ (Δ : env) φ,((Δ • φ) = ∅) -> False). +{ + intros Δ θ Heq. assert (Hm:= multiplicity_empty θ). + unfold base.empty in *. + rewrite <- Heq, union_mult, singleton_mult_in in Hm by trivial. lia. +} +(* non invertible left rules *) +assert(HImpLImp : ∀Γ2 Γ1, Γ1 ++ Γ2 = Γ -> {φ1 & {φ2 & {φ3 &{H2312 : ((list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • (Implies φ2 φ3)) ⊢ (Implies φ1 φ2)) + & {H3: (list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • φ3 ⊢ φ) & (Implies (Implies φ1 φ2) φ3) ∈ Γ2}}}}} + + {∀ φ1 φ2 φ3 (_ : (list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • (Implies φ2 φ3)) ⊢ (Implies φ1 φ2)) + (_: list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • φ3 ⊢ φ), + Implies (Implies φ1 φ2) φ3 ∈ Γ2 → False}). +{ + induction Γ2 as [|θ Γ2]; intros Γ1 Heq. + - right. intros φ1 φ2 φ3 _ _ Hin. inversion Hin. + - assert(Heq' : (Γ1 ++ [θ]) ++ Γ2 = Γ) by (subst; auto with *). + destruct (IHΓ2 (Γ1 ++ [θ]) Heq') as [(φ1 & φ2 & φ3 & Hp1 & Hp2 & Hin)|Hf]. + + left. repeat eexists; eauto. now right. + + destruct θ. + 5: destruct θ1. + 9 : { + destruct (Hind (Implies θ1_2 θ2 :: rm (Implies (Implies θ1_1 θ1_2) θ2) Γ) (Implies θ1_1 θ1_2)) + as [[Hp1 _] | Hf']. + - order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + - destruct (Hind (θ2 :: rm (Implies (Implies θ1_1 θ1_2) θ2) Γ) φ) as [[Hp2 _] | Hf'']. + + order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + + left. repeat eexists; try l_tac; eauto. ms. + + right; intros φ1 φ2 φ3 Hp1' Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf''. peapply Hp2. + - right; intros φ1 φ2 φ3 Hp1 Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf'. peapply Hp1. + } + all: (right; intros φ1 φ2 φ3 Hp1 Hp2 He; apply elem_of_list_In in He; destruct He as [Heq''| Hin]; + [discriminate|apply elem_of_list_In in Hin; eapply Hf; eauto]). +} +destruct (HImpLImp Γ [] (app_nil_l _)) as [(φ1 & φ2 & φ3 & Hp1 & Hp2 & Hin)|HfImpl]. +{ apply elem_of_list_to_set_disj in Hin. + left. eexists; trivial. exhibit Hin 0. rw (list_to_set_disj_rm Γ(Implies (Implies φ1 φ2) φ3)) 1. + apply ImpLImp; assumption. +} +(* ImpBox *) +assert(HImpLBox : ∀Γ2 Γ1, Γ1 ++ Γ2 = Γ -> {φ1 & {φ2 & {H2312 : ((⊗(list_to_set_disj ((rm (Implies (□ φ1) φ2) Γ))) • □ φ1 • φ2) ⊢φ1) + & {H3: (list_to_set_disj (rm (Implies (□ φ1) φ2) Γ) • φ2 ⊢ φ) & (Implies (□ φ1) φ2) ∈ Γ2}}}} + + {∀ φ1 φ2 (_ : ((⊗ (list_to_set_disj ((rm (Implies (□ φ1) φ2) Γ))) • □ φ1 • φ2) ⊢φ1)) + (_: list_to_set_disj (rm (Implies (□ φ1) φ2) Γ) • φ2 ⊢ φ), + Implies (□ φ1) φ2 ∈ Γ2 → False}). +{ + induction Γ2 as [|θ Γ2]; intros Γ1 Heq. + - right. intros φ1 φ2 _ _ Hin. inversion Hin. + - assert(Heq' : (Γ1 ++ [θ]) ++ Γ2 = Γ) by (subst; auto with *). + destruct (IHΓ2 (Γ1 ++ [θ]) Heq') as [(φ1 & φ2 & Hp1 & Hp2 & Hin)|Hf]. + + left. repeat eexists; eauto. now right. + + destruct θ. + 5: destruct θ1. + 10 : { + destruct (Hind (θ2 :: (□θ1) :: map open_box (rm (Implies (□ θ1) θ2) Γ)) θ1) + as [[Hp1 _] | Hf']. + - order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + - destruct (Hind (θ2 :: rm (Implies (□ θ1) θ2) Γ) φ) as [[Hp2 _] | Hf'']. + + order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + + left. repeat eexists; repeat l_tac; eauto. ms. + + right; intros φ1 φ2 Hp1' Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf''. peapply Hp2. + - right; intros φ1 φ2 Hp1 Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf'. + (erewrite proper_Provable; [| |reflexivity]); [eapply Hp1|]. + repeat rewrite <- ?list_to_set_disj_env_add, list_to_set_disj_open_boxes. trivial. + } + all: (right; intros φ1 φ2 Hp1 Hp2 He; apply elem_of_list_In in He; destruct He as [Heq''| Hin]; + [discriminate|apply elem_of_list_In in Hin; eapply Hf; eauto]). +} +destruct (HImpLBox Γ [] (app_nil_l _)) as [(φ1 & φ2 & Hp1 & Hp2 & Hin)|HfImpLBox]. +{ apply elem_of_list_to_set_disj in Hin. + left. eexists; trivial. exhibit Hin 0. rw (list_to_set_disj_rm Γ(Implies (□ φ1) φ2)) 1. + apply ImpBox; assumption. +} +clear Hind HImpLImp HImpLBox. +right. +Ltac eqt := match goal with | H : (_ • ?φ) = list_to_set_disj ?Γ |- _ => + let Heq := fresh "Heq" in assert(Heq := H); + assert(Hinφ : φ ∈ Γ) by (apply elem_of_list_to_set_disj; setoid_rewrite <- H; ms); + apply env_equiv_eq, env_add_inv', symmetry in Heq; rewrite list_to_set_disj_rm in Heq end. +intro Hp. inversion Hp; subst; try eqt; eauto 2. +- eapply HAndR; eauto. +- eapply HImpR; eauto. +- eapply HImpLVar; eauto. apply elem_of_list_to_set_disj. setoid_rewrite <- H; ms. +- eapply HfImpl; eauto. + + now rw Heq 1. + + now rw Heq 1. +- eapply HfImpLBox; eauto. + + now rw (proper_open_boxes _ _ Heq) 2. + + now rw Heq 1. +Defined. From 0e5b662bbcf4b76e53d046dfbd7d1e72e425f0a3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Tue, 3 Sep 2024 14:57:56 +0200 Subject: [PATCH 03/12] Extract the iSL decision procedure to OCaml --- bin/dune | 4 +- bin/isl_dec.ml | 27 +++ theories/extraction/UIML_extraction.v | 4 +- theories/iSL/DecisionProcedure.v | 295 ++++++++++++++++++++++++++ 4 files changed, 326 insertions(+), 4 deletions(-) create mode 100644 bin/isl_dec.ml diff --git a/bin/dune b/bin/dune index 4952ca8..a8987a2 100644 --- a/bin/dune +++ b/bin/dune @@ -1,6 +1,6 @@ (executables - (public_names uiml_demo benchmark uiml_cmdline) - (names uiml_demo benchmark uiml_cmdline) + (public_names uiml_demo benchmark uiml_cmdline isl_dec) + (names uiml_demo benchmark uiml_cmdline isl_dec) (modes js native) (preprocess (pps js_of_ocaml-ppx)) (libraries js_of_ocaml UIML angstrom exenum) diff --git a/bin/isl_dec.ml b/bin/isl_dec.ml new file mode 100644 index 0000000..d33a00e --- /dev/null +++ b/bin/isl_dec.ml @@ -0,0 +1,27 @@ +open Exenum +open Printer +open UIML.Formulas +open Sys +open Char +open UIML.DecisionProcedure +open UIML.Datatypes +open Stringconversion +open Modal_expressions_parser + + + +let nb_args = Array.length Sys.argv + +let form = if nb_args = 2 then (Sys.argv.(1)) else "T" + +let usage_string = + "isl_dec φ: decides the provability of the modal formula φ in iSL.\n" + +let print_decision = function + | Coq_inl _ -> "Probable" + | _ -> "Not provable" + +let () = + if nb_args = 2 then form |> eval |> coq_Provable_dec' [] |> + print_decision |> print_string |> print_newline + else print_string usage_string diff --git a/theories/extraction/UIML_extraction.v b/theories/extraction/UIML_extraction.v index 231e955..b449651 100644 --- a/theories/extraction/UIML_extraction.v +++ b/theories/extraction/UIML_extraction.v @@ -5,7 +5,7 @@ Require Import ExtrOcamlBasic ExtrOcamlString. Require Import K.Interpolation.UIK_braga. Require Import KS_export. -Require Import ISL.PropQuantifiers. +Require Import ISL.PropQuantifiers ISL.DecisionProcedure. Require Import ISL.Simp. @@ -41,5 +41,5 @@ Definition isl_simplified_A v f := A_simplified v f. Set Extraction Output Directory "extraction". -Separate Extraction gl_UI k_UI isl_E isl_A isl_simplified_E isl_simplified_A Formulas.weight isl_simp. +Separate Extraction Provable_dec' gl_UI k_UI isl_E isl_A isl_simplified_E isl_simplified_A Formulas.weight isl_simp. diff --git a/theories/iSL/DecisionProcedure.v b/theories/iSL/DecisionProcedure.v index 26e94e0..de48c30 100644 --- a/theories/iSL/DecisionProcedure.v +++ b/theories/iSL/DecisionProcedure.v @@ -319,3 +319,298 @@ intro Hp. inversion Hp; subst; try eqt; eauto 2. + now rw (proper_open_boxes _ _ Heq) 2. + now rw Heq 1. Defined. + + +Proposition Provable_dec' Γ φ : + (exists _ : list_to_set_disj Γ ⊢ φ, True) + (forall H : list_to_set_disj Γ ⊢ φ, False). +Proof. +remember (Γ, φ) as pe. +replace Γ with pe.1 by now inversion Heqpe. +replace φ with pe.2 by now inversion Heqpe. clear Heqpe Γ φ. +revert pe. +refine (@well_founded_induction _ _ wf_pointed_order _ _). +intros (Γ& φ) Hind; simpl. +assert(Hind' := λ Γ' φ', Hind(Γ', φ')). simpl in Hind'. clear Hind. rename Hind' into Hind. + +case (decide (⊥ ∈ Γ)); intro Hbot. +{ left. eexists; trivial. apply elem_of_list_to_set_disj in Hbot. exhibit Hbot 0. apply ExFalso. } +assert(HAndR : {φ1 & {φ2 & φ = (And φ1 φ2)}} + {∀ φ1 φ2, φ ≠ (And φ1 φ2)}) by (destruct φ; eauto). +destruct HAndR as [(φ1 & φ2 & Heq) | HAndR]. +{ subst. + destruct (Hind Γ φ1) as [Hp1| H1]. order_tac. + - destruct (Hind Γ φ2) as [Hp2| H2]. order_tac. + + left. destruct Hp1, Hp2. eexists; trivial. apply AndR; assumption. + + right. intro Hp. apply AndR_rev in Hp. tauto. + - right. intro Hp. apply AndR_rev in Hp. tauto. +} +assert(Hvar : {p & φ = Var p & Var p ∈ Γ} + {∀ p, φ = Var p -> Var p ∈ Γ -> False}). { + destruct φ. 2-6: right; auto with *. + case (decide (Var v ∈ Γ)); intro Hin. + - left. exists v; trivial. + - right; auto with *. } +destruct Hvar as [[p Heq Hp]|Hvar]. +{ subst. left. eexists; trivial. apply elem_of_list_to_set_disj in Hp. exhibit Hp 0. apply Atom. } +assert(HAndL : {ψ1 & {ψ2 & (And ψ1 ψ2) ∈ Γ}} + {∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False}). { + pose (fA := (fun θ => match θ with |And _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 3: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +} +destruct HAndL as [(ψ1 & ψ2 & Hin)|HAndL]. +{ destruct (Hind (ψ1 :: ψ2 :: rm (And ψ1 ψ2) Γ) φ) as [Hp' | Hf]. order_tac. + - left. destruct Hp' as [Hp' _]. eexists; trivial. apply elem_of_list_to_set_disj in Hin. + exhibit Hin 0. + rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_rm _ _)) _ _ eq_refl). + apply AndL. peapply Hp'. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add (ψ2 :: rm (And ψ1 ψ2) Γ) ψ1)) 0. + rw (symmetry (list_to_set_disj_env_add (rm (And ψ1 ψ2) Γ) ψ2)) 1. + exch 0. apply AndL_rev. + rw (symmetry (list_to_set_disj_rm Γ(And ψ1 ψ2))) 1. + apply elem_of_list_to_set_disj in Hin. + pose (difference_singleton (list_to_set_disj Γ) (And ψ1 ψ2)). + peapply Hf'. +} +assert(HImpR : {φ1 & {φ2 & φ = (Implies φ1 φ2)}} + {∀ φ1 φ2, φ ≠ (Implies φ1 φ2)}) by (destruct φ; eauto). +destruct HImpR as [(φ1 & φ2 & Heq) | HImpR]. +{ subst. + destruct (Hind (φ1 :: Γ) φ2) as [Hp1| H1]. order_tac. + - left. destruct Hp1 as [Hp1 _]. eexists; trivial. apply ImpR. peapply Hp1. + - right. intro Hf. apply H1. apply ImpR_rev in Hf. peapply Hf. +} +assert(HOrL : {ψ1 & {ψ2 & (Or ψ1 ψ2) ∈ Γ}} + {∀ ψ1 ψ2, (Or ψ1 ψ2) ∈ Γ -> False}). { + pose (fA := (fun θ => match θ with |Or _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 4: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +} +destruct HOrL as [(ψ1 & ψ2 & Hin)|HOrL]. +{ apply elem_of_list_to_set_disj in Hin. + destruct (Hind (ψ1 :: rm (Or ψ1 ψ2) Γ) φ) as [Hp1| Hf]. order_tac. + - destruct (Hind (ψ2 :: rm (Or ψ1 ψ2) Γ) φ) as [Hp2| Hf]. order_tac. + + left. destruct Hp1 as [Hp1 _]. destruct Hp2 as [Hp2 _]. eexists; trivial. exhibit Hin 0. + rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_rm _ _)) _ _ eq_refl). + apply OrL. peapply Hp1. peapply Hp2. + + right; intro Hf'. assert(Hf'' :list_to_set_disj (rm (Or ψ1 ψ2) Γ) • Or ψ1 ψ2 ⊢ φ). { + rw (symmetry (list_to_set_disj_rm Γ(Or ψ1 ψ2))) 1. + pose (difference_singleton (list_to_set_disj Γ) (Or ψ1 ψ2)). peapply Hf'. + } + apply OrL_rev in Hf''. apply Hf. peapply Hf''. + - right; intro Hf'. assert(Hf'' :list_to_set_disj (rm (Or ψ1 ψ2) Γ) • Or ψ1 ψ2 ⊢ φ). { + rw (symmetry (list_to_set_disj_rm Γ(Or ψ1 ψ2))) 1. + pose (difference_singleton (list_to_set_disj Γ) (Or ψ1 ψ2)). peapply Hf'. + } + apply OrL_rev in Hf''. apply Hf. peapply Hf''.1. +} +assert(HImpLVar : {p & {ψ & Var p ∈ Γ /\ (Implies (Var p) ψ) ∈ Γ}} + + {∀ p ψ, Var p ∈ Γ -> (Implies (Var p) ψ) ∈ Γ -> False}). { + pose (fIp :=λ p θ, match θ with | Implies (Var q) _ => if decide (p = q) then true else false | _ => false end). + pose (fp:= (fun θ => match θ with |Var p => if (exists_dec (fIp p) Γ) then true else false | _ => false end)). + destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fp. destruct θ. 2-6: auto with *. + case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. + unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. + destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. + subst. apply elem_of_list_In in Hinψ, Hin. + do 2 eexists. split; eauto. + - right. intros p ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. + simpl in Hp. case exists_dec as [|Hf'] in Hp. auto with *. + apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. +} +destruct HImpLVar as [[p [ψ [Hinp Hinψ]]]|HImpLVar]. +{ apply elem_of_list_to_set_disj in Hinp. + apply elem_of_list_to_set_disj in Hinψ. + assert(Hinp' : Var p ∈ (list_to_set_disj Γ ∖ {[Implies p ψ]} : env)) + by (apply in_difference; [discriminate| assumption]). + destruct (Hind (ψ :: rm (Implies (Var p) ψ) Γ) φ) as [Hp|Hf]. order_tac. + - left. destruct Hp as [Hp _]. eexists; trivial. exhibit Hinψ 0. + exhibit Hinp' 1. apply ImpLVar. + rw (symmetry (difference_singleton (list_to_set_disj Γ ∖ {[Implies p ψ]}) (Var p) Hinp')) 1. + rw (list_to_set_disj_rm Γ(Implies p ψ)) 1. l_tac. exact Hp. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add (rm (Implies p ψ) Γ) ψ)) 0. + rw (symmetry (list_to_set_disj_rm Γ(Implies p ψ))) 1. + exhibit Hinp' 1. apply ImpLVar_rev. + rw (symmetry (difference_singleton _ _ Hinp')) 1. + rw (symmetry (difference_singleton _ _ Hinψ)) 0. + exact Hf'. +} +assert(HImpLAnd : {φ1 & {φ2 & {φ3 & (Implies (And φ1 φ2) φ3) ∈ Γ}}} + + {∀ φ1 φ2 φ3, (Implies (And φ1 φ2) φ3) ∈ Γ -> False}). { + pose (fII := (fun θ => match θ with |Implies (And _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-2,4-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +} +destruct HImpLAnd as [(φ1&φ2&φ3&Hin)|HImpLAnd]. +{ apply elem_of_list_to_set_disj in Hin. + destruct (Hind (Implies φ1 (Implies φ2 φ3) :: rm (Implies (And φ1 φ2) φ3) Γ) φ) as [Hp|Hf]. order_tac. + - left. destruct Hp as [Hp _]. eexists; trivial. exhibit Hin 0. apply ImpLAnd. + rw (list_to_set_disj_rm Γ(Implies (And φ1 φ2) φ3)) 1. l_tac. exact Hp. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add (rm (Implies (And φ1 φ2) φ3) Γ) (Implies φ1 (Implies φ2 φ3)))) 0. + rw (symmetry (list_to_set_disj_rm Γ(Implies (And φ1 φ2) φ3))) 1. + apply ImpLAnd_rev. + rw (symmetry (difference_singleton _ _ Hin)) 0. exact Hf'. +} +assert(HImpLOr : {φ1 & {φ2 & {φ3 & (Implies (Or φ1 φ2) φ3) ∈ Γ}}} + + {∀ φ1 φ2 φ3, (Implies (Or φ1 φ2) φ3) ∈ Γ -> False}). { + pose (fII := (fun θ => match θ with |Implies (Or _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-3, 5-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +} +destruct HImpLOr as [(φ1&φ2&φ3&Hin)|HImpLOr]. +{ apply elem_of_list_to_set_disj in Hin. + destruct (Hind (Implies φ2 φ3 :: Implies φ1 φ3 :: rm (Implies (Or φ1 φ2) φ3) Γ) φ) as [Hp|Hf]. order_tac. + - left. destruct Hp as [Hp _]. eexists; trivial. exhibit Hin 0. apply ImpLOr. + rw (list_to_set_disj_rm Γ(Implies (Or φ1 φ2) φ3)) 2. do 2 l_tac. exact Hp. + - right. intro Hf'. apply Hf. + rw (symmetry (list_to_set_disj_env_add ( Implies φ1 φ3 :: rm (Implies (Or φ1 φ2) φ3) Γ) (Implies φ2 φ3))) 0. + rw (symmetry (list_to_set_disj_env_add (rm (Implies (Or φ1 φ2) φ3) Γ) (Implies φ1 φ3))) 1. + rw (symmetry (list_to_set_disj_rm Γ(Implies (Or φ1 φ2) φ3))) 2. + apply ImpLOr_rev. + rw (symmetry (difference_singleton _ _ Hin)) 0. exact Hf'. +} +(* non invertible right rules *) +assert(HOrR1 : {φ1 & {φ2 & (exists (_ : list_to_set_disj Γ ⊢ φ1), φ = (Or φ1 φ2))}} + + {∀ φ1 φ2, ∀ (H : list_to_set_disj Γ ⊢ φ1), φ = (Or φ1 φ2) -> False}). +{ + destruct φ. 4: { destruct (Hind Γ φ1)as [Hp|Hf]. order_tac. + - left. do 2 eexists. destruct Hp as [Hp _]. eexists; eauto. + - right. intros ? ? Hp Heq. inversion Heq. subst. tauto. + } + all: right; auto with *. +} +destruct HOrR1 as [(φ1 & φ2 & Hp)| HOrR1]. +{ left. destruct Hp as (Hp & Heq). subst. eexists; trivial. apply OrR1, Hp. } +assert(HOrR2 : {φ1 & {φ2 & exists (_ : list_to_set_disj Γ ⊢ φ2), φ = (Or φ1 φ2)}} + + {∀ φ1 φ2, ∀ (H : list_to_set_disj Γ ⊢ φ2), φ = (Or φ1 φ2) -> False}). +{ + destruct φ. 4: { destruct (Hind Γ φ2)as [Hp|Hf]. order_tac. + - left. do 2 eexists. destruct Hp as [Hp _]; eauto. + - right. intros ? ? Hp Heq. inversion Heq. subst. tauto. + } + all: right; auto with *. +} +destruct HOrR2 as [(φ1 & φ2 & Hp)| HOrR2 ]. +{ left. destruct Hp as [Hp Heq]. subst. eexists; trivial. apply OrR2, Hp. } +assert(HBoxR : {φ' & exists (_ : (⊗ (list_to_set_disj Γ) • □ φ' ⊢ φ')), φ = (□ φ')} + + {∀ φ', ∀ (H : ⊗ (list_to_set_disj Γ) • □ φ' ⊢ φ'), φ = (□ φ') -> False}). +{ + destruct φ. 6: { destruct (Hind ((□ φ) :: map open_box Γ) φ)as [Hp|Hf]. order_tac. + - left. eexists. destruct Hp as [Hp _]. eexists; eauto. l_tac. exact Hp. + - right. intros ? Hp Heq. inversion Heq. subst. apply Hf. + rw (symmetry (list_to_set_disj_env_add (map open_box Γ) (□ φ'))) 0. + rewrite <- list_to_set_disj_open_boxes. exact Hp. + } + all: right; auto with *. +} +destruct HBoxR as [(φ' & Hp)| HBoxR ]. +{ left. destruct Hp as [Hp Heq]. subst. eexists; trivial. apply BoxR, Hp. } +assert(Hempty: ∀ (Δ : env) φ,((Δ • φ) = ∅) -> False). +{ + intros Δ θ Heq. assert (Hm:= multiplicity_empty θ). + unfold base.empty in *. + rewrite <- Heq, union_mult, singleton_mult_in in Hm by trivial. lia. +} +(* non invertible left rules *) +assert(HImpLImp : ∀Γ2 Γ1, Γ1 ++ Γ2 = Γ -> {φ1 & {φ2 & {φ3 & exists (_ : (list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • (Implies φ2 φ3)) ⊢ (Implies φ1 φ2)), + exists (_: list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • φ3 ⊢ φ), (Implies (Implies φ1 φ2) φ3) ∈ Γ2}}} + + {∀ φ1 φ2 φ3 (_ : (list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • (Implies φ2 φ3)) ⊢ (Implies φ1 φ2)) + (_: list_to_set_disj (rm (Implies (Implies φ1 φ2) φ3) Γ) • φ3 ⊢ φ), + Implies (Implies φ1 φ2) φ3 ∈ Γ2 → False}). +{ + induction Γ2 as [|θ Γ2]; intros Γ1 Heq. + - right. intros φ1 φ2 φ3 _ _ Hin. inversion Hin. + - assert(Heq' : (Γ1 ++ [θ]) ++ Γ2 = Γ) by (subst; auto with *). + destruct (IHΓ2 (Γ1 ++ [θ]) Heq') as [(φ1 & φ2 & φ3 & Hp)|Hf]. + + left. do 3 eexists. destruct Hp as ( Hp1 & Hp2 & Hin). do 2 (eexists; eauto). now right. + + destruct θ. + 5: destruct θ1. + 9 : { + destruct (Hind (Implies θ1_2 θ2 :: rm (Implies (Implies θ1_1 θ1_2) θ2) Γ) (Implies θ1_1 θ1_2)) + as [Hp1| Hf']. + - order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + - destruct (Hind (θ2 :: rm (Implies (Implies θ1_1 θ1_2) θ2) Γ) φ) as [Hp2| Hf'']. + + order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + + left. do 3 eexists. destruct Hp1 as [Hp1 _]. destruct Hp2 as [Hp2 _]. + eexists; try l_tac; eauto. ms. + + right; intros φ1 φ2 φ3 Hp1' Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf''. peapply Hp2. + - right; intros φ1 φ2 φ3 Hp1 Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf'. peapply Hp1. + } + all: (right; intros φ1 φ2 φ3 Hp1 Hp2 He; apply elem_of_list_In in He; destruct He as [Heq''| Hin]; + [discriminate|apply elem_of_list_In in Hin; eapply Hf; eauto]). +} +destruct (HImpLImp Γ [] (app_nil_l _)) as [(φ1 & φ2 & φ3 & Hp1)|HfImpl]. +{ left. destruct Hp1 as (Hp1 & Hp2 & Hin). eexists; trivial. + apply elem_of_list_to_set_disj in Hin. exhibit Hin 0. + rw (list_to_set_disj_rm Γ(Implies (Implies φ1 φ2) φ3)) 1. + apply ImpLImp; assumption. +} +(* ImpBox *) +assert(HImpLBox : ∀Γ2 Γ1, Γ1 ++ Γ2 = Γ -> {φ1 & {φ2 & exists (_ : (⊗(list_to_set_disj ((rm (Implies (□ φ1) φ2) Γ))) • □ φ1 • φ2) ⊢φ1), + exists (_ : list_to_set_disj (rm (Implies (□ φ1) φ2) Γ) • φ2 ⊢ φ), + (Implies (□ φ1) φ2) ∈ Γ2}} + + {∀ φ1 φ2 (_ : ((⊗ (list_to_set_disj ((rm (Implies (□ φ1) φ2) Γ))) • □ φ1 • φ2) ⊢φ1)) + (_: list_to_set_disj (rm (Implies (□ φ1) φ2) Γ) • φ2 ⊢ φ), + Implies (□ φ1) φ2 ∈ Γ2 → False}). +{ + induction Γ2 as [|θ Γ2]; intros Γ1 Heq. + - right. intros φ1 φ2 _ _ Hin. inversion Hin. + - assert(Heq' : (Γ1 ++ [θ]) ++ Γ2 = Γ) by (subst; auto with *). + destruct (IHΓ2 (Γ1 ++ [θ]) Heq') as [(φ1 & φ2 & Hp1)|Hf]. + + left. do 2 eexists. destruct Hp1 as (Hp1 & Hp2 & Hin). do 2 (eexists; eauto). now right. + + destruct θ. + 5: destruct θ1. + 10 : { + destruct (Hind (θ2 :: (□θ1) :: map open_box (rm (Implies (□ θ1) θ2) Γ)) θ1) + as [Hp1|Hf']. + - order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + - destruct (Hind (θ2 :: rm (Implies (□ θ1) θ2) Γ) φ) as [Hp2| Hf'']. + + order_tac. rewrite <- Permutation_middle. unfold rm. + destruct form_eq_dec; [|tauto]. order_tac. + + left. do 2 eexists. destruct Hp1 as [Hp1 _]. destruct Hp2 as [Hp2 _]. + repeat eexists; repeat l_tac; eauto. ms. + + right; intros φ1 φ2 Hp1' Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf''. peapply Hp2. + - right; intros φ1 φ2 Hp1 Hp2 He; apply elem_of_list_In in He; + destruct He as [Heq''| Hin]; [|apply elem_of_list_In in Hin; eapply Hf; eauto]. + inversion Heq''. subst. apply Hf'. + (erewrite proper_Provable; [| |reflexivity]); [eapply Hp1|]. + repeat rewrite <- ?list_to_set_disj_env_add, list_to_set_disj_open_boxes. trivial. + } + all: (right; intros φ1 φ2 Hp1 Hp2 He; apply elem_of_list_In in He; destruct He as [Heq''| Hin]; + [discriminate|apply elem_of_list_In in Hin; eapply Hf; eauto]). +} +destruct (HImpLBox Γ [] (app_nil_l _)) as [(φ1 & φ2 & Hp1)|HfImpLBox]. +{ left. destruct Hp1 as (Hp1 & Hp2 & Hin). eexists; trivial. + apply elem_of_list_to_set_disj in Hin. exhibit Hin 0. + rw (list_to_set_disj_rm Γ(Implies (□ φ1) φ2)) 1. + apply ImpBox; assumption. +} +clear Hind HImpLImp HImpLBox. +right. +intro Hp. inversion Hp; subst; try eqt; eauto 2. +- eapply HAndR; eauto. +- eapply HImpR; eauto. +- eapply HImpLVar; eauto. apply elem_of_list_to_set_disj. setoid_rewrite <- H; ms. +- eapply HfImpl; eauto. + + now rw Heq 1. + + now rw Heq 1. +- eapply HfImpLBox; eauto. + + now rw (proper_open_boxes _ _ Heq) 2. + + now rw Heq 1. +Defined. From 8521f2f672b3aa6c856193c9aa655ca9449cc5d2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Tue, 3 Sep 2024 15:07:01 +0200 Subject: [PATCH 04/12] Decision procedure: name cleanup --- bin/isl_dec.ml | 2 +- theories/extraction/UIML_extraction.v | 2 +- theories/iSL/DecisionProcedure.v | 6 ++++-- 3 files changed, 6 insertions(+), 4 deletions(-) diff --git a/bin/isl_dec.ml b/bin/isl_dec.ml index d33a00e..ef3cf4d 100644 --- a/bin/isl_dec.ml +++ b/bin/isl_dec.ml @@ -22,6 +22,6 @@ let print_decision = function | _ -> "Not provable" let () = - if nb_args = 2 then form |> eval |> coq_Provable_dec' [] |> + if nb_args = 2 then form |> eval |> coq_Provable_dec [] |> print_decision |> print_string |> print_newline else print_string usage_string diff --git a/theories/extraction/UIML_extraction.v b/theories/extraction/UIML_extraction.v index b449651..1c6ec99 100644 --- a/theories/extraction/UIML_extraction.v +++ b/theories/extraction/UIML_extraction.v @@ -41,5 +41,5 @@ Definition isl_simplified_A v f := A_simplified v f. Set Extraction Output Directory "extraction". -Separate Extraction Provable_dec' gl_UI k_UI isl_E isl_A isl_simplified_E isl_simplified_A Formulas.weight isl_simp. +Separate Extraction Provable_dec gl_UI k_UI isl_E isl_A isl_simplified_E isl_simplified_A Formulas.weight isl_simp. diff --git a/theories/iSL/DecisionProcedure.v b/theories/iSL/DecisionProcedure.v index de48c30..a7802ea 100644 --- a/theories/iSL/DecisionProcedure.v +++ b/theories/iSL/DecisionProcedure.v @@ -21,7 +21,8 @@ induction l as [|x l]. * right. simpl. intros z [Hz|Hz]; subst; try rewrite Heq; auto with *. Defined. -Proposition Provable_dec Γ φ : +(* This function computes a proof tree of a sequent, if there is one, or produces a proof that there is none *) +Proposition Proof_tree_dec Γ φ : {_ : list_to_set_disj Γ ⊢ φ & True} + {forall H : list_to_set_disj Γ ⊢ φ, False}. Proof. (* duplicate *) @@ -321,7 +322,8 @@ intro Hp. inversion Hp; subst; try eqt; eauto 2. Defined. -Proposition Provable_dec' Γ φ : +(* This function decides whether a sequent is provable *) +Proposition Provable_dec Γ φ : (exists _ : list_to_set_disj Γ ⊢ φ, True) + (forall H : list_to_set_disj Γ ⊢ φ, False). Proof. remember (Γ, φ) as pe. From 8c0a2fbbcb106777bff71f5aa1b69176b0494949 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Wed, 4 Sep 2024 18:13:34 +0200 Subject: [PATCH 05/12] Compute interpolants by apply invertible rules first ; proofs missing --- theories/iSL/InvPropQuantifiers.v | 303 ++++++++++++++++++++++++++++++ theories/iSL/Order.v | 1 + 2 files changed, 304 insertions(+) create mode 100644 theories/iSL/InvPropQuantifiers.v diff --git a/theories/iSL/InvPropQuantifiers.v b/theories/iSL/InvPropQuantifiers.v new file mode 100644 index 0000000..21d3d14 --- /dev/null +++ b/theories/iSL/InvPropQuantifiers.v @@ -0,0 +1,303 @@ +Require Import ISL.Sequents ISL.SequentProps ISL.Order ISL.Optimizations ISL.DecisionProcedure. +Require Import Coq.Program.Equality. (* for dependent induction *) + +(** * Overview: Propositional Quantifiers + +The main theorem proved in this file was first proved as Theorem 1 in: + +(Pitts 1992). A. M. Pitts. On an interpretation of second order quantification in first order intuitionistic propositional logic. J. Symb. Log., 57(1):33–52. +It has been further extended to handle iSL + +It consists of two parts: + +1) the inductive construction of the propositional quantifiers; + +2) a proof of its correctness. *) + +Section UniformInterpolation. + +(** Throughout the construction and proof, we fix a variable p, with respect to + which the propositional quantifier will be computed. *) +Variable p : variable. + + +(** * Definition of propositional quantifiers. *) + +Section PropQuantDefinition. + +(** We define the formulas Eφ and Aφ associated to any formula φ. This + is an implementation of Pitts' Table 5, together with a (mostly automatic) + proof that the definition terminates*) + +Local Notation "Δ '•' φ" := (cons φ Δ). + +Notation "□⁻¹ Γ" := (map open_box Γ) (at level 75). + +Definition applicable_AndR φ: {φ1 & {φ2 | φ = (And φ1 φ2)}} + (∀ φ1 φ2, φ ≠ (And φ1 φ2)). +Proof. (destruct φ; eauto). Defined. + +Definition applicable_other_var (Γ : list form): {q | p ≠ q /\ Var q ∈ Γ} + (∀ q, p ≠ q -> Var q ∈ Γ -> False). +Proof. + pose (fA := (fun θ => match θ with |Var q => if decide (p = q) then false else true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. + 1: { case decide in Hθ. auto with *. eexists; split; eauto. apply elem_of_list_In. assumption. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. + case decide in Hψ; auto with *. +Defined. + +Definition applicable_AndL (Γ : list form): {ψ1 & {ψ2 | (And ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False). +Proof. + pose (fA := (fun θ => match θ with |And _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 3: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +Defined. + +Definition applicable_ImpR φ: {φ1 & {φ2 | φ = (Implies φ1 φ2)}} + (∀ φ1 φ2, φ ≠ (Implies φ1 φ2)). +Proof. destruct φ; eauto. Defined. + + +Definition applicable_OrL (Γ : list form): {ψ1 & {ψ2 | (Or ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (Or ψ1 ψ2) ∈ Γ -> False). +Proof. + pose (fA := (fun θ => match θ with |Or _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 4: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +Defined. + +Definition applicable_ImpL_other_Var (Γ : list form): + {q & {ψ | q ≠ p /\ (Implies (Var q) ψ) ∈ Γ}} + + (∀ q ψ, q ≠ p -> (Implies (Var q) ψ) ∈ Γ -> False). +Proof. + pose (fIp :=λ θ, match θ with | Implies (Var q) _ => if decide (p = q) then false else true | _ => false end). + destruct (exists_dec fIp Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fIp. destruct θ; auto with *. + destruct θ1; auto with *. + case decide in Hθ; auto with *. + apply elem_of_list_In in Hin. + do 2 eexists; split; eauto. + - right. intros q ψ Hp Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fIp. + simpl in Hψ. rewrite decide_False in Hψ; trivial; auto with *. +Defined. + +Definition applicable_ImpLVar (Γ : list form): + {q & {ψ | Var q ∈ Γ /\ (Implies (Var q) ψ) ∈ Γ}} + + (∀ q ψ, Var q ∈ Γ -> (Implies (Var q) ψ) ∈ Γ -> False). +Proof. + pose (fIp :=λ p θ, match θ with | Implies (Var q) _ => if decide (p = q) then true else false | _ => false end). + pose (fp:= (fun θ => match θ with |Var p => if (exists_dec (fIp p) Γ) then true else false | _ => false end)). + destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fp. destruct θ. 2-6: auto with *. + case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. + unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. + destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. + subst. apply elem_of_list_In in Hinψ, Hin. + do 2 eexists. split; eauto. + - right. intros q ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. + simpl in Hp. case exists_dec as [|Hf'] in Hp. auto with *. + apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. +Defined. + +Definition applicable_ImpLVar_spec (Γ : list form) q: + {ψ | Var q ∈ Γ /\ (Implies (Var q) ψ) ∈ Γ} + + (∀ ψ, Var q ∈ Γ -> (Implies (Var q) ψ) ∈ Γ -> False). +Proof. + pose (fIp :=λ θ, match θ with | Implies (Var p) _ => if decide (p = q) then true else false | _ => false end). + pose (fp:= (fun θ => match θ with |Var p => if decide (p = q) then if (exists_dec fIp Γ) then true else false else false| _ => false end)). + destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fp. destruct θ. 2-6: auto with *. + case decide in Hθ; [|auto with *]. subst. + case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. + unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. + destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. + subst. apply elem_of_list_In in Hinψ, Hin. + eexists. split; eauto. + - right. intros ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. + simpl in Hp. rewrite decide_True in Hp by trivial. + case exists_dec as [|Hf'] in Hp. auto with *. + apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. +Defined. + +Definition applicable_ImpLAnd (Γ : list form): + {φ1 & {φ2 & {φ3 | (Implies (And φ1 φ2) φ3) ∈ Γ}}} + + (∀ φ1 φ2 φ3, (Implies (And φ1 φ2) φ3) ∈ Γ -> False). +Proof. + pose (fII := (fun θ => match θ with |Implies (And _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-2,4-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +Defined. + +Definition applicable_ImpLOr (Γ : list form): + {φ1 & {φ2 & {φ3 | (Implies (Or φ1 φ2) φ3) ∈ Γ}}} + + (∀ φ1 φ2 φ3, (Implies (Or φ1 φ2) φ3) ∈ Γ -> False). +Proof. + pose (fII := (fun θ => match θ with |Implies (Or _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-3, 5-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +Defined. + +Definition applicable_OrR φ: + {φ1 & {φ2 | φ = (Or φ1 φ2)}} + (∀ φ1 φ2, φ = (Or φ1 φ2) -> False). +Proof. destruct φ; eauto with *. Defined. + +Definition applicable_BoxR φ: {φ' | φ = (□ φ')} + (∀ φ', φ = (□ φ') -> False). +Proof. destruct φ; eauto with *. Defined. + +Definition applicable_ImpLImp (Γ : list form): + {φ1 & {φ2 & {φ3 | ((φ1 → φ2) → φ3) ∈ Γ}}} +(forall φ1 φ2 φ3, ((φ1 → φ2) → φ3) ∈ Γ -> False). +Proof. + pose (fII := (fun θ => match θ with |Implies (Implies _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-4, 6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +Defined. + +(* can this be replaced with generalised axiom?*) +Definition applicable_Atom q (Δ : list form) ϕ : + (ϕ = Var q /\ Var q ∈ Δ) + (ϕ ≠ Var q \/ (Var q ∈ Δ -> False)). +Proof. +case (decide (ϕ = Var q)). +case (decide (Var q ∈ Δ)). +all: auto with *. +Defined. + +Definition sumor_bind0 {A} {B : Prop} {C}: A + B -> (A -> C) -> C -> C := +λ qH f c, +match qH with +| inl a => f a +| inr _ => c +end. + +Definition sumor_bind1 {A A'} {B : Prop} {C}: {q : A | A' q} + B -> (forall x (_ : A' x), C) -> C -> C := +λ qH f c, +match qH with +| inl (exist _ q Hq) => f q Hq +| inr _ => c +end. + +Definition sumor_bind2 {A A' A''} {B : Prop} {C}: {q : A & {r : A' | A'' q r}} + B -> (forall x y (_ : A'' x y), C) -> C -> C := +λ qH f c, +match qH with +| inl (existT q (exist _ r Hq)) => f q r Hq +| inr _ => c +end. + +Definition sumor_bind3 {A A' A'' A'''} {B : Prop} {C}: + {q : A & {r : A' & { s : A'' | A''' q r s}}} + B -> (forall x y z (_ : A''' x y z), C) -> C -> C := +λ qH f c, +match qH with +| inl (existT q (existT r (exist _ s Hq))) => f q r s Hq +| inr _ => c +end. + +Notation "cond '?' A ':0' B" := (sumor_bind0 cond A B) (at level 150, right associativity). +Notation "cond '?' A ':1' B" := (sumor_bind1 cond A B) (at level 150, right associativity). +Notation "cond '?' A ':2' B" := (sumor_bind2 cond A B) (at level 150, right associativity). +Notation "cond '?' A ':3' B" := (sumor_bind3 cond A B) (at level 150, right associativity). + +(* solves the obligations of the following programs *) +Obligation Tactic := simpl; intros; order_tac. +Program Definition e_rule {Δ : list form } {ϕ : form} + (EA0 : ∀ bool pe (Hpe : pe ≺· (Δ, ϕ)), form) + (θ: form) (Hin : θ ∈ Δ) : form := +let E Δ H := EA0 true (Δ, ϕ) H in +let A pe0 H := EA0 false pe0 H in +let Δ' := rm θ Δ in +match θ with +| Var q → δ => if decide (p = q) then ⊤ else q → E (Δ'•δ) _ (* E5 modified *) +(* E8 modified *) +| ((δ₁→ δ₂)→ δ₃) => + (E (Δ'•(δ₂ → δ₃) • δ₁) _⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) ⇢ E (Δ'•δ₃) _ +| □ φ => □(E ((□⁻¹ Δ') • φ ) _) (* very redundant ; identical for each box *) +| (□δ1 → δ2) => (□(E((□⁻¹ Δ') • δ2 • □δ1) _ ⇢ A((□⁻¹ Δ') • δ2 • □δ1, δ1) _)) ⇢ E(Δ' • δ2) _ +| _ => ⊤ +end. + +(** The implementation of the rules for defining A is separated into two pieces. + Referring to Table 5 in Pitts, the definition a_rule_env handles A1-8 and A10, + and the definition a_rule_form handles A9 and A11-13. *) +Program Definition a_rule_env {Δ : list form} {ϕ : form} + (EA0 : ∀ bool pe (Hpe : pe ≺· (Δ, ϕ)), form) + (θ: form) (Hin : θ ∈ Δ) : form := +let E Δ H := EA0 true (Δ, ϕ) H in +let A pe0 H := EA0 false pe0 H in +let Δ' := rm θ Δ in +match θ with +| Var q → δ => if decide (p = q) then ⊥ else q ⊼ A (Δ'•δ, ϕ) _ (* A4 *) +(* A6 *) +(* A8 modified*) +| ((δ₁→ δ₂)→ δ₃) => + (E (Δ'•(δ₂ → δ₃) • δ₁) _ ⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) + ⊼ A (Δ'•δ₃, ϕ) _ +| ((□δ1) → δ2) => (□(E((□⁻¹ Δ')• δ2 • □δ1) _ ⇢ A((□⁻¹ Δ')• δ2 • □δ1, δ1) _)) ∧ A(Δ' • δ2, ϕ) _ +(* using ⊼ here breaks congruence *) +| _ => ⊥ +end. + +Program Definition a_rule_form {Δ : list form} {ϕ : form} + (EA0 : ∀ bool pe (Hpe : pe ≺· (Δ, ϕ)), form) := +let E pe0:= EA0 true pe0 in +let A pe0 H := EA0 false pe0 H in +match ϕ with +| Var q => + if decide (p = q) + then ⊥ + else Var q (* A9 *) +(* A12 *) +| ϕ₁ ∨ ϕ₂ => A (Δ, ϕ₁) _ ⊻ A (Δ, ϕ₂) _ +| □δ => □((E ((□⁻¹ Δ) • □δ, δ) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) +| _ => ⊥ +end. + +Obligation Tactic := solve[simpl; intros; simpl in *; intuition; order_tac] || (try (apply Wf.measure_wf, wf_pointed_order)). +Program Fixpoint EA (b : bool) (pe : list form * form) {wf pointed_env_order pe} := + let Δ := fst pe in + let ϕ := snd pe in + let E := EA true in + let A := EA false in + (* E *) + if b then + if decide (⊥ ∈ Δ) then ⊥ else + applicable_other_var Δ ? λ q _, q ⊼ E (rm (Var q) Δ, ϕ) _ :1 + applicable_AndL Δ ? λ δ₁ δ₂ _, E ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂, ϕ) _ :2 + applicable_OrL Δ ? λ δ₁ δ₂ _, E ((rm (δ₁ ∨ δ₂) Δ•δ₁, ϕ)) _ ⊻E ((rm (δ₁ ∨ δ₂) Δ•δ₂, ϕ)) _ :2 + applicable_ImpLVar Δ ? λ q ψ _, E ((rm (Var q → ψ) Δ • ψ, ϕ)) _ :2 + applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, E ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)), ϕ)) _ :3 + applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, E (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃), ϕ) _ :3 + ⋀ (in_map Δ (e_rule EA)) (* the non-invertible rules *) + + (* A *) + else + applicable_other_var Δ ? λ q _, A (rm (Var q) Δ, ϕ) _ :1 + applicable_AndL Δ ? λ δ₁ δ₂ _, A ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂, ϕ) _ :2 + applicable_OrL Δ ? λ δ₁ δ₂ _, (E ((rm (δ₁ ∨ δ₂) Δ•δ₁, ϕ)) _ ⇢ A ((rm (δ₁ ∨ δ₂) Δ•δ₁, ϕ)) _) ⊼ + (E ((rm (δ₁ ∨ δ₂) Δ•δ₂, ϕ)) _ ⇢ A ((rm (δ₁ ∨ δ₂) Δ•δ₂, ϕ)) _) :2 + applicable_ImpLVar Δ ? λ q ψ _, A (rm (Var q → ψ) Δ • ψ, ϕ) _ :2 + applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, A ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)), ϕ)) _ :3 + applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, A (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃), ϕ) _ :3 + applicable_Atom p Δ ϕ ? λ _, ⊤ :0 + applicable_AndR ϕ ? λ ϕ₁ ϕ₂ _, A (Δ, ϕ₁) _ ⊼ A (Δ, ϕ₂) _ :2 + applicable_ImpR ϕ ? λ ϕ₁ ϕ₂ _, E (Δ • ϕ₁, ϕ₂) _ ⇢ A (Δ • ϕ₁, ϕ₂) _ :2 + (⋁ (in_map Δ (a_rule_env EA)) ⊻ a_rule_form EA) (* the non-invertible rules *) +. + + +Definition E pe := EA true pe. +Definition A pe := EA false pe. + +Definition Ef (ψ : form) := E ([ψ], ⊥). +Definition Af (ψ : form) := A ([], ψ). + +End PropQuantDefinition. + +End UniformInterpolation. + diff --git a/theories/iSL/Order.v b/theories/iSL/Order.v index 2d9a5be..b640f40 100644 --- a/theories/iSL/Order.v +++ b/theories/iSL/Order.v @@ -369,6 +369,7 @@ unfold pointed_env_order; subst; simpl; repeat rewrite open_boxes_add; try match apply (env_order_equiv_right_compat (equiv_disj_union_compat_r(equiv_disj_union_compat_r(difference_singleton Γ ψ' H)))) || (eapply env_order_le_lt_trans; [| apply env_order_add_compat; eapply env_order_lt_le_trans; [| (apply env_order_eq_add; apply (remove_In_env_order_refl _ ψ'); try apply elem_of_list_In; trivial) ] ] ) +|H : ?a = _ |- context[?a] => rewrite H; try prepare_order end. From d76ad489e0ba0aba3724c3b786b4c43e5452087b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Mon, 9 Sep 2024 13:23:39 +0200 Subject: [PATCH 06/12] decorrelate E and A --- theories/iSL/PropQuantifiers.v | 267 +++++++++++++++------------------ 1 file changed, 124 insertions(+), 143 deletions(-) diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index f91c7a6..c00eda6 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -32,47 +32,19 @@ Section PropQuantDefinition. Local Notation "Δ '•' φ" := (cons φ Δ). -Notation "□⁻¹ Γ" := (map open_box Γ) (at level 75). - -(** TODO: Experimental: weight_preserving simplifications *) - -Definition isimp (φ : form) : list form := [φ]. -Opaque isimp. - -Lemma env_order_refl_isimp φ: env_order_refl (isimp φ) [φ]. -Proof. -Admitted. - - -Lemma isim_vars {φ ψ x}: occurs_in x ψ -> ψ ∈ isimp φ -> occurs_in x φ. -Proof. -Admitted. - -Lemma isimp_complete_L (Γ : env) φ θ: Γ ⊎ list_to_set_disj (isimp φ) ⊢ θ -> Γ ⊎ {[φ]} ⊢ θ. -Proof. Admitted. - -Ltac isimp_order := match goal with -| |- ?Δ ++ isimp ?φ ≺ _ => eapply env_order_le_lt_trans; -[apply env_order_refl_disj_union_compat ; [ isimp_order |apply (env_order_refl_isimp φ)]| -repeat rewrite <- Permutation_cons_append] -| |- env_order_refl (_ ++ isimp _) _ => -apply env_order_refl_disj_union_compat; [ isimp_order|apply env_order_refl_isimp] -| |- _ • _ ≺ _ => eapply env_order_le_lt_trans; [repeat apply env_order_eq_add; isimp_order| -repeat rewrite <- Permutation_cons_append] -| _ => right; reflexivity -end. - (* solves the obligations of the following programs *) -Obligation Tactic := try solve[intros; order_tac; try isimp_order; order_tac]. +Obligation Tactic := intros; order_tac. + +Notation "□⁻¹ Γ" := (map open_box Γ) (at level 75). (** First, the implementation of the rules for calculating E. The names of the rules refer to the table in Pitts' paper. *) (** note the use of "lazy" conjunctions, disjunctions and implications *) Program Definition e_rule {Δ : list form } {ϕ : form} - (EA0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form * form) + (E : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (θ: form) (Hin : θ ∈ Δ) : form := -let E Δ H := fst (EA0 (Δ, ϕ) H) in -let A pe0 H := snd (EA0 pe0 H) in +let E Δ H := E (Δ, ϕ) H in let Δ' := rm θ Δ in match θ with | Bot => ⊥ (* E0 *) @@ -80,23 +52,23 @@ match θ with if decide (p = q) then ⊤ (* default *) else q (* E1 modified *) (* E2 *) -| δ₁ ∧ δ₂ => E ((Δ' ++ isimp δ₁) ++ isimp δ₂) _ +| δ₁ ∧ δ₂ => E ((Δ'•δ₁)•δ₂) _ (* E3 *) -| δ₁ ∨ δ₂ => E (Δ'•δ₁) _ ⊻ E (Δ' ++ isimp δ₂) _ +| δ₁ ∨ δ₂ => E (Δ'•δ₁) _ ⊻ E (Δ' •δ₂) _ | Var q → δ => - if decide (Var q ∈ Δ) then E (Δ' ++ isimp δ) _ (* E5 modified *) + if decide (Var q ∈ Δ) then E (Δ'•δ) _ (* E5 modified *) else if decide (p = q) then ⊤ - else q ⇢ E (Δ' ++ isimp δ) _ (* E4 *) + else q ⇢ E (Δ'•δ) _ (* E4 *) (* E6 *) -| (δ₁ ∧ δ₂)→ δ₃ => E (Δ' ++ isimp (δ₁ → (δ₂ → δ₃))) _ +| (δ₁ ∧ δ₂)→ δ₃ => E (Δ'•(δ₁ → (δ₂ → δ₃))) _ (* E7 *) -| (δ₁ ∨ δ₂)→ δ₃ => E ((Δ' ++ isimp (δ₁ → δ₃)) ++ isimp (δ₂ → δ₃)) _ +| (δ₁ ∨ δ₂)→ δ₃ => E (Δ'•(δ₁ → δ₃)•(δ₂ → δ₃)) _ (* E8 modified *) | ((δ₁→ δ₂)→ δ₃) => - (E ((Δ' ++ isimp (δ₂ → δ₃)) ++ isimp δ₁) _⇢ A ((Δ' ++ isimp (δ₂ → δ₃)) ++ isimp δ₁, δ₂) _) ⇢ E (Δ' ++ isimp δ₃) _ + (E (Δ'•(δ₂ → δ₃) • δ₁) _⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) ⇢ E (Δ'•δ₃) _ | Bot → _ => ⊤ -| □ φ => □(E ((□⁻¹ Δ') ++ isimp φ ) _) (* very redundant ; identical for each box *) -| (□δ1 → δ2) => (□(E(((□⁻¹ Δ') ++ isimp δ2) ++ isimp (□δ1)) _ ⇢ A(((□⁻¹ Δ') ++ isimp δ2) ++ isimp (□δ1), δ1) _)) ⇢ E(Δ' ++ isimp δ2) _ +| □ φ => □(E ((□⁻¹ Δ') • φ ) _) (* very redundant ; identical for each box *) +| (□δ1 → δ2) => (□(E((□⁻¹ Δ') • δ2 • □δ1) _ ⇢ A((□⁻¹ Δ') • δ2 • □δ1, δ1) _)) ⇢ E(Δ' • δ2) _ end. @@ -104,10 +76,10 @@ end. Referring to Table 5 in Pitts, the definition a_rule_env handles A1-8 and A10, and the definition a_rule_form handles A9 and A11-13. *) Program Definition a_rule_env {Δ : list form} {ϕ : form} - (EA0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form * form) + (E : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (θ: form) (Hin : θ ∈ Δ) : form := -let E Δ H := fst (EA0 (Δ, ϕ) H) in -let A pe0 H := snd (EA0 pe0 H) in +let E Δ H := E (Δ, ϕ) H in let Δ' := rm θ Δ in match θ with | Var q => @@ -117,41 +89,36 @@ match θ with else ⊥ else ⊥ (* A1 modified : A (Δ', ϕ) can be removed *) (* A2 *) -| δ₁ ∧ δ₂ => A ((Δ' ++ isimp δ₁) ++ isimp δ₂, ϕ) _ +| δ₁ ∧ δ₂ => A ((Δ'•δ₁)•δ₂, ϕ) _ (* A3 *) | δ₁ ∨ δ₂ => - (E (Δ' ++ isimp δ₁) _ ⇢ A (Δ' ++ isimp δ₁, ϕ) _) - ⊼ (E (Δ' ++ isimp δ₂) _ ⇢ A (Δ' ++ isimp δ₂, ϕ) _) + (E (Δ'•δ₁) _ ⇢ A (Δ'•δ₁, ϕ) _) + ⊼ (E (Δ'•δ₂) _ ⇢ A (Δ'•δ₂, ϕ) _) | Var q → δ => if decide (Var q ∈ Δ) then A (Δ'•δ, ϕ) _ (* A5 modified *) else if decide (p = q) then ⊥ - else q ⊼ A (Δ' ++ isimp δ, ϕ) _ (* A4 *) + else q ⊼ A (Δ'•δ, ϕ) _ (* A4 *) (* A6 *) | (δ₁ ∧ δ₂)→ δ₃ => - A (Δ' ++ isimp (δ₁ → (δ₂ → δ₃)), ϕ) _ + A (Δ'•(δ₁ → (δ₂ → δ₃)), ϕ) _ (* A7 *) | (δ₁ ∨ δ₂)→ δ₃ => - A ((Δ' ++ isimp (δ₁ → δ₃)) ++ isimp (δ₂ → δ₃), ϕ) _ + A ((Δ'•(δ₁ → δ₃))•(δ₂ → δ₃), ϕ) _ (* A8 modified*) | ((δ₁→ δ₂)→ δ₃) => - (E ((Δ' ++ isimp (δ₂ → δ₃)) ++ isimp δ₁) _ ⇢ A ((Δ' ++ isimp (δ₂ → δ₃) )++ isimp δ₁, δ₂) _) - ⊼ A (Δ' ++ isimp δ₃, ϕ) _ + (E (Δ'•(δ₂ → δ₃) • δ₁) _ ⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) + ⊼ A (Δ'•δ₃, ϕ) _ | Bot => ⊥ | Bot → _ => ⊥ | □δ => ⊥ -| ((□δ1) → δ2) => (□(E(((□⁻¹ Δ') ++ isimp δ2 ) ++ isimp (□δ1)) _ - ⇢ A(((□⁻¹ Δ') ++ isimp δ2 ) ++ isimp (□δ1), δ1) _)) - ∧ A(Δ' ++ isimp δ2, ϕ) _ +| ((□δ1) → δ2) => (□(E((□⁻¹ Δ')• δ2 • □δ1) _ ⇢ A((□⁻¹ Δ')• δ2 • □δ1, δ1) _)) ∧ A(Δ' • δ2, ϕ) _ (* using ⊼ here breaks congruence *) end. - -(* make sure that the proof obligations do not depend on EA0 *) -Obligation Tactic := intros Δ ϕ _ _ _; intros; order_tac. +Obligation Tactic := intros Δ ϕ _ _; intros; order_tac. Program Definition a_rule_form {Δ : list form} {ϕ : form} - (EA0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form * form) : form := -let E pe0 H := fst (EA0 pe0 H) in -let A pe0 H := snd (EA0 pe0 H) in + (E : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) := match ϕ with | Var q => if decide (p = q) (* TODO : change this to p∈Vars(ϕ) *) @@ -167,17 +134,18 @@ match ϕ with | Bot => ⊥ | □δ => □((E ((□⁻¹ Δ) • □δ, δ) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) end. -(* TODO: use isimp here too? *) Obligation Tactic := intros; order_tac. -Program Fixpoint EA (pe : list form * form) {wf pointed_env_order pe} := - let Δ := fst pe in - (⋀ (in_map Δ (e_rule EA)), - ⋁ (in_map Δ (a_rule_env EA)) ⊻ a_rule_form EA). -Next Obligation. apply wf_pointed_order. Defined. +Program Fixpoint EA (b : bool) (pe : list form * form) {wf pointed_env_order pe} : form := + let Δ := pe.1 in + let E := EA true in + let A := EA false in + if b then⋀ (in_map Δ (e_rule E A)) + else ⋁ (in_map Δ (a_rule_env E A)) ⊻ a_rule_form E A. +Next Obligation. apply Wf.measure_wf, wf_pointed_order. Defined. -Definition E pe := (EA pe).1. -Definition A pe := (EA pe).2. +Definition E := EA true. +Definition A := EA false. Definition Ef (ψ : form) := E ([ψ], ⊥). Definition Af (ψ : form) := A ([], ψ). @@ -185,78 +153,92 @@ Definition Af (ψ : form) := A ([], ψ). (** Congruence lemmas: if we apply any of e_rule, a_rule_env, or a_rule_form to two equal environments, then they give the same results. *) -Lemma e_rule_cong Δ ϕ θ (Hin: θ ∈ Δ) EA1 EA2: - (forall pe Hpe, EA1 pe Hpe = EA2 pe Hpe) -> - @e_rule Δ ϕ EA1 θ Hin = @e_rule Δ ϕ EA2 θ Hin. +Lemma e_rule_cong Δ ϕ θ (Hin: θ ∈ Δ) E1 A1 E2 A2: + (forall pe Hpe, E1 pe Hpe = E2 pe Hpe) -> + (forall pe Hpe, A1 pe Hpe = A2 pe Hpe) -> + @e_rule Δ ϕ E1 A1 θ Hin = @e_rule Δ ϕ E2 A2 θ Hin. Proof. - intro Heq. + intros HeqE HeqA. destruct θ; simpl; try (destruct θ1); repeat (destruct decide); - f_equal; repeat erewrite Heq; trivial. + f_equal; repeat erewrite ?HeqE, ?HeqA; trivial. Qed. -Lemma e_rule_cong_strong Δ ϕ θ (Hin1 Hin2: θ ∈ Δ) EA1 EA2: - (forall pe Hpe1 Hpe2, EA1 pe Hpe1 = EA2 pe Hpe2) -> - @e_rule Δ ϕ EA1 θ Hin1 = @e_rule Δ ϕ EA2 θ Hin2. +Lemma e_rule_cong_strong Δ ϕ θ (Hin1 Hin2: θ ∈ Δ) E1 A1 E2 A2: + (forall pe Hpe1 Hpe2, E1 pe Hpe1 = E2 pe Hpe2) -> + (forall pe Hpe1 Hpe2, A1 pe Hpe1 = A2 pe Hpe2) -> + @e_rule Δ ϕ E1 A1 θ Hin1 = @e_rule Δ ϕ E2 A2 θ Hin2. Proof. - intro Heq. + intros HeqE HeqA. destruct θ; simpl; try (destruct θ1); repeat (destruct decide); - f_equal; repeat erewrite Heq; trivial. + f_equal; repeat erewrite ?HeqE, ?HeqA; trivial. Qed. -Lemma a_rule_env_cong Δ ϕ θ Hin EA1 EA2: - (forall pe Hpe, EA1 pe Hpe = EA2 pe Hpe) -> - @a_rule_env Δ ϕ EA1 θ Hin = @a_rule_env Δ ϕ EA2 θ Hin. +Lemma a_rule_env_cong Δ ϕ θ Hin E1 A1 E2 A2: + (forall pe Hpe, E1 pe Hpe = E2 pe Hpe) -> + (forall pe Hpe, A1 pe Hpe = A2 pe Hpe) -> + @a_rule_env Δ ϕ E1 A1 θ Hin = @a_rule_env Δ ϕ E2 A2 θ Hin. Proof. - intro Heq. + intros HeqE HeqA. destruct θ; simpl; trivial; repeat (destruct decide); - f_equal; repeat erewrite Heq; trivial; + f_equal; repeat erewrite ?HeqE, ?HeqA; trivial; destruct θ1; try (destruct decide); trivial; simpl; - repeat erewrite Heq; trivial; (destruct decide); trivial. + repeat erewrite ?HeqE, ?HeqA; trivial; (destruct decide); trivial. Qed. -Lemma a_rule_env_cong_strong Δ ϕ θ Hin1 Hin2 EA1 EA2: - (forall pe Hpe1 Hpe2, EA1 pe Hpe1 = EA2 pe Hpe2) -> - @a_rule_env Δ ϕ EA1 θ Hin1 = @a_rule_env Δ ϕ EA2 θ Hin2. +Lemma a_rule_env_cong_strong Δ ϕ θ Hin1 Hin2 E1 A1 E2 A2: + (forall pe Hpe1 Hpe2, E1 pe Hpe1 = E2 pe Hpe2) -> + (forall pe Hpe1 Hpe2, A1 pe Hpe1 = A2 pe Hpe2) -> + @a_rule_env Δ ϕ E1 A1 θ Hin1 = @a_rule_env Δ ϕ E2 A2 θ Hin2. Proof. - intro Heq. + intros HeqE HeqA. destruct θ; simpl; trivial; repeat (destruct decide); - f_equal; repeat erewrite Heq; trivial; + f_equal; repeat erewrite ?HeqE, ?HeqA; trivial; destruct θ1; try (destruct decide); trivial; simpl; - repeat erewrite Heq; trivial; (destruct decide); trivial. + repeat erewrite ?HeqE, ?HeqA; trivial; (destruct decide); trivial. Qed. -Lemma a_rule_form_cong Δ ϕ EA1 EA2: - (forall pe Hpe, EA1 pe Hpe = EA2 pe Hpe) -> - @a_rule_form Δ ϕ EA1 = @a_rule_form Δ ϕ EA2. + +Lemma a_rule_form_cong Δ ϕ E1 A1 E2 A2: + (forall pe Hpe, E1 pe Hpe = E2 pe Hpe) -> + (forall pe Hpe, A1 pe Hpe = A2 pe Hpe) -> + @a_rule_form Δ ϕ E1 A1 = @a_rule_form Δ ϕ E2 A2. Proof. - intro Heq. + intros HeqE HeqA. destruct ϕ; simpl; repeat (destruct decide); trivial; - repeat (erewrite Heq; eauto); trivial. + repeat (erewrite ?HeqE, ?HeqA; eauto); trivial. Qed. -Lemma a_rule_form_cong_strong Δ ϕ EA1 EA2: - (forall pe Hpe1 Hpe2, EA1 pe Hpe1 = EA2 pe Hpe2) -> - @a_rule_form Δ ϕ EA1 = @a_rule_form Δ ϕ EA2. + +Lemma a_rule_form_cong_strong Δ ϕ E1 A1 E2 A2: + (forall pe Hpe1 Hpe2, E1 pe Hpe1 = E2 pe Hpe2) -> + (forall pe Hpe1 Hpe2, A1 pe Hpe1 = A2 pe Hpe2) -> + @a_rule_form Δ ϕ E1 A1 = @a_rule_form Δ ϕ E2 A2. Proof. - intro Heq. + intros HeqE HeqA. destruct ϕ; simpl; repeat (destruct decide); trivial; - repeat (erewrite Heq; eauto); trivial. + repeat (erewrite ?HeqE, ?HeqA; eauto); trivial. Qed. Lemma EA_eq Δ ϕ: - (E (Δ, ϕ) = ⋀ (in_map Δ (@e_rule Δ ϕ (λ pe _, EA pe)))) /\ - (A (Δ, ϕ) = (⋁ (in_map Δ (@a_rule_env Δ ϕ (λ pe _, EA pe)))) ⊻ - @a_rule_form Δ ϕ (λ pe _, EA pe)). + (E (Δ, ϕ) = ⋀ (in_map Δ (@e_rule Δ ϕ (λ pe _, E pe) (λ pe _, A pe)))) /\ + (A (Δ, ϕ) = (⋁ (in_map Δ (@a_rule_env Δ ϕ (λ pe _, E pe) (λ pe _, A pe)))) ⊻ + @a_rule_form Δ ϕ (λ pe _, E pe) (λ pe _, A pe)). Proof. simpl. unfold E, A, EA. simpl. -rewrite -> Wf.Fix_eq. -- simpl. split; trivial. -- intros Δ' f g Heq. f_equal; f_equal. - + apply in_map_ext. intros. apply e_rule_cong; intros. +unfold EA_func at 1. +rewrite -> Wf.Fix_eq; simpl. +- split; trivial. + unfold EA_func at 1. rewrite -> Wf.Fix_eq; simpl; trivial. + intros [[|] Δ'] f g Heq; simpl; f_equal. + + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. + + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; now rewrite Heq. - + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros. + + apply a_rule_form_cong; intros; now rewrite Heq. +- intros [[|] Δ'] f g Heq; simpl; f_equal. + + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. + + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; now rewrite Heq. - + apply a_rule_form_cong; intros. now rewrite Heq. + + apply a_rule_form_cong; intros; now rewrite Heq. Qed. Definition E_eq Δ ϕ := proj1 (EA_eq Δ ϕ). @@ -293,8 +275,6 @@ try match goal with | H : occurs_in _ (?a ⇢ ?b) |- _ => apply occurs_in_make_impl in H | H : occurs_in _ (?a ⊻ ?b) |- _ => apply occurs_in_make_disj in H | H : occurs_in _ (?a ⊼ ?b) |- _ => apply occurs_in_make_conj in H -|H1 : ?x0 ∈ (?Δ ++ _), H2 : occurs_in ?x ?x0 |- _ => repeat rewrite elem_of_app in H1 -|H1 : ?x0 ∈ (isimp _), H2 : occurs_in ?x ?x0 |- _ => apply (isim_vars H2) in H1 |H1 : ?x0 ∈ (⊗ ?Δ), H2 : occurs_in ?x ?x0 |- _ => apply (occurs_in_open_boxes _ _ _ H2) in H1 |H1 : ?x0 ∈ (map open_box ?Δ), H2 : occurs_in ?x ?x0 |- _ => @@ -309,11 +289,13 @@ try multimatch goal with (** *** (a) *) Lemma e_rule_vars Δ (θ : form) (Hin : θ ∈ Δ) (ϕ : form) - (EA0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form * form) x + (E0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + x (HEA0 : ∀ pe Hpe, - (occurs_in x (fst (EA0 pe Hpe)) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ - (occurs_in x (snd (EA0 pe Hpe)) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))) : -occurs_in x (e_rule EA0 θ Hin) -> x ≠ p ∧ ∃ θ, θ ∈ Δ /\ occurs_in x θ. + (occurs_in x (E0 pe Hpe) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ + (occurs_in x (A0 pe Hpe) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))) : +occurs_in x (e_rule E0 A0 θ Hin) -> x ≠ p ∧ ∃ θ, θ ∈ Δ /\ occurs_in x θ. Proof. destruct θ; unfold e_rule; simpl; try tauto; try solve [repeat case decide; repeat vars_tac]. destruct θ1; repeat case decide; repeat vars_tac. @@ -323,11 +305,13 @@ Qed. (** *** (b) *) Lemma a_rule_env_vars Δ θ Hin ϕ - (EA0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form * form) x + (E0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + x (HEA0 : ∀ pe Hpe, - (occurs_in x (fst (EA0 pe Hpe)) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ - (occurs_in x (snd (EA0 pe Hpe)) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))): -occurs_in x (a_rule_env EA0 θ Hin) -> x ≠ p ∧ (occurs_in x ϕ \/ ∃ θ, θ ∈ Δ /\ occurs_in x θ). + (occurs_in x (E0 pe Hpe) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ + (occurs_in x (A0 pe Hpe) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))): +occurs_in x (a_rule_env E0 A0 θ Hin) -> x ≠ p ∧ (occurs_in x ϕ \/ ∃ θ, θ ∈ Δ /\ occurs_in x θ). Proof. destruct θ; unfold a_rule_env; simpl; try tauto; try solve [repeat case decide; repeat vars_tac]. destruct θ1; repeat case decide; repeat vars_tac. @@ -335,11 +319,13 @@ Qed. Lemma a_rule_form_vars Δ ϕ - (EA0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form * form) x + (E0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + x (HEA0 : ∀ pe Hpe, - (occurs_in x (fst (EA0 pe Hpe)) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ - (occurs_in x (snd (EA0 pe Hpe)) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))): - occurs_in x (a_rule_form EA0) -> x ≠ p ∧ (occurs_in x ϕ \/ ∃ θ, θ ∈ Δ /\ occurs_in x θ). + (occurs_in x (E0 pe Hpe) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ + (occurs_in x (A0 pe Hpe) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))): + occurs_in x (a_rule_form E0 A0) -> x ≠ p ∧ (occurs_in x ϕ \/ ∃ θ, θ ∈ Δ /\ occurs_in x θ). Proof. destruct ϕ; unfold a_rule_form; simpl; try tauto; try solve [repeat case decide; repeat vars_tac]. Qed. @@ -397,9 +383,11 @@ Local Ltac l_tac := repeat rewrite list_to_set_disj_open_boxes; || rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (equiv_disj_union_compat_r (list_to_set_disj_env_add _ _))) _ _ eq_refl) || rewrite (proper_Provable _ _ (equiv_disj_union_compat_r(equiv_disj_union_compat_r (equiv_disj_union_compat_r (list_to_set_disj_env_add _ _)))) _ _ eq_refl). -Lemma a_rule_env_spec (Δ : list form) θ ϕ Hin (EA0 : ∀ pe, (pe ≺· (Δ, ϕ)) → form * form) - (HEA : forall Δ ϕ Hpe, (list_to_set_disj Δ ⊢ fst (EA0 (Δ, ϕ) Hpe)) * (list_to_set_disj Δ• snd(EA0 (Δ, ϕ) Hpe) ⊢ ϕ)) : - (list_to_set_disj Δ • a_rule_env EA0 θ Hin ⊢ ϕ). +Lemma a_rule_env_spec (Δ : list form) θ ϕ Hin + (E0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (A0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) + (HEA : forall Δ ϕ Hpe, (list_to_set_disj Δ ⊢ E0 (Δ, ϕ) Hpe) * (list_to_set_disj Δ• A0 (Δ, ϕ) Hpe ⊢ ϕ)) : + (list_to_set_disj Δ • a_rule_env E0 A0 θ Hin ⊢ ϕ). Proof with (auto with proof). assert (HE := λ Δ0 ϕ0 Hpe, fst (HEA Δ0 ϕ0 Hpe)). assert (HA := λ Δ0 ϕ0 Hpe, snd (HEA Δ0 ϕ0 Hpe)). @@ -411,14 +399,7 @@ rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (equiv_disj_union_compat + subst. case_decide; subst; auto with proof. + exch 0... - constructor 2. -- simpl; exch 0. apply AndL. -(* requires a specific lemma using isimp_complete_L, equiv_assoc and list_to_set_disj_app *) -assert((list_to_set_disj (((rm (θ1 ∧ θ2) Δ) ++ isimp θ1) ++ isimp θ2) - • (EA0 ((rm (θ1 ∧ θ2) Δ ++ isimp θ1) ++ isimp θ2, ϕ) - (a_rule_env_obligation_1 Δ ϕ (θ1 ∧ θ2) Hin θ1 θ2 eq_refl)).2) ⊢ ϕ). - -auto with proof. - exch 1; exch 0. do 2 l_tac... +- simpl; exch 0. apply AndL. exch 1; exch 0. do 2 l_tac... - apply make_conj_sound_L. exch 0. apply OrL; exch 0. + apply AndL. apply make_impl_sound_L. exch 0. apply make_impl_sound_L. @@ -542,21 +523,21 @@ destruct φ'1; repeat rewrite (Hind _ φ) by (trivial; order_tac); trivial. Qed. Lemma E_left {Γ} {θ} {Δ} {φ φ'} (Hin : φ ∈ Δ) : -(Γ•e_rule (λ pe (_ : pe ≺· (Δ, φ')), EA pe) φ Hin) ⊢ θ -> -Γ•E (Δ, φ') ⊢ θ. +(Γ • e_rule (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe) φ Hin) ⊢ θ -> +Γ • E (Δ, φ') ⊢ θ. Proof. intro Hp. rewrite E_eq. -destruct (@in_map_in _ _ _ (e_rule (λ pe (_ : pe ≺· (Δ, φ')), EA pe)) _ Hin) as [Hin' Hrule]. +destruct (@in_map_in _ _ _ (e_rule (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe) ) _ Hin) as [Hin' Hrule]. eapply conjunction_L. - apply Hrule. - exact Hp. Qed. Lemma A_right {Γ} {Δ} {φ φ'} (Hin : φ ∈ Δ) : -Γ ⊢ a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), EA pe) φ Hin -> +Γ ⊢ a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe) φ Hin -> Γ ⊢ A (Δ, φ'). Proof. intro Hp. rewrite A_eq. -destruct (@in_map_in _ _ _ (a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), EA pe)) _ Hin) as [Hin' Hrule]. +destruct (@in_map_in _ _ _ (a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe)) _ Hin) as [Hin' Hrule]. eapply make_disj_sound_R, OrR1, disjunction_R. - exact Hrule. - exact Hp. @@ -1041,4 +1022,4 @@ intros Hp φ Hvarsφ; repeat split. * intros φ0 Hφ0. apply gmultiset_elem_of_singleton in Hφ0. subst. auto with *. * peapply Hyp. * now rewrite E_of_empty. -Qed. +Qed. \ No newline at end of file From 5df78ae1d3714e801771f984b82cd75ba5138b5c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Mon, 9 Sep 2024 14:26:00 +0200 Subject: [PATCH 07/12] Remove extra formula argument from E --- theories/iSL/PropQuantifiers.v | 174 +++++++++++++++------------------ 1 file changed, 77 insertions(+), 97 deletions(-) diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index c00eda6..70507e8 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -37,6 +37,7 @@ Obligation Tactic := intros; order_tac. Notation "□⁻¹ Γ" := (map open_box Γ) (at level 75). + (** First, the implementation of the rules for calculating E. The names of the rules refer to the table in Pitts' paper. *) (** note the use of "lazy" conjunctions, disjunctions and implications *) @@ -44,7 +45,7 @@ Program Definition e_rule {Δ : list form } {ϕ : form} (E : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (A : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (θ: form) (Hin : θ ∈ Δ) : form := -let E Δ H := E (Δ, ϕ) H in +let E Δ H := E (Δ, ⊥) H in let Δ' := rm θ Δ in match θ with | Bot => ⊥ (* E0 *) @@ -71,7 +72,6 @@ match θ with | (□δ1 → δ2) => (□(E((□⁻¹ Δ') • δ2 • □δ1) _ ⇢ A((□⁻¹ Δ') • δ2 • □δ1, δ1) _)) ⇢ E(Δ' • δ2) _ end. - (** The implementation of the rules for defining A is separated into two pieces. Referring to Table 5 in Pitts, the definition a_rule_env handles A1-8 and A10, and the definition a_rule_form handles A9 and A11-13. *) @@ -79,7 +79,7 @@ Program Definition a_rule_env {Δ : list form} {ϕ : form} (E : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (A : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (θ: form) (Hin : θ ∈ Δ) : form := -let E Δ H := E (Δ, ϕ) H in +let E Δ H := E (Δ, ⊥) H in let Δ' := rm θ Δ in match θ with | Var q => @@ -129,10 +129,10 @@ match ϕ with (* A12 *) | ϕ₁ ∨ ϕ₂ => A (Δ, ϕ₁) _ ⊻ A (Δ, ϕ₂) _ (* A13 *) -| ϕ₁→ ϕ₂ => E (Δ•ϕ₁, ϕ₂) _ ⇢ A (Δ•ϕ₁, ϕ₂) _ +| ϕ₁→ ϕ₂ => E (Δ•ϕ₁, ⊥) _ ⇢ A (Δ•ϕ₁, ϕ₂) _ (* TODO would a general right-implication rule be useful efficient? *) | Bot => ⊥ -| □δ => □((E ((□⁻¹ Δ) • □δ, δ) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) +| □δ => □((E ((□⁻¹ Δ) • □δ, ⊥) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) end. Obligation Tactic := intros; order_tac. @@ -144,10 +144,10 @@ Program Fixpoint EA (b : bool) (pe : list form * form) {wf pointed_env_order pe} else ⋁ (in_map Δ (a_rule_env E A)) ⊻ a_rule_form E A. Next Obligation. apply Wf.measure_wf, wf_pointed_order. Defined. -Definition E := EA true. +Definition E Δ:= EA true (Δ, ⊥). Definition A := EA false. -Definition Ef (ψ : form) := E ([ψ], ⊥). +Definition Ef (ψ : form) := E ([ψ]). Definition Af (ψ : form) := A ([], ψ). @@ -220,9 +220,9 @@ Proof. Qed. Lemma EA_eq Δ ϕ: - (E (Δ, ϕ) = ⋀ (in_map Δ (@e_rule Δ ϕ (λ pe _, E pe) (λ pe _, A pe)))) /\ - (A (Δ, ϕ) = (⋁ (in_map Δ (@a_rule_env Δ ϕ (λ pe _, E pe) (λ pe _, A pe)))) ⊻ - @a_rule_form Δ ϕ (λ pe _, E pe) (λ pe _, A pe)). + (E Δ = ⋀ (in_map Δ (@e_rule Δ ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) /\ + (A (Δ, ϕ) = (⋁ (in_map Δ (@a_rule_env Δ ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) ⊻ + @a_rule_form Δ ϕ (λ pe _, E pe.1) (λ pe _, A pe)). Proof. simpl. unfold E, A, EA. simpl. unfold EA_func at 1. @@ -241,7 +241,7 @@ rewrite -> Wf.Fix_eq; simpl. + apply a_rule_form_cong; intros; now rewrite Heq. Qed. -Definition E_eq Δ ϕ := proj1 (EA_eq Δ ϕ). +Definition E_eq Δ := proj1 (EA_eq Δ ⊥). Definition A_eq Δ ϕ := proj2 (EA_eq Δ ϕ). End PropQuantDefinition. @@ -292,7 +292,7 @@ Lemma e_rule_vars Δ (θ : form) (Hin : θ ∈ Δ) (ϕ : form) (E0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) (A0 : ∀ pe (Hpe : pe ≺· (Δ, ϕ)), form) x - (HEA0 : ∀ pe Hpe, + (HE0 : ∀ pe Hpe, (occurs_in x (E0 pe Hpe) -> x ≠ p ∧ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ) /\ (occurs_in x (A0 pe Hpe) -> x ≠ p ∧ (occurs_in x pe.2 \/ ∃ θ, θ ∈ pe.1 /\ occurs_in x θ))) : occurs_in x (e_rule E0 A0 θ Hin) -> x ≠ p ∧ ∃ θ, θ ∈ Δ /\ occurs_in x θ. @@ -330,8 +330,20 @@ Proof. destruct ϕ; unfold a_rule_form; simpl; try tauto; try solve [repeat case decide; repeat vars_tac]. Qed. +Lemma pointed_env_order_bot_R pe Δ φ: (pe ≺· (Δ, ⊥)) -> pe ≺· (Δ, φ). +Proof. +Admitted. + +Hint Resolve pointed_env_order_bot_R : order. + +Lemma pointed_env_order_bot_L pe Δ φ: ((Δ, φ) ≺· pe) -> (Δ, ⊥) ≺· pe. +Proof. +Admitted. + +Hint Resolve pointed_env_order_bot_L : order. + Proposition EA_vars Δ ϕ x: - (occurs_in x (E (Δ, ϕ)) -> x <> p /\ ∃ θ, θ ∈ Δ /\ occurs_in x θ) /\ + (occurs_in x (E Δ) -> x <> p /\ ∃ θ, θ ∈ Δ /\ occurs_in x θ) /\ (occurs_in x (A (Δ, ϕ)) -> x <> p /\ (occurs_in x ϕ \/ (∃ θ, θ ∈ Δ /\ occurs_in x θ))). Proof. remember (Δ, ϕ) as pe. @@ -346,7 +358,7 @@ split. apply in_in_map in Hin as (ψ&Hin&Heq). subst φ. apply e_rule_vars in Hφ. + trivial. - + intros; now apply Hind. + + intros; apply Hind. now apply pointed_env_order_bot_R. (* A *) - intro Hocc. apply occurs_in_make_disj in Hocc as [Hocc|Hocc]. (* disjunction *) @@ -362,8 +374,8 @@ Qed. End VariablesCorrect. Ltac foldEA := repeat match goal with -| |- context C [(EA ?pe).1] => fold (E pe) -| |- context C [(EA ?pe).2] => fold (A pe) +| |- context C [EA true ?pe] => fold (E pe) +| |- context C [EA false ?pe] => fold (A pe) end. (** ** (ii) Entailment *) @@ -431,7 +443,7 @@ rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (equiv_disj_union_compat - auto with proof. Qed. -Proposition entail_correct Δ ϕ : (list_to_set_disj Δ ⊢ E (Δ, ϕ)) * (list_to_set_disj Δ•A (Δ, ϕ) ⊢ ϕ). +Proposition entail_correct Δ ϕ : (list_to_set_disj Δ ⊢ E Δ) * (list_to_set_disj Δ•A (Δ, ϕ) ⊢ ϕ). Proof with (auto with proof). remember (Δ, ϕ) as pe. replace Δ with pe.1 by now subst. @@ -441,7 +453,7 @@ unfold pointed_env_order. intros (Δ, ϕ) Hind. simpl. rewrite E_eq, A_eq. (* uncurry the induction hypothesis for convenience *) -assert (HE := fun d f x=> fst (Hind (d, f) x)). +assert (HE := fun d x=> fst (Hind (d, ⊥) x)). assert (HA := fun d f x=> snd (Hind (d, f) x)). unfold E, A in *; simpl in HE, HA. simpl in *. clear Hind. @@ -478,16 +490,18 @@ destruct ψ; unfold e_rule; exhibit Hi 0; rewrite (proper_Provable _ _ (equiv_di ++ do 2 l_tac. apply HA. order_tac. -- exch 0. l_tac. apply weakening, HE. order_tac. - apply BoxR. apply weakening. box_tac. l_tac. apply HE. order_tac. + } (* A *) apply make_disj_sound_L, OrL. - apply disjunction_L. intros φ Hin. apply in_in_map in Hin as (φ' & Heq & Hφ'). subst φ. apply a_rule_env_spec; intros; split ; apply HE || apply HA; order_tac. + eapply pointed_env_order_bot_L in Hpe. exact Hpe. - destruct ϕ; simpl; auto using HE with proof. + case decide; intro; subst; [constructor 2|constructor 1]. + apply make_conj_sound_L, AndR; apply AndL; auto using HE with proof. - + apply ImpR. exch 0. l_tac. apply make_impl_sound_L, ImpL; auto using HE, HA with proof. + + apply ImpR. exch 0. l_tac. foldEA. apply make_impl_sound_L, ImpL; auto using HE, HA with proof. + foldEA. apply BoxR. box_tac. exch 0. l_tac. apply make_impl_sound_L, ImpL. * apply weakening, HE. order_tac. * apply HA. order_tac. @@ -506,38 +520,22 @@ Hint Extern 5 (?a ≺· ?b) => order_tac : proof. the E-quantified version, and similarly on the other side for the A-quantifier. *) -(** E's second argument is mostly irrelevant and is only there for uniform treatment with A *) -Lemma E_irr ϕ' Δ ϕ : E (Δ, ϕ) = E (Δ, ϕ'). -Proof. -remember ((Δ, ϕ) : pointed_env) as pe. -replace Δ with pe.1 by now subst. -replace ϕ with pe.2 by now subst. clear Heqpe Δ ϕ. - induction pe as [[Γ φ] Hind0] using (well_founded_induction_type wf_pointed_order). -assert(Hind : ∀ Γ0 φ0, ((Γ0, φ0) ≺· (Γ, φ)) → E (Γ0, φ0) = E (Γ0, ϕ')) - by exact (fun Γ0 φ0 => (Hind0 (Γ0, φ0))). -unfold E in Hind. simpl in Hind. -do 2 rewrite E_eq. f_equal. apply in_map_ext. -intros φ' Hin. unfold e_rule. -destruct φ'; repeat rewrite (Hind _ φ) by (trivial; order_tac); trivial. -destruct φ'1; repeat rewrite (Hind _ φ) by (trivial; order_tac); trivial. -Qed. - -Lemma E_left {Γ} {θ} {Δ} {φ φ'} (Hin : φ ∈ Δ) : -(Γ • e_rule (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe) φ Hin) ⊢ θ -> -Γ • E (Δ, φ') ⊢ θ. +Lemma E_left {Γ} {θ} {Δ} {φ} (Hin : φ ∈ Δ) : +(Γ • e_rule (λ pe (_ : pe ≺· (Δ, ⊥)), E pe.1) (λ pe (_ : pe ≺· (Δ, ⊥)), A pe) φ Hin) ⊢ θ -> +Γ • E Δ ⊢ θ. Proof. intro Hp. rewrite E_eq. -destruct (@in_map_in _ _ _ (e_rule (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe) ) _ Hin) as [Hin' Hrule]. +destruct (@in_map_in _ _ _ (e_rule (λ pe (_ : pe ≺· (Δ, ⊥)), E pe.1) (λ pe (_ : pe ≺· (Δ, ⊥)), A pe) ) _ Hin) as [Hin' Hrule]. eapply conjunction_L. - apply Hrule. - exact Hp. Qed. Lemma A_right {Γ} {Δ} {φ φ'} (Hin : φ ∈ Δ) : -Γ ⊢ a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe) φ Hin -> +Γ ⊢ a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe.1) (λ pe (_ : pe ≺· (Δ, φ')), A pe) φ Hin -> Γ ⊢ A (Δ, φ'). Proof. intro Hp. rewrite A_eq. -destruct (@in_map_in _ _ _ (a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe) (λ pe (_ : pe ≺· (Δ, φ')), A pe)) _ Hin) as [Hin' Hrule]. +destruct (@in_map_in _ _ _ (a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe.1) (λ pe (_ : pe ≺· (Δ, φ')), A pe)) _ Hin) as [Hin' Hrule]. eapply make_disj_sound_R, OrR1, disjunction_R. - exact Hrule. - exact Hp. @@ -546,13 +544,13 @@ Qed. Proposition pq_correct Γ Δ ϕ: (∀ φ0, φ0 ∈ Γ -> ¬ occurs_in p φ0) -> (Γ ⊎ list_to_set_disj Δ ⊢ ϕ) -> - (¬ occurs_in p ϕ -> Γ•E (Δ, ϕ) ⊢ ϕ) * (Γ•E (Δ, ϕ) ⊢ A (Δ, ϕ)). + (¬ occurs_in p ϕ -> Γ•E Δ ⊢ ϕ) * (Γ•E Δ ⊢ A (Δ, ϕ)). Proof. (* This proof goes by induction on the ordering w.r.t (Γ ⊎Δ, ϕ) instead of on the structure of Γ ⊎Δ ⊢ ϕ, to allow better rules *) (* we want to use an E rule *) Local Ltac Etac := foldEA; intros; match goal with -| Hin : ?a ∈ list_to_set_disj ?Δ |- _•E (?Δ, _) ⊢ _=> +| Hin : ?a ∈ list_to_set_disj ?Δ |- _•E ?Δ ⊢ _=> apply (E_left (proj1 (elem_of_list_to_set_disj _ _) Hin)); unfold e_rule end. (* we want to use an A rule defined in a_rule_env *) @@ -628,25 +626,25 @@ end; simpl. - split; Etac; auto with proof. (* AndR *) - split. - + intro Hocc. apply AndR; erewrite E_irr; apply Hind; auto with proof. - + Atac'. foldEA. apply make_conj_sound_R, AndR; erewrite E_irr; + + intro Hocc. apply AndR; apply Hind; auto with proof. + + Atac'. foldEA. apply make_conj_sound_R, AndR; (eapply Hind; [order_tac | occ | trivial]). (* AndL *) - exch 0. apply AndL. exch 1; exch 0. peapply Hind; auto with proof. occ. peapply' Hp. - exch 0. apply AndL. exch 1; exch 0. peapply Hind; auto with proof. occ. peapply' Hp. - split. - + Etac. foldEA. apply Hind; auto with proof. peapply' Hp. + + Etac. foldEA. apply Hind; auto with proof. simpl. peapply' Hp. + Atac. Etac. apply Hind; auto; auto with proof. peapply' Hp. (* OrR1 & OrR2 *) - split. - + intro Hocc. apply OrR1. erewrite E_irr; apply Hind; auto with proof. + + intro Hocc. apply OrR1. apply Hind; auto with proof. + rewrite A_eq. apply make_disj_sound_R, OrR2. - apply make_disj_sound_R, OrR1; erewrite E_irr. + apply make_disj_sound_R, OrR1. apply Hind; auto with proof. - simpl. split. - + intro Hocc. apply OrR2. erewrite E_irr; apply Hind; auto with proof. + + intro Hocc. apply OrR2. apply Hind; auto with proof. + rewrite A_eq. apply make_disj_sound_R, OrR2. - apply make_disj_sound_R, OrR2; erewrite E_irr. + apply make_disj_sound_R, OrR2. apply Hind; auto with proof. (* OrL *) - exch 0. apply OrL; exch 0. @@ -656,7 +654,7 @@ end; simpl. + apply Hind; auto with proof. occ. peapply' Hp1. + apply Hind; auto with proof. occ. peapply' Hp2. - split. - + Etac. apply make_disj_sound_L, OrL; apply Hind; auto with proof. + + Etac. apply make_disj_sound_L, OrL; apply Hind; auto with proof; simpl. * peapply' Hp1. * peapply' Hp2. + Atac. apply weakening. apply make_conj_sound_R,AndR, make_impl_sound_R. @@ -665,7 +663,7 @@ end; simpl. (* ImpR *) - split. + intro Hocc. apply ImpR. exch 0. - erewrite E_irr. apply Hind; auto with proof. occ. peapply Hp. + apply Hind; auto with proof. occ. peapply Hp. + Atac'. apply weakening, make_impl_sound_R, ImpR, Hind; auto with proof. order_tac. * rewrite <- Permutation_middle. order_tac. * peapply Hp. @@ -702,9 +700,9 @@ end; simpl. do 2 exhibit Hin1 1. split; [intro Hocc|]. -- Etac. case decide; intro Hp0;[|case decide; intro; subst; [auto with *|]]. - ++ foldEA. apply Hind; auto with proof; [order_tac; order_tac|occ | peapply' Hp]. + ++ foldEA. simpl. apply Hind; auto with proof; [order_tac; order_tac|occ | peapply' Hp]. ++ apply make_impl_sound_L, ImpLVar. exch 0. backward. rewrite env_add_remove. - apply Hind; auto with proof. peapply' Hp. + apply Hind; auto with proof. simpl. peapply' Hp. -- Etac. Atac. case decide; intro Hp0;[|case decide; intro; subst; [auto with *|]]. ++ foldEA. apply Hind; auto with proof; [order_tac; order_tac|occ | peapply' Hp]. ++ apply make_impl_sound_L, ImpLVar. @@ -719,7 +717,7 @@ end; simpl. -- (* subsubcase p = p0 *) apply elem_of_list_to_set_disj in Hin1'. subst. split; Etac; repeat rewrite decide_True by trivial. - ++ clear Heq. apply Hind; auto with proof. peapply' Hp. + ++ clear Heq. apply Hind; auto with proof. simpl. peapply' Hp. ++ Atac. repeat rewrite decide_True by trivial. clear Heq. apply Hind; auto with proof. peapply' Hp. -- (* subsubcase p ≠ p0 *) @@ -727,7 +725,7 @@ end; simpl. ++ apply contraction. Etac. rewrite decide_False by trivial. exch 0. assert((p0 → φ) ∈ list_to_set_disj Δ) by ms. Etac. rewrite decide_True by now apply elem_of_list_to_set_disj. - exch 0. apply weakening. apply Hind; auto with proof. peapply' Hp. + exch 0. apply weakening. apply Hind; auto with proof. simpl. peapply' Hp. ++ apply contraction. Etac. exch 0. assert((p0 → φ) ∈ list_to_set_disj Δ) by ms. Etac; Atac. rewrite decide_False by trivial. do 2 rewrite decide_True by now apply elem_of_list_to_set_disj. @@ -737,40 +735,39 @@ end; simpl. - exch 0. apply ImpLAnd. exch 0. apply Hind; auto with proof; [occ|peapply' Hp]. - exch 0. apply ImpLAnd. exch 0. apply Hind; auto with proof; [occ|peapply' Hp]. - split; Etac. - + apply Hind; auto with proof. peapply' Hp. + + apply Hind; auto with proof. simpl. peapply' Hp. + Atac. simpl. apply Hind; auto with proof. peapply' Hp. (* ImpLOr *) - exch 0; apply ImpLOr. exch 1; exch 0. apply Hind; auto with proof. occ. peapply' Hp. - exch 0; apply ImpLOr. exch 1; exch 0. apply Hind; [order_tac | occ|peapply' Hp]. - split; Etac. - + apply Hind; auto with proof. peapply' Hp. + + apply Hind; auto with proof. simpl. peapply' Hp. + Atac. apply Hind; auto with proof. peapply' Hp. (* ImpLImp *) - (* subcase 1: ((φ1 → φ2) → φ3) ∈ Γ *) assert(ψ ≠ Var p) by (intro; subst; simpl in *; tauto). exch 0; apply ImpLImp; exch 0. - + erewrite E_irr. apply Hind; auto with proof. occ. peapply' Hp1. + + apply Hind; auto with proof. occ. peapply' Hp1. simpl. apply Hnin in Hin0. simpl in *. tauto. - + erewrite E_irr. apply Hind; auto with proof. occ. peapply' Hp2. + + apply Hind; auto with proof. occ. peapply' Hp2. - exch 0; apply ImpLImp; exch 0. - + erewrite E_irr. apply Hind; auto with proof. occ. peapply' Hp1. + + apply Hind; auto with proof. occ. peapply' Hp1. simpl. apply Hnin in Hin0. simpl in Hin0. tauto. + apply Hind; auto with proof. occ. peapply' Hp2. - (* subcase 2: ((φ1 → φ2) → φ3) ∈ Δ *) split. + Etac. apply make_impl_sound_L2'. apply ImpLImp. * apply weakening. apply ImpR. foldEA. - rewrite (E_irr φ2). apply Hind; auto with proof. -- order_tac. repeat rewrite Permutation_middle. order_tac. -- repeat setoid_rewrite gmultiset_disj_union_assoc. setoid_rewrite gmultiset_disj_union_comm. repeat setoid_rewrite gmultiset_disj_union_assoc. exch 0. apply ImpR_rev. peapply' Hp1. - * apply Hind; auto with proof. peapply' Hp2. + * apply Hind; auto with proof. simpl. peapply' Hp2. + Atac. apply make_conj_sound_R, AndR. * apply weakening. apply make_impl_sound_R, ImpR. foldEA. - rewrite (E_irr φ2). apply Hind; auto with proof. + apply Hind; auto with proof. -- order_tac. repeat rewrite Permutation_middle. order_tac. -- repeat setoid_rewrite gmultiset_disj_union_assoc. setoid_rewrite gmultiset_disj_union_comm. @@ -778,7 +775,7 @@ end; simpl. peapply' Hp1. * Etac. simpl. apply make_impl_sound_L2', ImpLImp. -- apply weakening. apply ImpR. foldEA. - rewrite (E_irr φ2). apply Hind; auto with proof. + apply Hind; auto with proof. ++ order_tac. repeat rewrite Permutation_middle. order_tac. ++ repeat setoid_rewrite gmultiset_disj_union_assoc. setoid_rewrite gmultiset_disj_union_comm. @@ -789,7 +786,6 @@ end; simpl. destruct (open_boxes_case (list_to_set_disj Δ)) as [[θ Hθ] | Hequiv]. + apply contraction. Etac. exch 1; exch 0; apply ImpBox. * box_tac. box_tac. exch 2; exch 1; exch 0. apply weakening. foldEA. - erewrite E_irr with (ϕ' := φ1). exch 1; exch 0. apply Hind. -- order_tac. (* cannot handle it. manual proof *) rewrite elements_open_boxes. @@ -815,9 +811,7 @@ end; simpl. rewrite <- Hequiv, open_boxes_disj_union. rewrite open_boxes_remove by trivial. ms. } exch 0. apply ImpBox. - * box_tac. exch 1; exch 0. apply open_box_L. - erewrite E_irr with (ϕ' := φ1). - apply Hind. + * box_tac. exch 1; exch 0. apply open_box_L, Hind. -- order_tac. rewrite elements_open_boxes. repeat rewrite Permutation_middle. order_tac. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. @@ -829,7 +823,6 @@ end; simpl. - destruct (open_boxes_case (list_to_set_disj Δ)) as [[θ Hθ] | Hequiv]. + apply contraction. Etac. exch 1; exch 0; apply ImpBox. * box_tac. box_tac. exch 2; exch 1; exch 0. apply weakening. foldEA. - erewrite E_irr with (ϕ' := φ1). exch 1; exch 0. apply Hind. -- order_tac. rewrite elements_open_boxes. @@ -850,8 +843,7 @@ end; simpl. rewrite open_boxes_add. simpl. ms. -- intro HF. apply (Hnin _ Hin0). simpl. tauto. * exch 0. apply weakening. exch 0. apply Hind ; [order_tac|occ|peapply' Hp2]. - + erewrite E_irr with (ϕ' := φ1). - exch 0. apply ImpBox. + + exch 0. apply ImpBox. * box_tac. exch 1; exch 0. apply open_box_L. apply Hind. -- order_tac. rewrite elements_open_boxes. order_tac. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). @@ -862,13 +854,10 @@ end; simpl. rewrite open_boxes_remove by trivial. ms. } peapply Hp1. -- intro HF. apply (Hnin _ Hin0). simpl. tauto. - * erewrite E_irr with (ϕ' := ψ). - exch 0. apply Hind. order_tac. occ. clear Hequiv. peapply' Hp2. + * exch 0. apply Hind. order_tac. occ. clear Hequiv. peapply' Hp2. - split; Etac. + foldEA. apply make_impl_sound_L, ImpBox. - -- do 2 apply weakening. apply make_impl_sound_R, ImpR. - erewrite E_irr with (ϕ' := φ1) by ms. - apply Hind. + -- do 2 apply weakening. apply make_impl_sound_R, ImpR, Hind. ++ order_tac. rewrite elements_open_boxes. do 2 apply env_order_cancel_right. repeat rewrite Permutation_middle. @@ -882,7 +871,7 @@ end; simpl. simpl. rewrite union_difference_R by auto with proof. rewrite open_boxes_remove by ms. ms. } peapply Hp1. - -- apply Hind. order_tac. occ. peapply' Hp2. trivial. + -- apply Hind. order_tac. occ. simpl. peapply' Hp2. trivial. + assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ) ⊎ list_to_set_disj (map open_box (rm (□ φ1 → φ2) Δ)))). { rewrite Heq'. repeat rewrite open_boxes_remove by ms. simpl. @@ -891,10 +880,8 @@ end; simpl. } foldEA. Atac. apply AndR. * apply make_impl_sound_L, ImpBox. - -- do 2 apply weakening. apply make_impl_sound_R, ImpR. - erewrite E_irr with (ϕ' := φ1). - apply Hind. - ++ order_tac. rewrite elements_open_boxes. + -- do 2 apply weakening. apply make_impl_sound_R, ImpR, Hind. + ++ order_tac. rewrite elements_open_boxes. do 2 apply env_order_cancel_right. repeat rewrite Permutation_middle. apply env_order_disj_union_compat_strong_left; order_tac. @@ -902,7 +889,6 @@ end; simpl. apply (Hnin θ0); ms. ++ peapply Hp1. -- apply BoxR. box_tac. do 2 apply weakening. apply make_impl_sound_R, ImpR. foldEA. - erewrite E_irr with (ϕ' := φ1) by ms. apply Hind. ++ order_tac. rewrite elements_open_boxes. do 2 apply env_order_cancel_right. @@ -912,9 +898,7 @@ end; simpl. apply (Hnin θ0); ms. ++ peapply Hp1. * foldEA. apply make_impl_sound_L, ImpBox. - -- do 2 apply weakening. apply make_impl_sound_R, ImpR. - erewrite E_irr with (ϕ' := φ1). - apply Hind. + -- do 2 apply weakening. apply make_impl_sound_R, ImpR, Hind. ++ order_tac. rewrite elements_open_boxes. do 2 apply env_order_cancel_right. repeat rewrite Permutation_middle. @@ -925,9 +909,8 @@ end; simpl. -- clear Heq''. apply Hind; [order_tac|occ|peapply' Hp2]. - split. + intro Hocc. destruct (open_boxes_case (list_to_set_disj Δ)) as [[θ Hθ] | Hequiv]. - * Etac. foldEA. apply BoxR. box_tac. exch 0. - erewrite E_irr with (ϕ' := φ); apply Hind. - -- order_tac. + * Etac. foldEA. apply BoxR. box_tac. exch 0. apply Hind. + -- order_tac. rewrite elements_open_boxes. repeat rewrite Permutation_middle. apply env_order_disj_union_compat_strong_left. order_tac. @@ -943,9 +926,7 @@ end; simpl. rewrite open_boxes_remove by ms. rewrite <- difference_singleton by auto with proof. ms. -- trivial. - * apply BoxR. box_tac. exch 0. apply open_box_L. - erewrite E_irr with (ϕ' := φ). - apply Hind. + * apply BoxR. box_tac. exch 0. apply open_box_L, Hind. -- order_tac. rewrite elements_open_boxes. repeat rewrite Permutation_middle. @@ -957,8 +938,7 @@ end; simpl. -- (erewrite proper_Provable; [| |reflexivity]); [eapply Hp|]. rewrite open_boxes_disj_union, <- Hequiv. ms. -- trivial. - + Atac'. foldEA. apply BoxR. box_tac. apply weakening. erewrite E_irr with (ϕ' := φ). - apply weakening. apply make_impl_sound_R, ImpR. + + Atac'. foldEA. apply BoxR. box_tac. apply weakening, weakening, make_impl_sound_R, ImpR. apply Hind. * order_tac. rewrite elements_open_boxes. @@ -984,7 +964,7 @@ End UniformInterpolation. Open Scope type_scope. -Lemma E_of_empty p φ : E p ([], φ) = (Implies Bot Bot). +Lemma E_of_empty p : E p [] = (Implies Bot Bot). Proof. rewrite E_eq. rewrite in_map_empty. now unfold conjunction, nodup, foldl. Qed. @@ -1006,8 +986,8 @@ intros Hp φ Hvarsφ; repeat split. + intros x Hx. apply (EA_vars p _ ⊥ x) in Hx. destruct Hx as [Hneq [θ [Hθ Hocc]]]. apply elem_of_list_singleton in Hθ. subst. apply Hvarsφ in Hocc. destruct Hocc; subst; tauto. - + replace {[φ]} with (list_to_set_disj [φ] : env). apply entail_correct. simpl. ms. - + intros ψ Hψ Hyp. rewrite elem_of_list_In in Hp. unfold Ef. rewrite E_irr with (ϕ' := ψ). + + replace {[φ]} with (list_to_set_disj [φ] : env). apply (entail_correct _ _ ⊥). ms. + + intros ψ Hψ Hyp. rewrite elem_of_list_In in Hp. unfold Ef. * peapply (pq_correct p ∅ [φ] ψ). -- intros θ Hin. inversion Hin. -- peapply Hyp. From b9fcb92dff4f0fc2fe4b808d86b6662d213d116d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Mon, 9 Sep 2024 18:01:36 +0200 Subject: [PATCH 08/12] make progress in defining a correct simplification for environments, interleaved during the computation of Pitts Interpolants --- theories/extraction/UIML_extraction.v | 3 +- theories/iSL/DecisionProcedure.v | 1 + theories/iSL/InvPropQuantifiers.v | 180 ++++++++++---------------- theories/iSL/Order.v | 11 +- theories/iSL/PropQuantifiers.v | 3 +- theories/iSL/Simp.v | 6 +- 6 files changed, 83 insertions(+), 121 deletions(-) diff --git a/theories/extraction/UIML_extraction.v b/theories/extraction/UIML_extraction.v index 1c6ec99..ba82573 100644 --- a/theories/extraction/UIML_extraction.v +++ b/theories/extraction/UIML_extraction.v @@ -5,9 +5,8 @@ Require Import ExtrOcamlBasic ExtrOcamlString. Require Import K.Interpolation.UIK_braga. Require Import KS_export. -Require Import ISL.PropQuantifiers ISL.DecisionProcedure. +Require Import ISL.InvPropQuantifiers ISL.DecisionProcedure ISL.Simp. -Require Import ISL.Simp. Fixpoint MPropF_of_form (f : Formulas.form) : MPropF := match f with diff --git a/theories/iSL/DecisionProcedure.v b/theories/iSL/DecisionProcedure.v index a7802ea..66bc1d5 100644 --- a/theories/iSL/DecisionProcedure.v +++ b/theories/iSL/DecisionProcedure.v @@ -21,6 +21,7 @@ induction l as [|x l]. * right. simpl. intros z [Hz|Hz]; subst; try rewrite Heq; auto with *. Defined. + (* This function computes a proof tree of a sequent, if there is one, or produces a proof that there is none *) Proposition Proof_tree_dec Γ φ : {_ : list_to_set_disj Γ ⊢ φ & True} + {forall H : list_to_set_disj Γ ⊢ φ, False}. diff --git a/theories/iSL/InvPropQuantifiers.v b/theories/iSL/InvPropQuantifiers.v index 21d3d14..224bc78 100644 --- a/theories/iSL/InvPropQuantifiers.v +++ b/theories/iSL/InvPropQuantifiers.v @@ -1,41 +1,12 @@ -Require Import ISL.Sequents ISL.SequentProps ISL.Order ISL.Optimizations ISL.DecisionProcedure. Require Import Coq.Program.Equality. (* for dependent induction *) - -(** * Overview: Propositional Quantifiers - -The main theorem proved in this file was first proved as Theorem 1 in: - -(Pitts 1992). A. M. Pitts. On an interpretation of second order quantification in first order intuitionistic propositional logic. J. Symb. Log., 57(1):33–52. -It has been further extended to handle iSL - -It consists of two parts: - -1) the inductive construction of the propositional quantifiers; - -2) a proof of its correctness. *) +Require Import ISL.PropQuantifiers ISL.Sequents ISL.SequentProps. +Require Import ISL.Order ISL.DecisionProcedure. Section UniformInterpolation. -(** Throughout the construction and proof, we fix a variable p, with respect to - which the propositional quantifier will be computed. *) -Variable p : variable. - - -(** * Definition of propositional quantifiers. *) - -Section PropQuantDefinition. - -(** We define the formulas Eφ and Aφ associated to any formula φ. This - is an implementation of Pitts' Table 5, together with a (mostly automatic) - proof that the definition terminates*) - -Local Notation "Δ '•' φ" := (cons φ Δ). - -Notation "□⁻¹ Γ" := (map open_box Γ) (at level 75). - Definition applicable_AndR φ: {φ1 & {φ2 | φ = (And φ1 φ2)}} + (∀ φ1 φ2, φ ≠ (And φ1 φ2)). Proof. (destruct φ; eauto). Defined. - +(* Definition applicable_other_var (Γ : list form): {q | p ≠ q /\ Var q ∈ Γ} + (∀ q, p ≠ q -> Var q ∈ Γ -> False). Proof. pose (fA := (fun θ => match θ with |Var q => if decide (p = q) then false else true | _ => false end)). @@ -46,6 +17,7 @@ Proof. - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. case decide in Hψ; auto with *. Defined. +*) Definition applicable_AndL (Γ : list form): {ψ1 & {ψ2 | (And ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False). Proof. @@ -69,6 +41,7 @@ Proof. - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. Defined. +(* Definition applicable_ImpL_other_Var (Γ : list form): {q & {ψ | q ≠ p /\ (Implies (Var q) ψ) ∈ Γ}} + (∀ q ψ, q ≠ p -> (Implies (Var q) ψ) ∈ Γ -> False). @@ -83,6 +56,7 @@ Proof. - right. intros q ψ Hp Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fIp. simpl in Hψ. rewrite decide_False in Hψ; trivial; auto with *. Defined. +*) Definition applicable_ImpLVar (Γ : list form): {q & {ψ | Var q ∈ Γ /\ (Implies (Var q) ψ) ∈ Γ}} + @@ -151,6 +125,7 @@ Proof. destruct φ; eauto with *. Defined. Definition applicable_BoxR φ: {φ' | φ = (□ φ')} + (∀ φ', φ = (□ φ') -> False). Proof. destruct φ; eauto with *. Defined. +(* Definition applicable_ImpLImp (Γ : list form): {φ1 & {φ2 & {φ3 | ((φ1 → φ2) → φ3) ∈ Γ}}} +(forall φ1 φ2 φ3, ((φ1 → φ2) → φ3) ∈ Γ -> False). Proof. @@ -160,6 +135,7 @@ Proof. destruct θ1. 1-4, 6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. Defined. +*) (* can this be replaced with generalised axiom?*) Definition applicable_Atom q (Δ : list form) ϕ : @@ -204,100 +180,76 @@ Notation "cond '?' A ':1' B" := (sumor_bind1 cond A B) (at level 150, right asso Notation "cond '?' A ':2' B" := (sumor_bind2 cond A B) (at level 150, right associativity). Notation "cond '?' A ':3' B" := (sumor_bind3 cond A B) (at level 150, right associativity). -(* solves the obligations of the following programs *) -Obligation Tactic := simpl; intros; order_tac. -Program Definition e_rule {Δ : list form } {ϕ : form} - (EA0 : ∀ bool pe (Hpe : pe ≺· (Δ, ϕ)), form) - (θ: form) (Hin : θ ∈ Δ) : form := -let E Δ H := EA0 true (Δ, ϕ) H in -let A pe0 H := EA0 false pe0 H in -let Δ' := rm θ Δ in -match θ with -| Var q → δ => if decide (p = q) then ⊤ else q → E (Δ'•δ) _ (* E5 modified *) -(* E8 modified *) -| ((δ₁→ δ₂)→ δ₃) => - (E (Δ'•(δ₂ → δ₃) • δ₁) _⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) ⇢ E (Δ'•δ₃) _ -| □ φ => □(E ((□⁻¹ Δ') • φ ) _) (* very redundant ; identical for each box *) -| (□δ1 → δ2) => (□(E((□⁻¹ Δ') • δ2 • □δ1) _ ⇢ A((□⁻¹ Δ') • δ2 • □δ1, δ1) _)) ⇢ E(Δ' • δ2) _ -| _ => ⊤ -end. +(* TODO: plan: + - reuse propquantifiers (esp, don't redefine erule); + - define propquant' as the same fixpoint as propquant, but with a simplification step + - the simplification shold look like the monadic simplifications done here, but more general + - to ensure that' it's still correct, prove that simplified sequents are "equivalent" to the original ones + - show (by induction on the definition of EA and EA') that EA equiv EA' + *) + +Local Notation "Δ '•' φ" := (cons φ Δ). -(** The implementation of the rules for defining A is separated into two pieces. - Referring to Table 5 in Pitts, the definition a_rule_env handles A1-8 and A10, - and the definition a_rule_form handles A9 and A11-13. *) -Program Definition a_rule_env {Δ : list form} {ϕ : form} - (EA0 : ∀ bool pe (Hpe : pe ≺· (Δ, ϕ)), form) - (θ: form) (Hin : θ ∈ Δ) : form := -let E Δ H := EA0 true (Δ, ϕ) H in -let A pe0 H := EA0 false pe0 H in -let Δ' := rm θ Δ in -match θ with -| Var q → δ => if decide (p = q) then ⊥ else q ⊼ A (Δ'•δ, ϕ) _ (* A4 *) -(* A6 *) -(* A8 modified*) -| ((δ₁→ δ₂)→ δ₃) => - (E (Δ'•(δ₂ → δ₃) • δ₁) _ ⇢ A (Δ'•(δ₂ → δ₃) • δ₁, δ₂) _) - ⊼ A (Δ'•δ₃, ϕ) _ -| ((□δ1) → δ2) => (□(E((□⁻¹ Δ')• δ2 • □δ1) _ ⇢ A((□⁻¹ Δ')• δ2 • □δ1, δ1) _)) ∧ A(Δ' • δ2, ϕ) _ -(* using ⊼ here breaks congruence *) -| _ => ⊥ -end. +Infix "⊢?" := Provable_dec (at level 80). -Program Definition a_rule_form {Δ : list form} {ϕ : form} - (EA0 : ∀ bool pe (Hpe : pe ≺· (Δ, ϕ)), form) := -let E pe0:= EA0 true pe0 in -let A pe0 H := EA0 false pe0 H in -match ϕ with -| Var q => - if decide (p = q) - then ⊥ - else Var q (* A9 *) -(* A12 *) -| ϕ₁ ∨ ϕ₂ => A (Δ, ϕ₁) _ ⊻ A (Δ, ϕ₂) _ -| □δ => □((E ((□⁻¹ Δ) • □δ, δ) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) -| _ => ⊥ -end. +(* Probably very costly *) +Definition applicable_strong_weakening (Γ : list form): + {φ : form | φ ∈ Γ /\ exists (_ : list_to_set_disj (rm φ Γ) ⊢ φ), True} + + (∀ φ, φ ∈ Γ -> forall (_ : list_to_set_disj (rm φ Γ) ⊢ φ), False). +Proof. +destruct (exists_dec (λ φ, if Provable_dec (rm φ Γ) φ then true else false) Γ) as [[φ [Hin Hφ]]| Hf]. +- left. exists φ; split. + + now apply elem_of_list_In. + + destruct ((rm φ Γ) ⊢? φ). trivial. contradict Hφ. +- right. intros φ Hin Hφ. apply (Hf φ). now apply elem_of_list_In. + destruct ((rm φ Γ) ⊢? φ). auto with *. tauto. +Defined. + +Obligation Tactic := (simpl; intuition order_tac). +Program Fixpoint simp_env (Δ : list form) {wf env_order Δ} : list form := + (* invertible left rules *) + applicable_AndL Δ ? λ δ₁ δ₂ _, simp_env ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂) :2 + applicable_ImpLVar Δ ? λ q ψ _, simp_env ((rm (Var q → ψ) Δ • ψ)) :2 + applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, simp_env ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)))) :3 + applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, simp_env (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃)) :3 + (* remove redundant assumptions *) + applicable_strong_weakening Δ ? λ φ _, simp_env (rm φ Δ) :1 + Δ +. +Next Obligation. apply Wf.measure_wf, wf_env_order. Qed. + +Definition simp_form φ:= ⋀ (simp_env [φ]). + +Definition pointed_env_order_refl pe1 pe2 := env_order_refl (pe1.2 :: pe1.2 :: pe1.1) (pe2.2 :: pe2.2 :: pe2.1). + +Lemma simp_env_order pe : env_order_refl (simp_env pe) pe. +Proof. +Admitted. + +Variable p : variable. Obligation Tactic := solve[simpl; intros; simpl in *; intuition; order_tac] || (try (apply Wf.measure_wf, wf_pointed_order)). -Program Fixpoint EA (b : bool) (pe : list form * form) {wf pointed_env_order pe} := - let Δ := fst pe in - let ϕ := snd pe in - let E := EA true in - let A := EA false in +Program Fixpoint EA (b : bool) (pe : pointed_env) {wf pointed_env_order pe} : form := + let Δ := simp_env pe.1 in (* because pattern matchin in lets is not broken *) + let ϕ := pe.2 in + let E pe H := EA true pe in + let A pe H := EA false pe in (* E *) if b then - if decide (⊥ ∈ Δ) then ⊥ else - applicable_other_var Δ ? λ q _, q ⊼ E (rm (Var q) Δ, ϕ) _ :1 - applicable_AndL Δ ? λ δ₁ δ₂ _, E ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂, ϕ) _ :2 - applicable_OrL Δ ? λ δ₁ δ₂ _, E ((rm (δ₁ ∨ δ₂) Δ•δ₁, ϕ)) _ ⊻E ((rm (δ₁ ∨ δ₂) Δ•δ₂, ϕ)) _ :2 - applicable_ImpLVar Δ ? λ q ψ _, E ((rm (Var q → ψ) Δ • ψ, ϕ)) _ :2 - applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, E ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)), ϕ)) _ :3 - applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, E (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃), ϕ) _ :3 - ⋀ (in_map Δ (e_rule EA)) (* the non-invertible rules *) - + ⋀ (in_map Δ (@e_rule p Δ ϕ E A)) (* A *) else - applicable_other_var Δ ? λ q _, A (rm (Var q) Δ, ϕ) _ :1 - applicable_AndL Δ ? λ δ₁ δ₂ _, A ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂, ϕ) _ :2 - applicable_OrL Δ ? λ δ₁ δ₂ _, (E ((rm (δ₁ ∨ δ₂) Δ•δ₁, ϕ)) _ ⇢ A ((rm (δ₁ ∨ δ₂) Δ•δ₁, ϕ)) _) ⊼ - (E ((rm (δ₁ ∨ δ₂) Δ•δ₂, ϕ)) _ ⇢ A ((rm (δ₁ ∨ δ₂) Δ•δ₂, ϕ)) _) :2 - applicable_ImpLVar Δ ? λ q ψ _, A (rm (Var q → ψ) Δ • ψ, ϕ) _ :2 - applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, A ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)), ϕ)) _ :3 - applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, A (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃), ϕ) _ :3 - applicable_Atom p Δ ϕ ? λ _, ⊤ :0 - applicable_AndR ϕ ? λ ϕ₁ ϕ₂ _, A (Δ, ϕ₁) _ ⊼ A (Δ, ϕ₂) _ :2 - applicable_ImpR ϕ ? λ ϕ₁ ϕ₂ _, E (Δ • ϕ₁, ϕ₂) _ ⇢ A (Δ • ϕ₁, ϕ₂) _ :2 - (⋁ (in_map Δ (a_rule_env EA)) ⊻ a_rule_form EA) (* the non-invertible rules *) -. + ⋁ (in_map Δ (@a_rule_env p Δ ϕ E A)) ⊻ @a_rule_form p Δ ϕ E A. +Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. +order_tac. do 2 apply env_order_eq_add. apply simp_env_order. Qed. +Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. +order_tac. do 2 apply env_order_eq_add. apply simp_env_order. Qed. -Definition E pe := EA true pe. -Definition A pe := EA false pe. +Definition E := EA true. +Definition A := EA false. Definition Ef (ψ : form) := E ([ψ], ⊥). Definition Af (ψ : form) := A ([], ψ). -End PropQuantDefinition. - End UniformInterpolation. - diff --git a/theories/iSL/Order.v b/theories/iSL/Order.v index b640f40..80442f4 100644 --- a/theories/iSL/Order.v +++ b/theories/iSL/Order.v @@ -34,7 +34,7 @@ Qed. Definition env_order_refl Δ Δ' := (Δ ≺ Δ') \/ Δ ≡ₚ Δ'. -Local Notation "Δ ≼ Δ'" := (env_order_refl Δ Δ') (at level 150). +Global Notation "Δ ≼ Δ'" := (env_order_refl Δ Δ') (at level 150). Global Instance Proper_env_weight: Proper ((≡ₚ) ==> (=)) env_weight. Proof. @@ -392,6 +392,15 @@ Qed. Global Hint Resolve openboxes_env_order : order. +Lemma weight_Arrow_1 φ ψ : 1 < weight (φ → ψ). +Proof. simpl. pose (weight_pos φ). lia. Qed. + +Lemma weight_Box_1 φ: 1 < weight (□ φ). +Proof. simpl. pose (weight_pos φ). lia. Qed. + +Global Hint Resolve weight_Arrow_1 : order. +Global Hint Resolve weight_Box_1 : order. + Ltac order_tac := try unfold pointed_env_ms_order; prepare_order; repeat rewrite elements_env_add; diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index 70507e8..2cddd1e 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -1,7 +1,8 @@ -Require Import ISL.Sequents. +Require Import ISL.Sequents ISL.Formulas. Require Import ISL.SequentProps ISL.Order ISL.Optimizations. Require Import Coq.Program.Equality. (* for dependent induction *) + (** * Overview: Propositional Quantifiers The main theorem proved in this file was first proved as Theorem 1 in: diff --git a/theories/iSL/Simp.v b/theories/iSL/Simp.v index 3126de4..823a65b 100644 --- a/theories/iSL/Simp.v +++ b/theories/iSL/Simp.v @@ -527,10 +527,10 @@ match goal with end. Qed. -Require Import ISL.PropQuantifiers. +Require Import ISL.InvPropQuantifiers. -Definition E_simplified (p: variable) (ψ: form) := simp (E p ([ψ], ⊥)). -Definition A_simplified (p: variable) (ψ: form) := simp (Af p (ψ)). +Definition E_simplified (p: variable) (ψ: form) := simp_form (E p ([ψ], ⊥)). +Definition A_simplified (p: variable) (ψ: form) := simp_form (Af p (ψ)). (* Lemma bot_vars_incl V: vars_incl ⊥ V. From c3ceeca7b88cbeb5adea7f83e78a540bb145d2b0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Tue, 10 Sep 2024 21:20:21 +0200 Subject: [PATCH 09/12] stash progress on simplification --- theories/iSL/InvPropQuantifiers.v | 87 ++++++++++- theories/iSL/Order.v | 15 ++ theories/iSL/PropQuantifiers.v | 204 ++++++++++++------------- theories/iSL/Simp.v | 2 +- theories/iSL/Simp_env.v | 243 ++++++++++++++++++++++++++++++ 5 files changed, 445 insertions(+), 106 deletions(-) create mode 100644 theories/iSL/Simp_env.v diff --git a/theories/iSL/InvPropQuantifiers.v b/theories/iSL/InvPropQuantifiers.v index 224bc78..983ad70 100644 --- a/theories/iSL/InvPropQuantifiers.v +++ b/theories/iSL/InvPropQuantifiers.v @@ -1,7 +1,7 @@ Require Import Coq.Program.Equality. (* for dependent induction *) Require Import ISL.PropQuantifiers ISL.Sequents ISL.SequentProps. Require Import ISL.Order ISL.DecisionProcedure. - +Require Import Coq.Classes.RelationClasses. Section UniformInterpolation. Definition applicable_AndR φ: {φ1 & {φ2 | φ = (And φ1 φ2)}} + (∀ φ1 φ2, φ ≠ (And φ1 φ2)). @@ -245,11 +245,90 @@ order_tac. do 2 apply env_order_eq_add. apply simp_env_order. Qed. Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. order_tac. do 2 apply env_order_eq_add. apply simp_env_order. Qed. - -Definition E := EA true. +Definition E Δ := EA true (Δ, ⊥). Definition A := EA false. -Definition Ef (ψ : form) := E ([ψ], ⊥). +Lemma EA_eq Δ ϕ: let Δ' := simp_env Δ in + (E Δ = ⋀ (in_map Δ' (@e_rule p Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) /\ + (A (Δ, ϕ) = (⋁ (in_map Δ' (@a_rule_env p Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) ⊻ + @a_rule_form p Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)). +Proof. +simpl. unfold E, A, EA. simpl. +unfold EA_func at 1. +rewrite -> Wf.Fix_eq; simpl. +- split; trivial. + unfold EA_func at 1. rewrite -> Wf.Fix_eq; simpl; trivial. + intros [[|] Δ'] f g Heq; simpl; f_equal. + + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. + + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; + now rewrite Heq. + + apply a_rule_form_cong; intros; now rewrite Heq. +- intros [[|] Δ'] f g Heq; simpl; f_equal. + + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. + + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; + now rewrite Heq. + + apply a_rule_form_cong; intros; now rewrite Heq. +Qed. + +Definition Ef (ψ : form) := E [ψ]. Definition Af (ψ : form) := A ([], ψ). End UniformInterpolation. + +Section Equivalence. + +Definition equiv_form φ ψ : Type := (φ ≼ ψ) * (ψ ≼ φ). + +Definition equiv_env Δ Γ: Set := (list_to_set_disj Δ ⊢ ⋀ Γ) * (list_to_set_disj Γ ⊢ ⋀ Δ). + +Lemma symmetric_equiv_env Δ Γ: equiv_env Δ Γ -> equiv_env Γ Δ . +Proof. intros [H1 H2]. split; trivial. Qed. + +Lemma equiv_env_L1 Γ Δ Δ' φ: (equiv_env Δ Δ') -> + Γ ⊎ list_to_set_disj Δ ⊢ φ -> Γ ⊎ list_to_set_disj Δ' ⊢ φ. +Proof. +intros [H1 H2]. +Admitted. + +Lemma equiv_env_L2 Γ Δ Δ' φ: (equiv_env Δ Δ') -> + list_to_set_disj Δ ⊎ Γ ⊢ φ -> list_to_set_disj Δ' ⊎ Γ ⊢ φ. +Proof. +Admitted. + +End Equivalence. + +Infix "≡f" := equiv_form (at level 120). + +Global Infix "≡e" := equiv_env (at level 120). +(* +Section Soundness. + +Variable p : variable. + +Proposition entail_correct_equiv_env Δ Δ' ϕ : (Δ ≡e Δ') -> + (list_to_set_disj Δ ⊢ PropQuantifiers.E p Δ') * + (list_to_set_disj Δ • PropQuantifiers.A p (Δ', ϕ) ⊢ ϕ). +Proof. +intro Heq. split. +- peapply (equiv_env_L1 gmultiset_empty Δ' Δ); eauto. now apply symmetric_equiv_env. + peapply (entail_correct p Δ' ϕ).1 . +- peapply (equiv_env_L2 {[PropQuantifiers.A p (Δ', ϕ)]} Δ' Δ); eauto. now apply symmetric_equiv_env. + peapply (entail_correct p Δ' ϕ).2. +Qed. + +Lemma pq_correct_equiv_env Γ Δ ϕ Δ': + (∀ φ0, φ0 ∈ Γ -> ¬ occurs_in p φ0) -> + (Γ ⊎ list_to_set_disj Δ ⊢ ϕ) -> + (Δ ≡e Δ') -> + (¬ occurs_in p ϕ -> Γ• PropQuantifiers.E p Δ' ⊢ ϕ) * + (Γ• PropQuantifiers.E p Δ' ⊢ PropQuantifiers.A p (Δ', ϕ)). +Proof. +intros Hocc Hp Heq. split. +- intro Hocc'. apply pq_correct; trivial. eapply equiv_env_L1; eauto. +- apply pq_correct; trivial. eapply equiv_env_L1; eauto. +Qed. + + + +End Soundness. +*) \ No newline at end of file diff --git a/theories/iSL/Order.v b/theories/iSL/Order.v index 80442f4..e780708 100644 --- a/theories/iSL/Order.v +++ b/theories/iSL/Order.v @@ -406,3 +406,18 @@ try unfold pointed_env_ms_order; prepare_order; repeat rewrite elements_env_add; simpl; auto 10 with order; try (apply env_order_disj_union_compat_right; order_tac). + +Global Hint Extern 5 (?a ≺ ?b) => order_tac : proof. +Hint Extern 5 (?a ≺· ?b) => order_tac : proof. + +Lemma pointed_env_order_bot_R pe Δ φ: (pe ≺· (Δ, ⊥)) -> pe ≺· (Δ, φ). +Proof. +Admitted. + +Hint Resolve pointed_env_order_bot_R : order. + +Lemma pointed_env_order_bot_L pe Δ φ: ((Δ, φ) ≺· pe) -> (Δ, ⊥) ≺· pe. +Proof. +Admitted. + +Hint Resolve pointed_env_order_bot_L : order. diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index 2cddd1e..b03fe6c 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -1,6 +1,7 @@ Require Import ISL.Sequents ISL.Formulas. Require Import ISL.SequentProps ISL.Order ISL.Optimizations. Require Import Coq.Program.Equality. (* for dependent induction *) +Require Import ISL.Simp_env. (** * Overview: Propositional Quantifiers @@ -136,14 +137,19 @@ match ϕ with | □δ => □((E ((□⁻¹ Δ) • □δ, ⊥) _) ⇢ A((□⁻¹ Δ) • □δ, δ) _) end. -Obligation Tactic := intros; order_tac. -Program Fixpoint EA (b : bool) (pe : list form * form) {wf pointed_env_order pe} : form := - let Δ := pe.1 in - let E := EA true in - let A := EA false in - if b then⋀ (in_map Δ (e_rule E A)) - else ⋁ (in_map Δ (a_rule_env E A)) ⊻ a_rule_form E A. -Next Obligation. apply Wf.measure_wf, wf_pointed_order. Defined. + +Obligation Tactic := (try (apply Wf.measure_wf, wf_pointed_order)). +Program Fixpoint EA (b : bool) (pe : pointed_env) {wf pointed_env_order pe} : form := + let Δ := simp_env pe.1 in (* new: weight-decreasing simplification *) + let ϕ := pe.2 in + let E pe H := EA true pe in + let A pe H := EA false pe in + (* E *) + if b then ⋀ (in_map Δ (@e_rule Δ ϕ E A)) + (* A *) + else ⋁ (in_map Δ (@a_rule_env Δ ϕ E A)) ⊻ @a_rule_form Δ ϕ E A. +Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. order_tac. Qed. +Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. order_tac. Qed. Definition E Δ:= EA true (Δ, ⊥). Definition A := EA false. @@ -220,26 +226,13 @@ Proof. repeat (erewrite ?HeqE, ?HeqA; eauto); trivial. Qed. -Lemma EA_eq Δ ϕ: - (E Δ = ⋀ (in_map Δ (@e_rule Δ ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) /\ - (A (Δ, ϕ) = (⋁ (in_map Δ (@a_rule_env Δ ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) ⊻ - @a_rule_form Δ ϕ (λ pe _, E pe.1) (λ pe _, A pe)). +Lemma EA_eq Δ ϕ: let Δ' := simp_env Δ in + (E Δ = ⋀ (in_map Δ' (@e_rule Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) /\ + (A (Δ, ϕ) = (⋁ (in_map Δ' (@a_rule_env Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) ⊻ + @a_rule_form Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)). Proof. -simpl. unfold E, A, EA. simpl. -unfold EA_func at 1. -rewrite -> Wf.Fix_eq; simpl. -- split; trivial. - unfold EA_func at 1. rewrite -> Wf.Fix_eq; simpl; trivial. - intros [[|] Δ'] f g Heq; simpl; f_equal. - + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. - + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; - now rewrite Heq. - + apply a_rule_form_cong; intros; now rewrite Heq. -- intros [[|] Δ'] f g Heq; simpl; f_equal. - + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. - + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; - now rewrite Heq. - + apply a_rule_form_cong; intros; now rewrite Heq. +simpl. unfold E, A, EA, EA_func. +repeat rewrite Wf.WfExtensionality.fix_sub_eq_ext; simpl. split; trivial. Qed. Definition E_eq Δ := proj1 (EA_eq Δ ⊥). @@ -331,18 +324,6 @@ Proof. destruct ϕ; unfold a_rule_form; simpl; try tauto; try solve [repeat case decide; repeat vars_tac]. Qed. -Lemma pointed_env_order_bot_R pe Δ φ: (pe ≺· (Δ, ⊥)) -> pe ≺· (Δ, φ). -Proof. -Admitted. - -Hint Resolve pointed_env_order_bot_R : order. - -Lemma pointed_env_order_bot_L pe Δ φ: ((Δ, φ) ≺· pe) -> (Δ, ⊥) ≺· pe. -Proof. -Admitted. - -Hint Resolve pointed_env_order_bot_L : order. - Proposition EA_vars Δ ϕ x: (occurs_in x (E Δ) -> x <> p /\ ∃ θ, θ ∈ Δ /\ occurs_in x θ) /\ (occurs_in x (A (Δ, ϕ)) -> x <> p /\ (occurs_in x ϕ \/ (∃ θ, θ ∈ Δ /\ occurs_in x θ))). @@ -358,18 +339,21 @@ split. - intros Hocc. apply variables_conjunction in Hocc as (φ&Hin&Hφ). apply in_in_map in Hin as (ψ&Hin&Heq). subst φ. apply e_rule_vars in Hφ. - + trivial. - + intros; apply Hind. now apply pointed_env_order_bot_R. + + intuition. now apply equiv_env_vars. + + intros; apply Hind. auto with order. (* A *) - intro Hocc. apply occurs_in_make_disj in Hocc as [Hocc|Hocc]. (* disjunction *) + apply variables_disjunction in Hocc as (φ&Hin&Hφ). apply in_in_map in Hin as (ψ&Hin&Heq). subst φ. - apply a_rule_env_vars in Hφ; trivial. + apply a_rule_env_vars in Hφ. + * intuition. right. now apply equiv_env_vars. + * intros pe Hpe. apply simp_env_pointed_env_order_L in Hpe. apply Hind in Hpe. tauto. (* pointer rule *) + apply a_rule_form_vars in Hocc. * destruct Hocc as [Hneq [Hocc | Hocc]]; vars_tac. - * vars_tac; apply Hind in H; trivial; tauto. + right. apply equiv_env_vars. vars_tac. + * vars_tac; apply Hind in H; trivial; intuition; auto with order. Qed. End VariablesCorrect. @@ -386,7 +370,6 @@ Section EntailmentCorrect. entailed by the original formula and entail the original formula, respectively. *) -Hint Extern 5 (?a ≺ ?b) => order_tac : proof. Opaque make_disj. Opaque make_conj. @@ -453,15 +436,20 @@ refine (@well_founded_induction _ _ wf_pointed_order _ _). unfold pointed_env_order. intros (Δ, ϕ) Hind. simpl. rewrite E_eq, A_eq. +remember (simp_env Δ) as Δ'. +assert (Hind' := λ y H, Hind y (simp_env_pointed_env_order_L y Δ ϕ H)). +clear Hind. (* uncurry the induction hypothesis for convenience *) -assert (HE := fun d x=> fst (Hind (d, ⊥) x)). -assert (HA := fun d f x=> snd (Hind (d, f) x)). +assert (HE := fun d x=> fst (Hind' (d, ⊥) x)). +assert (HA := fun d f x=> snd (Hind' (d, f) x)). unfold E, A in *; simpl in HE, HA. -simpl in *. clear Hind. +simpl in *. clear Hind'. split. { (* E *) apply conjunction_R1. intros φ Hin. apply in_in_map in Hin. -destruct Hin as (ψ&Hin&Heq). subst. +destruct Hin as (ψ&Hin&Heq). subst φ. +apply (equiv_env_L0 Δ'); [subst Δ'; apply simp_env_equiv_env|]. +rewrite <- HeqΔ' in *. clear HeqΔ' Δ. rename Δ' into Δ. assert(Hi : ψ ∈ list_to_set_disj Δ) by now apply elem_of_list_to_set_disj. destruct ψ; unfold e_rule; exhibit Hi 0; rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_rm _ _)) _ _ eq_refl). - case decide; intro; subst; simpl; auto using HE with proof. @@ -494,6 +482,8 @@ destruct ψ; unfold e_rule; exhibit Hi 0; rewrite (proper_Provable _ _ (equiv_di } (* A *) +apply (equiv_env_L2 _ Δ'); [subst Δ'; apply simp_env_equiv_env|]. +rewrite <- HeqΔ' in *. clear HeqΔ' Δ. rename Δ' into Δ. apply make_disj_sound_L, OrL. - apply disjunction_L. intros φ Hin. apply in_in_map in Hin as (φ' & Heq & Hφ'). subst φ. @@ -513,30 +503,35 @@ End EntailmentCorrect. (** ** Uniformity *) Section PropQuantCorrect. -Hint Extern 5 (?a ≺· ?b) => order_tac : proof. - (** The proof in this section, which is the most complex part of the argument, shows that the E- and A-formulas constructed above are indeed their propositionally quantified versions, that is, *any* formula entailed by the original formula and using only variables from that formula except p is already a consequence of the E-quantified version, and similarly on the other side for the A-quantifier. *) -Lemma E_left {Γ} {θ} {Δ} {φ} (Hin : φ ∈ Δ) : -(Γ • e_rule (λ pe (_ : pe ≺· (Δ, ⊥)), E pe.1) (λ pe (_ : pe ≺· (Δ, ⊥)), A pe) φ Hin) ⊢ θ -> -Γ • E Δ ⊢ θ. -Proof. -intro Hp. rewrite E_eq. -destruct (@in_map_in _ _ _ (e_rule (λ pe (_ : pe ≺· (Δ, ⊥)), E pe.1) (λ pe (_ : pe ≺· (Δ, ⊥)), A pe) ) _ Hin) as [Hin' Hrule]. +(* This holds by idempotence of simp_env *) +Lemma A_simp_env Δ φ: A(simp_env Δ, φ) = A (Δ, φ). +Proof. do 2 rewrite A_eq. now rewrite simp_env_idempotent. Qed. + +Lemma E_simp_env Δ: E(simp_env Δ) = E (Δ). +Proof. do 2 rewrite E_eq. now rewrite simp_env_idempotent. Qed. + +Lemma E_left {Γ} {θ} {Δ' Δ : list form} {φ : form}: (Δ' = simp_env Δ) -> ∀ (Hin : φ ∈ Δ'), +(Γ • e_rule (λ pe (_ : pe ≺· (Δ', ⊥)), E pe.1) (λ pe (_ : pe ≺· (Δ', ⊥)), A pe) φ Hin) ⊢ θ -> +Γ • E Δ' ⊢ θ. +Proof. +intros Heq Hin Hp. subst Δ'. rewrite E_simp_env, E_eq. +destruct (@in_map_in _ _ _ (e_rule (λ pe (_ : pe ≺· (simp_env Δ, ⊥)), E pe.1) (λ pe (_ : pe ≺· (simp_env Δ, ⊥)), A pe) ) _ Hin) as [Hin' Hrule]. eapply conjunction_L. - apply Hrule. - exact Hp. Qed. -Lemma A_right {Γ} {Δ} {φ φ'} (Hin : φ ∈ Δ) : -Γ ⊢ a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe.1) (λ pe (_ : pe ≺· (Δ, φ')), A pe) φ Hin -> -Γ ⊢ A (Δ, φ'). -Proof. intro Hp. rewrite A_eq. -destruct (@in_map_in _ _ _ (a_rule_env (λ pe (_ : pe ≺· (Δ, φ')), E pe.1) (λ pe (_ : pe ≺· (Δ, φ')), A pe)) _ Hin) as [Hin' Hrule]. +Lemma A_right {Γ} {Δ Δ'} {φ φ'} : (Δ' = simp_env Δ) -> ∀ (Hin : φ ∈ Δ'), +Γ ⊢ a_rule_env (λ pe (_ : pe ≺· (Δ', φ')), E pe.1) (λ pe (_ : pe ≺· (Δ', φ')), A pe) φ Hin -> +Γ ⊢ A (Δ', φ'). +Proof. intros Heq Hin Hp. subst Δ'. rewrite A_simp_env, A_eq. +destruct (@in_map_in _ _ _ (a_rule_env (λ pe (_ : pe ≺· (simp_env Δ, φ')), E pe.1) (λ pe (_ : pe ≺· (simp_env Δ, φ')), A pe)) _ Hin) as [Hin' Hrule]. eapply make_disj_sound_R, OrR1, disjunction_R. - exact Hrule. - exact Hp. @@ -551,12 +546,13 @@ Proof. instead of on the structure of Γ ⊎Δ ⊢ ϕ, to allow better rules *) (* we want to use an E rule *) Local Ltac Etac := foldEA; intros; match goal with -| Hin : ?a ∈ list_to_set_disj ?Δ |- _•E ?Δ ⊢ _=> - apply (E_left (proj1 (elem_of_list_to_set_disj _ _) Hin)); unfold e_rule end. +| HeqΔ': ?Δ' = simp_env ?Δ, Hin : ?a ∈ list_to_set_disj ?Δ' |- _•E ?Δ' ⊢ _=> + apply (E_left HeqΔ'(proj1 (elem_of_list_to_set_disj _ _) Hin)); unfold e_rule end. (* we want to use an A rule defined in a_rule_env *) -Local Ltac Atac := foldEA; match goal with | Hin : ?a ∈ list_to_set_disj ?Δ |- _ ⊢ A (?Δ, _) => - apply (A_right (proj1 (elem_of_list_to_set_disj _ _) Hin)); unfold a_rule_env end. +Local Ltac Atac := foldEA; match goal with +| HeqΔ': ?Δ' = simp_env ?Δ, Hin : ?a ∈ list_to_set_disj ?Δ' |- _ ⊢ A (?Δ', _) => + apply (A_right HeqΔ' (proj1 (elem_of_list_to_set_disj _ _) Hin)); unfold a_rule_env end. (* we want to use an A rule defined in a_rule_form *) Local Ltac Atac' := foldEA; rewrite A_eq; apply make_disj_sound_R, OrR2; simpl. @@ -597,19 +593,25 @@ revert pe Γ Δ ϕ Heqpe. refine (@well_founded_induction _ _ wf_pointed_order _ _). intros (ΓΔ, ϕ0) Hind Γ Δ ϕ Heq Hnin Hp. inversion Heq; subst; clear Heq. -assert(Hind' := λ Γ0 Δ0 ψ Ho, Hind (elements Γ0 ++ Δ0, ψ) Ho _ _ _ eq_refl). +assert(Ho' : pointed_env_order_refl (elements Γ ++ simp_env Δ, ϕ) (elements Γ ++ Δ, ϕ)) + by (unfold pointed_env_order_refl; simpl; order_tac). +assert(Hind' := λ Γ0 Δ0 ψ Ho, Hind (elements Γ0 ++ Δ0, ψ) (env_order_lt_le_trans _ _ _ Ho Ho') _ _ _ eq_refl). simpl in Hind'. clear Ho'. clear Hind. rename Hind' into Hind. +rewrite <- E_simp_env, <- A_simp_env. +assert(Hp' := equiv_env_L1 Γ _ _ _ (symmetric_equiv_env _ _ (simp_env_equiv_env Δ)) Hp). +remember (simp_env Δ) as Δ'. clear Hp. +rename Hp' into Hp. dependent destruction Hp; (* try and solve the easy case where the main formula is on the left *) try match goal with -| H : (?Γ0•?a•?b) = Γ ⊎ list_to_set_disj Δ |- _ => rename H into Heq; +| H : (?Γ0•?a•?b) = Γ ⊎ list_to_set_disj Δ' |- _ => rename H into Heq; pose(Heq' := env_equiv_eq _ _ Heq); apply env_add_inv' in Heq' -| H : ((?Γ0 • ?a) = Γ ⊎ list_to_set_disj Δ) |- _ => rename H into Heq; - assert(Hin : a ∈ (Γ ⊎ list_to_set_disj Δ)) by (rewrite <- Heq; ms); +| H : ((?Γ0 • ?a) = Γ ⊎ list_to_set_disj Δ') |- _ => rename H into Heq; + assert(Hin : a ∈ (Γ ⊎ list_to_set_disj Δ')) by (rewrite <- Heq; ms); pose(Heq' := env_equiv_eq _ _ Heq); apply env_add_inv' in Heq'; try (case (decide (a ∈ Γ)); intro Hin0; [split; intros; exhibit Hin0 1| - case (decide (a ∈ (list_to_set_disj Δ : env))); intro Hin0'; + case (decide (a ∈ (list_to_set_disj Δ' : env))); intro Hin0'; [|apply gmultiset_elem_of_disj_union in Hin; exfalso; tauto]]) end; simpl. @@ -628,7 +630,7 @@ end; simpl. (* AndR *) - split. + intro Hocc. apply AndR; apply Hind; auto with proof. - + Atac'. foldEA. apply make_conj_sound_R, AndR; + + Atac'. foldEA. do 2 rewrite A_simp_env. apply make_conj_sound_R, AndR; (eapply Hind; [order_tac | occ | trivial]). (* AndL *) - exch 0. apply AndL. exch 1; exch 0. peapply Hind; auto with proof. occ. peapply' Hp. @@ -640,13 +642,13 @@ end; simpl. - split. + intro Hocc. apply OrR1. apply Hind; auto with proof. + rewrite A_eq. apply make_disj_sound_R, OrR2. - apply make_disj_sound_R, OrR1. + apply make_disj_sound_R, OrR1. rewrite A_simp_env. apply Hind; auto with proof. - simpl. split. + intro Hocc. apply OrR2. apply Hind; auto with proof. + rewrite A_eq. apply make_disj_sound_R, OrR2. - apply make_disj_sound_R, OrR2. - apply Hind; auto with proof. + apply make_disj_sound_R, OrR2. rewrite A_simp_env. + apply Hind; auto with proof. (* OrL *) - exch 0. apply OrL; exch 0. + apply Hind; auto with proof. occ. peapply' Hp1. @@ -667,7 +669,7 @@ end; simpl. apply Hind; auto with proof. occ. peapply Hp. + Atac'. apply weakening, make_impl_sound_R, ImpR, Hind; auto with proof. order_tac. * rewrite <- Permutation_middle. order_tac. - * peapply Hp. + * simpl. subst Δ'. rewrite simp_env_idempotent. peapply Hp. (* ImpLVar *) - pose(Heq'' := Heq'); apply env_add_inv' in Heq''. case (decide ((Var p0 → φ) ∈ Γ)). @@ -680,11 +682,11 @@ end; simpl. split; [intro Hocc|]; exhibit Hin0 1; exhibit Hin2 2; exch 0; exch 1; apply ImpLVar; exch 1; exch 0; (apply Hind; auto with proof; [occ | peapply' Hp]). * assert(Hin0' : Var p0 ∈ (Γ0•Var p0•(p0 → φ))) by ms. rewrite Heq in Hin0'. - case (decide (Var p0 ∈ (list_to_set_disj Δ: env))); intro Hp0; + case (decide (Var p0 ∈ (list_to_set_disj Δ': env))); intro Hp0; [|apply gmultiset_elem_of_disj_union in Hin0'; exfalso; tauto]. (* subcase 3: p0 ∈ Δ ; (p0 → φ) ∈ Γ *) clear Heq''. - split; try case decide; intros; subst. + split; try case decide; intros. -- apply contraction. Etac. rewrite decide_False by tauto. exhibit Hin0 2. exch 1. exch 0. apply ImpLVar. exch 0. apply weakening. exch 0. apply Hind; auto with proof. occ. peapply' Hp. @@ -694,40 +696,40 @@ end; simpl. + intro. assert(Hin : (Var p0 → φ) ∈ (Γ0•Var p0•(p0 → φ))) by ms. rewrite Heq in Hin. - case (decide ((Var p0 → φ) ∈ (list_to_set_disj Δ : env))); intro Hin0; + case (decide ((Var p0 → φ) ∈ (list_to_set_disj Δ' : env))); intro Hin0; [|apply gmultiset_elem_of_disj_union in Hin; exfalso; tauto]. case (decide (Var p0 ∈ Γ)); intro Hin1. * (* subcase 2: p0 ∈ Γ ; (p0 → φ) ∈ Δ *) do 2 exhibit Hin1 1. split; [intro Hocc|]. -- Etac. case decide; intro Hp0;[|case decide; intro; subst; [auto with *|]]. - ++ foldEA. simpl. apply Hind; auto with proof; [order_tac; order_tac|occ | peapply' Hp]. + ++ foldEA. simpl. apply Hind; auto with proof. occ. peapply' Hp. ++ apply make_impl_sound_L, ImpLVar. exch 0. backward. rewrite env_add_remove. apply Hind; auto with proof. simpl. peapply' Hp. -- Etac. Atac. case decide; intro Hp0;[|case decide; intro; subst; [auto with *|]]. - ++ foldEA. apply Hind; auto with proof; [order_tac; order_tac|occ | peapply' Hp]. + ++ foldEA. apply Hind; auto with proof; [occ | peapply' Hp]. ++ apply make_impl_sound_L, ImpLVar. apply make_conj_sound_R, AndR; auto 2 with proof. exch 0. backward. rewrite env_add_remove. apply Hind; auto with proof. peapply' Hp. - * assert(Hin': Var p0 ∈ Γ ⊎ list_to_set_disj Δ) by (rewrite <- Heq; ms). + * assert(Hin': Var p0 ∈ Γ ⊎ list_to_set_disj Δ') by (rewrite <- Heq; ms). apply gmultiset_elem_of_disj_union in Hin'. - case (decide (Var p0 ∈ (list_to_set_disj Δ: env))); intro Hin1'; [|exfalso; tauto]. + case (decide (Var p0 ∈ (list_to_set_disj Δ': env))); intro Hin1'; [|exfalso; tauto]. (* subcase 4: p0,(p0 → φ) ∈ Δ *) case (decide (p = p0)); intro. -- (* subsubcase p = p0 *) apply elem_of_list_to_set_disj in Hin1'. - subst. split; Etac; repeat rewrite decide_True by trivial. + split; Etac; repeat rewrite decide_True by trivial. ++ clear Heq. apply Hind; auto with proof. simpl. peapply' Hp. ++ Atac. repeat rewrite decide_True by trivial. clear Heq. apply Hind; auto with proof. peapply' Hp. -- (* subsubcase p ≠ p0 *) split; [intro Hocc|]. ++ apply contraction. Etac. rewrite decide_False by trivial. exch 0. - assert((p0 → φ) ∈ list_to_set_disj Δ) by ms. Etac. + assert((p0 → φ) ∈ list_to_set_disj Δ') by ms. Etac. rewrite decide_True by now apply elem_of_list_to_set_disj. exch 0. apply weakening. apply Hind; auto with proof. simpl. peapply' Hp. - ++ apply contraction. Etac. exch 0. assert((p0 → φ) ∈ list_to_set_disj Δ) by ms. + ++ apply contraction. Etac. exch 0. assert((p0 → φ) ∈ list_to_set_disj Δ') by ms. Etac; Atac. rewrite decide_False by trivial. do 2 rewrite decide_True by now apply elem_of_list_to_set_disj. exch 0. apply weakening. @@ -784,7 +786,7 @@ end; simpl. peapply' Hp1. -- apply Hind; auto with proof. peapply' Hp2. - (* ImpBox, external *) - destruct (open_boxes_case (list_to_set_disj Δ)) as [[θ Hθ] | Hequiv]. + destruct (open_boxes_case (list_to_set_disj Δ')) as [[θ Hθ] | Hequiv]. + apply contraction. Etac. exch 1; exch 0; apply ImpBox. * box_tac. box_tac. exch 2; exch 1; exch 0. apply weakening. foldEA. exch 1; exch 0. apply Hind. @@ -795,7 +797,7 @@ end; simpl. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. -- assert(Hθ' := In_open_boxes _ _ Hθ). - assert(Heq'' : Γ0 ≡ (Γ ∖ {[□ φ1 → φ2]} ⊎ list_to_set_disj ((□θ) :: rm (□ θ) Δ))). { + assert(Heq'' : Γ0 ≡ (Γ ∖ {[□ φ1 → φ2]} ⊎ list_to_set_disj ((□θ) :: rm (□ θ) Δ'))). { rewrite Heq'. simpl. rewrite union_difference_L by trivial. rewrite (difference_singleton _ _ Hθ). rewrite list_to_set_disj_rm. ms. @@ -807,7 +809,7 @@ end; simpl. rewrite open_boxes_add. simpl. ms. -- intro HF. apply (Hnin _ Hin0). simpl. tauto. * exch 0. apply weakening. exch 0. apply Hind; [order_tac | occ|peapply' Hp2| trivial]. - + assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ ∖ {[□ φ1 → φ2]}) ⊎ ⊗ (list_to_set_disj Δ))). { + + assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ ∖ {[□ φ1 → φ2]}) ⊎ ⊗ (list_to_set_disj Δ'))). { rewrite Heq'. rewrite union_difference_L by trivial. rewrite <- Hequiv, open_boxes_disj_union. rewrite open_boxes_remove by trivial. ms. } @@ -821,7 +823,7 @@ end; simpl. * exch 0. apply Hind; [order_tac|occ | | trivial]. (erewrite proper_Provable; [| |reflexivity]); [eapply Hp2|]. rewrite Heq'. rewrite union_difference_L by trivial. ms. -- destruct (open_boxes_case (list_to_set_disj Δ)) as [[θ Hθ] | Hequiv]. +- destruct (open_boxes_case (list_to_set_disj Δ')) as [[θ Hθ] | Hequiv]. + apply contraction. Etac. exch 1; exch 0; apply ImpBox. * box_tac. box_tac. exch 2; exch 1; exch 0. apply weakening. foldEA. exch 1; exch 0. apply Hind. @@ -832,7 +834,7 @@ end; simpl. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. -- assert(Hθ' := In_open_boxes _ _ Hθ). - assert(Heq'' : Γ0 ≡ (Γ ∖ {[□ φ1 → φ2]} ⊎ list_to_set_disj ((□θ) :: rm (□ θ) Δ))). { + assert(Heq'' : Γ0 ≡ (Γ ∖ {[□ φ1 → φ2]} ⊎ list_to_set_disj ((□θ) :: rm (□ θ) Δ'))). { rewrite Heq'. simpl. rewrite union_difference_L by trivial. rewrite (difference_singleton _ _ Hθ). rewrite list_to_set_disj_rm. ms. @@ -849,7 +851,7 @@ end; simpl. -- order_tac. rewrite elements_open_boxes. order_tac. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. - -- assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ ∖ {[□ φ1 → φ2]}) ⊎ ⊗ (list_to_set_disj Δ))). { + -- assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ ∖ {[□ φ1 → φ2]}) ⊎ ⊗ (list_to_set_disj Δ'))). { rewrite Heq'. rewrite union_difference_L by trivial. rewrite <- Hequiv, open_boxes_disj_union. rewrite open_boxes_remove by trivial. ms. } @@ -865,7 +867,7 @@ end; simpl. apply env_order_disj_union_compat_strong_left; order_tac. ++ intros φ0 Hin1 HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. - ++ assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ) ⊎ list_to_set_disj (map open_box (rm (□ φ1 → φ2) Δ)))). { + ++ assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ) ⊎ list_to_set_disj (map open_box (rm (□ φ1 → φ2) Δ')))). { rewrite Heq'. repeat rewrite open_boxes_remove by ms. simpl. rewrite <- list_to_set_disj_open_boxes, <- list_to_set_disj_rm, open_boxes_disj_union. trivial. @@ -873,7 +875,7 @@ end; simpl. } peapply Hp1. -- apply Hind. order_tac. occ. simpl. peapply' Hp2. trivial. - + assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ) ⊎ list_to_set_disj (map open_box (rm (□ φ1 → φ2) Δ)))). { + + assert(Heq'' : (⊗ Γ0) ≡ ((⊗Γ) ⊎ list_to_set_disj (map open_box (rm (□ φ1 → φ2) Δ')))). { rewrite Heq'. repeat rewrite open_boxes_remove by ms. simpl. rewrite <- list_to_set_disj_open_boxes, <- list_to_set_disj_rm, open_boxes_disj_union. trivial. @@ -909,7 +911,7 @@ end; simpl. ++ peapply Hp1. -- clear Heq''. apply Hind; [order_tac|occ|peapply' Hp2]. - split. - + intro Hocc. destruct (open_boxes_case (list_to_set_disj Δ)) as [[θ Hθ] | Hequiv]. + + intro Hocc. destruct (open_boxes_case (list_to_set_disj Δ')) as [[θ Hθ] | Hequiv]. * Etac. foldEA. apply BoxR. box_tac. exch 0. apply Hind. -- order_tac. rewrite elements_open_boxes. @@ -941,18 +943,19 @@ end; simpl. -- trivial. + Atac'. foldEA. apply BoxR. box_tac. apply weakening, weakening, make_impl_sound_R, ImpR. apply Hind. - * order_tac. - rewrite elements_open_boxes. + * order_tac. rewrite elements_open_boxes, simp_env_idempotent. repeat rewrite Permutation_middle. order_tac. * intros φ0 Hφ0 HF. destruct (occurs_in_open_boxes _ _ _ HF Hφ0) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. * (erewrite proper_Provable; [| |reflexivity]); [eapply Hp|]. - rewrite open_boxes_disj_union. - rewrite <- list_to_set_disj_env_add. - rewrite <- list_to_set_disj_open_boxes. ms. + subst Δ'. rewrite simp_env_idempotent. + rewrite open_boxes_disj_union. + rewrite <- list_to_set_disj_env_add. + rewrite <- list_to_set_disj_open_boxes. ms. Qed. + End PropQuantCorrect. End Correctness. @@ -964,10 +967,9 @@ End UniformInterpolation. Open Scope type_scope. - Lemma E_of_empty p : E p [] = (Implies Bot Bot). Proof. - rewrite E_eq. rewrite in_map_empty. now unfold conjunction, nodup, foldl. + rewrite E_eq, simp_env_nil. simpl. rewrite in_map_empty. now unfold conjunction, nodup, foldl. Qed. Definition vars_incl φ l := forall x, occurs_in x φ -> In x l. diff --git a/theories/iSL/Simp.v b/theories/iSL/Simp.v index 823a65b..27f6434 100644 --- a/theories/iSL/Simp.v +++ b/theories/iSL/Simp.v @@ -529,7 +529,7 @@ Qed. Require Import ISL.InvPropQuantifiers. -Definition E_simplified (p: variable) (ψ: form) := simp_form (E p ([ψ], ⊥)). +Definition E_simplified (p: variable) (ψ: form) := simp_form (E p [ψ]). Definition A_simplified (p: variable) (ψ: form) := simp_form (Af p (ψ)). (* diff --git a/theories/iSL/Simp_env.v b/theories/iSL/Simp_env.v new file mode 100644 index 0000000..b69628c --- /dev/null +++ b/theories/iSL/Simp_env.v @@ -0,0 +1,243 @@ +Require Import Coq.Program.Equality. +Require Import ISL.Sequents ISL.SequentProps. +Require Import ISL.Order ISL.DecisionProcedure. +Require Import Coq.Classes.RelationClasses. + +Definition applicable_AndL (Γ : list form): {ψ1 & {ψ2 | (And ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False). +Proof. + pose (fA := (fun θ => match θ with |And _ _ => true | _ => false end)). + destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fA. destruct θ. 3: { eexists. eexists. apply elem_of_list_In. eauto. } + all: auto with *. + - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. +Defined. + +Definition applicable_ImpLVar (Γ : list form): + {q & {ψ | Var q ∈ Γ /\ (Implies (Var q) ψ) ∈ Γ}} + + (∀ q ψ, Var q ∈ Γ -> (Implies (Var q) ψ) ∈ Γ -> False). +Proof. + pose (fIp :=λ p θ, match θ with | Implies (Var q) _ => if decide (p = q) then true else false | _ => false end). + pose (fp:= (fun θ => match θ with |Var p => if (exists_dec (fIp p) Γ) then true else false | _ => false end)). + destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fp. destruct θ. 2-6: auto with *. + case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. + unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. + destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. + subst. apply elem_of_list_In in Hinψ, Hin. + do 2 eexists. split; eauto. + - right. intros q ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. + simpl in Hp. case exists_dec as [|Hf'] in Hp. auto with *. + apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. +Defined. + +Definition applicable_ImpLAnd (Γ : list form): + {φ1 & {φ2 & {φ3 | (Implies (And φ1 φ2) φ3) ∈ Γ}}} + + (∀ φ1 φ2 φ3, (Implies (And φ1 φ2) φ3) ∈ Γ -> False). +Proof. + pose (fII := (fun θ => match θ with |Implies (And _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-2,4-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +Defined. + +Definition applicable_ImpLOr (Γ : list form): + {φ1 & {φ2 & {φ3 | (Implies (Or φ1 φ2) φ3) ∈ Γ}}} + + (∀ φ1 φ2 φ3, (Implies (Or φ1 φ2) φ3) ∈ Γ -> False). +Proof. + pose (fII := (fun θ => match θ with |Implies (Or _ _) _ => true | _ => false end)). + destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. + - left. subst fII. destruct θ. 1-4, 6: auto with *. + destruct θ1. 1-3, 5-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. + - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. +Defined. + +Definition sumor_bind0 {A} {B : Prop} {C}: A + B -> (A -> C) -> C -> C := +λ qH f c, +match qH with +| inl a => f a +| inr _ => c +end. + +Definition sumor_bind1 {A A'} {B : Prop} {C}: {q : A | A' q} + B -> (forall x (_ : A' x), C) -> C -> C := +λ qH f c, +match qH with +| inl (exist _ q Hq) => f q Hq +| inr _ => c +end. + +Definition sumor_bind2 {A A' A''} {B : Prop} {C}: {q : A & {r : A' | A'' q r}} + B -> (forall x y (_ : A'' x y), C) -> C -> C := +λ qH f c, +match qH with +| inl (existT q (exist _ r Hq)) => f q r Hq +| inr _ => c +end. + +Definition sumor_bind3 {A A' A'' A'''} {B : Prop} {C}: + {q : A & {r : A' & { s : A'' | A''' q r s}}} + B -> (forall x y z (_ : A''' x y z), C) -> C -> C := +λ qH f c, +match qH with +| inl (existT q (existT r (exist _ s Hq))) => f q r s Hq +| inr _ => c +end. + +Notation "cond '?' A ':0' B" := (sumor_bind0 cond A B) (at level 150, right associativity). +Notation "cond '?' A ':1' B" := (sumor_bind1 cond A B) (at level 150, right associativity). +Notation "cond '?' A ':2' B" := (sumor_bind2 cond A B) (at level 150, right associativity). +Notation "cond '?' A ':3' B" := (sumor_bind3 cond A B) (at level 150, right associativity). + +Local Notation "Δ '•' φ" := (cons φ Δ). + +Infix "⊢?" := Provable_dec (at level 80). + +(* Probably very costly *) +Definition applicable_strong_weakening (Γ : list form): + {φ : form | φ ∈ Γ /\ exists (_ : list_to_set_disj (rm φ Γ) ⊢ φ), True} + + (∀ φ, φ ∈ Γ -> forall (_ : list_to_set_disj (rm φ Γ) ⊢ φ), False). +Proof. +destruct (exists_dec (λ φ, if Provable_dec (rm φ Γ) φ then true else false) Γ) as [[φ [Hin Hφ]]| Hf]. +- left. exists φ; split. + + now apply elem_of_list_In. + + destruct ((rm φ Γ) ⊢? φ). trivial. contradict Hφ. +- right. intros φ Hin Hφ. apply (Hf φ). now apply elem_of_list_In. + destruct ((rm φ Γ) ⊢? φ). auto with *. tauto. +Defined. + +Local Obligation Tactic := (simpl; intuition order_tac). +Program Fixpoint simp_env (Δ : list form) {wf env_order Δ} : list form := + (* invertible left rules *) + if Δ ⊢? ⊥ then [⊥] else + applicable_AndL Δ ? λ δ₁ δ₂ _, simp_env ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂) :2 + applicable_ImpLVar Δ ? λ q ψ _, simp_env ((rm (Var q → ψ) Δ • ψ)) :2 + applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, simp_env ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)))) :3 + applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, simp_env (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃)) :3 + (* remove redundant assumptions *) + applicable_strong_weakening Δ ? λ φ _, simp_env (rm φ Δ) :1 + Δ +. +Next Obligation. apply Wf.measure_wf, wf_env_order. Qed. + +Lemma simp_env_eq (Δ : list form): simp_env Δ = + (* invertible left rules *) + if Δ ⊢? ⊥ then [⊥] else + applicable_AndL Δ ? λ δ₁ δ₂ _, simp_env ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂) :2 + applicable_ImpLVar Δ ? λ q ψ _, simp_env ((rm (Var q → ψ) Δ • ψ)) :2 + applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, simp_env ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)))) :3 + applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, simp_env (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃)) :3 + (* remove redundant assumptions *) + applicable_strong_weakening Δ ? λ φ _, simp_env (rm φ Δ) :1 + Δ. +Proof. +simpl. unfold simp_env. simpl. +repeat rewrite Wf.WfExtensionality.fix_sub_eq_ext; simpl. split; trivial. +Qed. + +Definition simp_form φ:= ⋀ (simp_env [φ]). + +Definition pointed_env_order_refl pe1 pe2 := env_order_refl (pe1.2 :: pe1.2 :: pe1.1) (pe2.2 :: pe2.2 :: pe2.1). + +(* TODO: move *) + +Lemma empty_entails_not_bot : (∅ ⊢ ⊥) -> False. +Proof. +intro Hf. dependent destruction Hf; simpl in *; +match goal with x : _ ⊎ {[+?φ+]} = _ |- _ => +exfalso; eapply (gmultiset_elem_of_empty φ); setoid_rewrite <- x; ms end. +Qed. + +Lemma simp_env_order Δ : env_order_refl (simp_env Δ) Δ. +Proof. +revert Δ. apply (well_founded_induction_type wf_env_order). intros Δ Hind. +rewrite simp_env_eq. +case (Δ ⊢? ⊥). +{ intros [Hf _]. destruct Δ as [|φ Δ]. + + now apply empty_entails_not_bot in Hf. + + rewrite env_weight_add. + unfold env_order_refl. destruct φ. + * admit. + * destruct Δ as [|φ Δ]. + -- now right. + -- left. unfold env_order, ltof, env_weight. simpl. pose (Order.pow5_gt_0 (weight φ)). lia. + * left. unfold env_order, ltof, env_weight. simpl. + replace 5 with (5 ^ 1) at 1 by trivial. lia. + apply Nat.lt_le_trans with (5 ^ weight φ); [|lia]. + apply Nat.pow_lt_mono_r. lia. destruct φ. simpl. lia. + nia. + pose (Order.pow5_gt_0 (weight φ)). lia. + intro. simpl. auto with proof. + +Admitted. + +Global Hint Resolve simp_env_order : order. + +Lemma simp_env_pointed_env_order_L pe Δ φ: (pe ≺· (simp_env Δ, φ)) -> pe ≺· (Δ, φ). +Proof. +intro Hl. eapply env_order_lt_le_trans; eauto. simpl. auto with order. +Qed. + +Lemma simp_env_env_order_L Δ Δ0: (Δ0 ≺ simp_env Δ) -> Δ0 ≺ Δ. +Proof. +intro Hl. eapply env_order_lt_le_trans; eauto. auto with order. +Qed. + +Global Hint Resolve simp_env_pointed_env_order_L : order. +Global Hint Resolve simp_env_env_order_L : order. + +Section Equivalence. + +Definition equiv_form φ ψ : Type := (φ ≼ ψ) * (ψ ≼ φ). + +Definition equiv_env Δ Γ: Set := (list_to_set_disj Δ ⊢ ⋀ Γ) * (list_to_set_disj Γ ⊢ ⋀ Δ). + +Lemma symmetric_equiv_env Δ Γ: equiv_env Δ Γ -> equiv_env Γ Δ . +Proof. intros [H1 H2]. split; trivial. Qed. + +Lemma equiv_env_L0 Δ Δ' φ: (equiv_env Δ Δ') -> + list_to_set_disj Δ ⊢ φ -> list_to_set_disj Δ' ⊢ φ. +Proof. +intros [H1 H2]. +Admitted. + + +Lemma equiv_env_L1 Γ Δ Δ' φ: (equiv_env Δ Δ') -> + Γ ⊎ list_to_set_disj Δ ⊢ φ -> Γ ⊎ list_to_set_disj Δ' ⊢ φ. +Proof. +intros [H1 H2]. +Admitted. + +Lemma equiv_env_L2 Γ Δ Δ' φ: (equiv_env Δ Δ') -> + list_to_set_disj Δ ⊎ Γ ⊢ φ -> list_to_set_disj Δ' ⊎ Γ ⊢ φ. +Proof. +Admitted. + +Lemma simp_env_equiv_env Δ: equiv_env (simp_env Δ) Δ. +Proof. +Admitted. + +Lemma simp_env_L1 Γ Δ φ: + Γ ⊎ list_to_set_disj (simp_env Δ) ⊢ φ -> Γ ⊎ list_to_set_disj Δ ⊢ φ. +Proof. +apply equiv_env_L1, simp_env_equiv_env. +Qed. + +End Equivalence. + +Global Hint Resolve simp_env_L1 : proof. + +Infix "≡f" := equiv_form (at level 120). + +Global Infix "≡e" := equiv_env (at level 120). + +Section Variables. +Lemma equiv_env_vars Δ x: + (∃ θ : form, ((θ ∈ simp_env Δ) /\ occurs_in x θ)) -> + ∃ θ : form, ((θ ∈ Δ) /\ occurs_in x θ). +Proof. +intros [θ [Hin Hocc]]. +Admitted. + +End Variables. + +Lemma simp_env_idempotent Δ: simp_env (simp_env Δ) = simp_env Δ. +Proof. +Admitted. \ No newline at end of file From b218377ecf58fa5a01d554c03452e92a5764f21d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Tue, 10 Sep 2024 22:04:47 +0200 Subject: [PATCH 10/12] fix definition for env_order_ refl --- theories/iSL/Order.v | 44 ++++++++++++++------------------------------ 1 file changed, 14 insertions(+), 30 deletions(-) diff --git a/theories/iSL/Order.v b/theories/iSL/Order.v index e780708..acbf8cb 100644 --- a/theories/iSL/Order.v +++ b/theories/iSL/Order.v @@ -32,7 +32,7 @@ intro Hw. unfold env_order, ltof. do 2 rewrite env_weight_singleton. apply Nat.pow_lt_mono_r. lia. trivial. Qed. -Definition env_order_refl Δ Δ' := (Δ ≺ Δ') \/ Δ ≡ₚ Δ'. +Definition env_order_refl Δ Δ' := (env_weight Δ) ≤(env_weight Δ'). Global Notation "Δ ≼ Δ'" := (env_order_refl Δ Δ') (at level 150). @@ -44,12 +44,7 @@ Qed. Global Instance Proper_env_order_refl_env_weight: Proper ((env_order_refl) ==> le) env_weight. -Proof. -intros Γ Δ. -unfold env_order. intros [Hle | Heq]. -- auto with *. -- now rewrite Heq. -Qed. +Proof. intros Γ Δ Hle. auto with *. Qed. Global Hint Resolve Proper_env_order_refl_env_weight : order. @@ -164,8 +159,7 @@ Lemma env_order_refl_disj_union_compat Δ Δ' Δ'' Δ''': (env_order_refl Δ Δ'') -> (env_order_refl Δ' Δ''') -> env_order_refl (Δ ++ Δ') (Δ'' ++ Δ'''). Proof. unfold env_order_refl, env_order, ltof. repeat rewrite env_weight_disj_union. -intros [Hlt1 | Heq1] [Hlt2 | Heq2]; try rewrite Heq1; try rewrite Heq2; try lia. -right. trivial. +intros Hle1 Hle2; try rewrite Heq1; try rewrite Heq2; try lia. Qed. Global Hint Resolve env_order_refl_disj_union_compat : order. @@ -174,20 +168,13 @@ Hint Unfold env_order_refl : order. Lemma env_order_disj_union_compat_strong_right Δ Δ' Δ'' Δ''': (Δ ≺ Δ'') -> (Δ' ≼ Δ''') -> Δ ++ Δ' ≺ Δ'' ++ Δ'''. Proof. -intros H1 H2. -destruct H2 as [Hlt|Heq]. -- unfold env_order, ltof in *. do 2 rewrite env_weight_disj_union. lia. -- rewrite Heq. now apply env_order_disj_union_compat_left. +intros Hlt Hle. unfold env_order_refl, env_order, ltof in *. do 2 rewrite env_weight_disj_union. lia. Qed. Lemma env_order_disj_union_compat_strong_left Δ Δ' Δ'' Δ''': (Δ ≼ Δ'') -> (Δ' ≺ Δ''') -> Δ ++ Δ' ≺ Δ'' ++ Δ'''. Proof. -intros H1 H2. -destruct H1 as [Hlt|Heq]. -- unfold env_order, ltof in *. do 2 rewrite env_weight_disj_union. lia. -- rewrite Heq. now apply env_order_disj_union_compat_right. -Qed. +intros Hlt Hle. unfold env_order_refl, env_order, ltof in *. do 2 rewrite env_weight_disj_union. lia. Qed. Global Hint Resolve env_order_disj_union_compat_strong_left : order. Global Hint Rewrite elements_open_boxes : order. @@ -197,12 +184,9 @@ Proof. destruct φ; simpl; lia. Qed. Lemma open_boxes_env_order Δ : (map open_box Δ) ≼ Δ. Proof. -induction Δ as [|φ Δ]. -- right. trivial. -- destruct IHΔ as [Hlt | Heq]; simpl. - + left. apply env_order_compat'; trivial. apply weight_open_box. - + rewrite Heq. auto with order. destruct φ; simpl; try auto with order. - left. auto with order. +unfold env_order_refl. +induction Δ as [|φ Δ]; trivial. +simpl. do 2 rewrite env_weight_add. destruct φ; simpl; lia. Qed. Global Hint Resolve open_boxes_env_order : order. @@ -281,7 +265,7 @@ Proof. etransitivity; [|apply env_order_0]. assumption. Qed. Lemma env_order_eq_add Δ Δ' φ: (Δ ≼ Δ') -> (Δ • φ) ≼ (Δ' • φ). -Proof. intros[Heq|Hlt]. left. now apply env_order_add_compat. right. ms. Qed. +Proof. unfold env_order_refl. do 2 rewrite env_weight_add. lia. Qed. Global Hint Resolve env_order_0 : order. @@ -312,9 +296,9 @@ Global Hint Rewrite open_boxes_remove : order. Lemma remove_env_order Δ φ: rm φ Δ ≼ Δ. Proof. -induction Δ as [|ψ Δ]. -- simpl. right. auto. -- simpl. destruct form_eq_dec; auto with order. +unfold env_order_refl. +induction Δ as [|ψ Δ]. trivial. +simpl. destruct form_eq_dec; repeat rewrite env_weight_add; lia. Qed. Global Hint Resolve remove_env_order : order. @@ -333,10 +317,10 @@ Qed. Global Hint Resolve remove_In_env_order_refl : order. Lemma env_order_lt_le_trans Γ Γ' Γ'' : (Γ ≺ Γ') -> (Γ' ≼ Γ'') -> Γ ≺ Γ''. -Proof. intros Hlt [Hlt' | Heq]. transitivity Γ'; auto with order. now rewrite <- Heq. Qed. +Proof. unfold env_order, env_order_refl, ltof. lia. Qed. Lemma env_order_le_lt_trans Γ Γ' Γ'' : (Γ ≼ Γ') -> (Γ' ≺ Γ'') -> Γ ≺ Γ''. -Proof. intros [Hlt' | Heq] Hlt. transitivity Γ'; auto with order. now rewrite Heq. Qed. +Proof. unfold env_order, env_order_refl, ltof. lia. Qed. Lemma remove_In_env_order Δ φ: In φ Δ -> rm φ Δ ≺ Δ. Proof. From 9cd54814adfc143656225b7481250572b5a4e57e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Wed, 11 Sep 2024 13:30:39 +0200 Subject: [PATCH 11/12] =?UTF-8?q?Done=20simplifying=20=CE=94=20in=20E(?= =?UTF-8?q?=CE=94)=20and=20A(=CE=94,=20=CF=86)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- theories/extraction/UIML_extraction.v | 2 +- theories/iSL/InvPropQuantifiers.v | 334 -------------------------- theories/iSL/Order.v | 22 +- theories/iSL/PropQuantifiers.v | 17 +- theories/iSL/Simp.v | 2 +- theories/iSL/Simp_env.v | 240 +++++++++++++++--- 6 files changed, 237 insertions(+), 380 deletions(-) delete mode 100644 theories/iSL/InvPropQuantifiers.v diff --git a/theories/extraction/UIML_extraction.v b/theories/extraction/UIML_extraction.v index ba82573..d44aae4 100644 --- a/theories/extraction/UIML_extraction.v +++ b/theories/extraction/UIML_extraction.v @@ -5,7 +5,7 @@ Require Import ExtrOcamlBasic ExtrOcamlString. Require Import K.Interpolation.UIK_braga. Require Import KS_export. -Require Import ISL.InvPropQuantifiers ISL.DecisionProcedure ISL.Simp. +Require Import ISL.PropQuantifiers ISL.DecisionProcedure ISL.Simp. Fixpoint MPropF_of_form (f : Formulas.form) : MPropF := diff --git a/theories/iSL/InvPropQuantifiers.v b/theories/iSL/InvPropQuantifiers.v deleted file mode 100644 index 983ad70..0000000 --- a/theories/iSL/InvPropQuantifiers.v +++ /dev/null @@ -1,334 +0,0 @@ -Require Import Coq.Program.Equality. (* for dependent induction *) -Require Import ISL.PropQuantifiers ISL.Sequents ISL.SequentProps. -Require Import ISL.Order ISL.DecisionProcedure. -Require Import Coq.Classes.RelationClasses. -Section UniformInterpolation. - -Definition applicable_AndR φ: {φ1 & {φ2 | φ = (And φ1 φ2)}} + (∀ φ1 φ2, φ ≠ (And φ1 φ2)). -Proof. (destruct φ; eauto). Defined. -(* -Definition applicable_other_var (Γ : list form): {q | p ≠ q /\ Var q ∈ Γ} + (∀ q, p ≠ q -> Var q ∈ Γ -> False). -Proof. - pose (fA := (fun θ => match θ with |Var q => if decide (p = q) then false else true | _ => false end)). - destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fA. destruct θ. - 1: { case decide in Hθ. auto with *. eexists; split; eauto. apply elem_of_list_In. assumption. } - all: auto with *. - - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. - case decide in Hψ; auto with *. -Defined. -*) - -Definition applicable_AndL (Γ : list form): {ψ1 & {ψ2 | (And ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False). -Proof. - pose (fA := (fun θ => match θ with |And _ _ => true | _ => false end)). - destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fA. destruct θ. 3: { eexists. eexists. apply elem_of_list_In. eauto. } - all: auto with *. - - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. -Defined. - -Definition applicable_ImpR φ: {φ1 & {φ2 | φ = (Implies φ1 φ2)}} + (∀ φ1 φ2, φ ≠ (Implies φ1 φ2)). -Proof. destruct φ; eauto. Defined. - - -Definition applicable_OrL (Γ : list form): {ψ1 & {ψ2 | (Or ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (Or ψ1 ψ2) ∈ Γ -> False). -Proof. - pose (fA := (fun θ => match θ with |Or _ _ => true | _ => false end)). - destruct (exists_dec fA Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fA. destruct θ. 4: { eexists. eexists. apply elem_of_list_In. eauto. } - all: auto with *. - - right. intros ψ1 ψ2 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fA. simpl in Hψ. tauto. -Defined. - -(* -Definition applicable_ImpL_other_Var (Γ : list form): - {q & {ψ | q ≠ p /\ (Implies (Var q) ψ) ∈ Γ}} + - (∀ q ψ, q ≠ p -> (Implies (Var q) ψ) ∈ Γ -> False). -Proof. - pose (fIp :=λ θ, match θ with | Implies (Var q) _ => if decide (p = q) then false else true | _ => false end). - destruct (exists_dec fIp Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fIp. destruct θ; auto with *. - destruct θ1; auto with *. - case decide in Hθ; auto with *. - apply elem_of_list_In in Hin. - do 2 eexists; split; eauto. - - right. intros q ψ Hp Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fIp. - simpl in Hψ. rewrite decide_False in Hψ; trivial; auto with *. -Defined. -*) - -Definition applicable_ImpLVar (Γ : list form): - {q & {ψ | Var q ∈ Γ /\ (Implies (Var q) ψ) ∈ Γ}} + - (∀ q ψ, Var q ∈ Γ -> (Implies (Var q) ψ) ∈ Γ -> False). -Proof. - pose (fIp :=λ p θ, match θ with | Implies (Var q) _ => if decide (p = q) then true else false | _ => false end). - pose (fp:= (fun θ => match θ with |Var p => if (exists_dec (fIp p) Γ) then true else false | _ => false end)). - destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fp. destruct θ. 2-6: auto with *. - case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. - unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. - destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. - subst. apply elem_of_list_In in Hinψ, Hin. - do 2 eexists. split; eauto. - - right. intros q ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. - simpl in Hp. case exists_dec as [|Hf'] in Hp. auto with *. - apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. -Defined. - -Definition applicable_ImpLVar_spec (Γ : list form) q: - {ψ | Var q ∈ Γ /\ (Implies (Var q) ψ) ∈ Γ} + - (∀ ψ, Var q ∈ Γ -> (Implies (Var q) ψ) ∈ Γ -> False). -Proof. - pose (fIp :=λ θ, match θ with | Implies (Var p) _ => if decide (p = q) then true else false | _ => false end). - pose (fp:= (fun θ => match θ with |Var p => if decide (p = q) then if (exists_dec fIp Γ) then true else false else false| _ => false end)). - destruct (exists_dec fp Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fp. destruct θ. 2-6: auto with *. - case decide in Hθ; [|auto with *]. subst. - case exists_dec as [(ψ &Hinψ & Hψ)|] in Hθ; [|auto with *]. - unfold fIp in Hψ. destruct ψ. 1-4, 6: auto with *. - destruct ψ1. 2-6: auto with *. case decide in Hψ; [|auto with *]. - subst. apply elem_of_list_In in Hinψ, Hin. - eexists. split; eauto. - - right. intros ψ Hp Hψ. rewrite elem_of_list_In in Hp, Hψ. apply Hf in Hp. subst fp fIp. - simpl in Hp. rewrite decide_True in Hp by trivial. - case exists_dec as [|Hf'] in Hp. auto with *. - apply (Hf' _ Hψ). rewrite decide_True; trivial. auto with *. -Defined. - -Definition applicable_ImpLAnd (Γ : list form): - {φ1 & {φ2 & {φ3 | (Implies (And φ1 φ2) φ3) ∈ Γ}}} + - (∀ φ1 φ2 φ3, (Implies (And φ1 φ2) φ3) ∈ Γ -> False). -Proof. - pose (fII := (fun θ => match θ with |Implies (And _ _) _ => true | _ => false end)). - destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fII. destruct θ. 1-4, 6: auto with *. - destruct θ1. 1-2,4-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. - - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. -Defined. - -Definition applicable_ImpLOr (Γ : list form): - {φ1 & {φ2 & {φ3 | (Implies (Or φ1 φ2) φ3) ∈ Γ}}} + - (∀ φ1 φ2 φ3, (Implies (Or φ1 φ2) φ3) ∈ Γ -> False). -Proof. - pose (fII := (fun θ => match θ with |Implies (Or _ _) _ => true | _ => false end)). - destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fII. destruct θ. 1-4, 6: auto with *. - destruct θ1. 1-3, 5-6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. - - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. -Defined. - -Definition applicable_OrR φ: - {φ1 & {φ2 | φ = (Or φ1 φ2)}} + (∀ φ1 φ2, φ = (Or φ1 φ2) -> False). -Proof. destruct φ; eauto with *. Defined. - -Definition applicable_BoxR φ: {φ' | φ = (□ φ')} + (∀ φ', φ = (□ φ') -> False). -Proof. destruct φ; eauto with *. Defined. - -(* -Definition applicable_ImpLImp (Γ : list form): - {φ1 & {φ2 & {φ3 | ((φ1 → φ2) → φ3) ∈ Γ}}} +(forall φ1 φ2 φ3, ((φ1 → φ2) → φ3) ∈ Γ -> False). -Proof. - pose (fII := (fun θ => match θ with |Implies (Implies _ _) _ => true | _ => false end)). - destruct (exists_dec fII Γ) as [(θ & Hin & Hθ) | Hf]. - - left. subst fII. destruct θ. 1-4, 6: auto with *. - destruct θ1. 1-4, 6: auto with *. do 3 eexists; apply elem_of_list_In; eauto. - - right. intros ψ1 ψ2 ψ3 Hψ. rewrite elem_of_list_In in Hψ. apply Hf in Hψ. subst fII. simpl in Hψ. tauto. -Defined. -*) - -(* can this be replaced with generalised axiom?*) -Definition applicable_Atom q (Δ : list form) ϕ : - (ϕ = Var q /\ Var q ∈ Δ) + (ϕ ≠ Var q \/ (Var q ∈ Δ -> False)). -Proof. -case (decide (ϕ = Var q)). -case (decide (Var q ∈ Δ)). -all: auto with *. -Defined. - -Definition sumor_bind0 {A} {B : Prop} {C}: A + B -> (A -> C) -> C -> C := -λ qH f c, -match qH with -| inl a => f a -| inr _ => c -end. - -Definition sumor_bind1 {A A'} {B : Prop} {C}: {q : A | A' q} + B -> (forall x (_ : A' x), C) -> C -> C := -λ qH f c, -match qH with -| inl (exist _ q Hq) => f q Hq -| inr _ => c -end. - -Definition sumor_bind2 {A A' A''} {B : Prop} {C}: {q : A & {r : A' | A'' q r}} + B -> (forall x y (_ : A'' x y), C) -> C -> C := -λ qH f c, -match qH with -| inl (existT q (exist _ r Hq)) => f q r Hq -| inr _ => c -end. - -Definition sumor_bind3 {A A' A'' A'''} {B : Prop} {C}: - {q : A & {r : A' & { s : A'' | A''' q r s}}} + B -> (forall x y z (_ : A''' x y z), C) -> C -> C := -λ qH f c, -match qH with -| inl (existT q (existT r (exist _ s Hq))) => f q r s Hq -| inr _ => c -end. - -Notation "cond '?' A ':0' B" := (sumor_bind0 cond A B) (at level 150, right associativity). -Notation "cond '?' A ':1' B" := (sumor_bind1 cond A B) (at level 150, right associativity). -Notation "cond '?' A ':2' B" := (sumor_bind2 cond A B) (at level 150, right associativity). -Notation "cond '?' A ':3' B" := (sumor_bind3 cond A B) (at level 150, right associativity). - -(* TODO: plan: - - reuse propquantifiers (esp, don't redefine erule); - - define propquant' as the same fixpoint as propquant, but with a simplification step - - the simplification shold look like the monadic simplifications done here, but more general - - to ensure that' it's still correct, prove that simplified sequents are "equivalent" to the original ones - - show (by induction on the definition of EA and EA') that EA equiv EA' - *) - -Local Notation "Δ '•' φ" := (cons φ Δ). - -Infix "⊢?" := Provable_dec (at level 80). - -(* Probably very costly *) -Definition applicable_strong_weakening (Γ : list form): - {φ : form | φ ∈ Γ /\ exists (_ : list_to_set_disj (rm φ Γ) ⊢ φ), True} - + (∀ φ, φ ∈ Γ -> forall (_ : list_to_set_disj (rm φ Γ) ⊢ φ), False). -Proof. -destruct (exists_dec (λ φ, if Provable_dec (rm φ Γ) φ then true else false) Γ) as [[φ [Hin Hφ]]| Hf]. -- left. exists φ; split. - + now apply elem_of_list_In. - + destruct ((rm φ Γ) ⊢? φ). trivial. contradict Hφ. -- right. intros φ Hin Hφ. apply (Hf φ). now apply elem_of_list_In. - destruct ((rm φ Γ) ⊢? φ). auto with *. tauto. -Defined. - -Obligation Tactic := (simpl; intuition order_tac). -Program Fixpoint simp_env (Δ : list form) {wf env_order Δ} : list form := - (* invertible left rules *) - applicable_AndL Δ ? λ δ₁ δ₂ _, simp_env ((rm (δ₁ ∧ δ₂) Δ•δ₁)•δ₂) :2 - applicable_ImpLVar Δ ? λ q ψ _, simp_env ((rm (Var q → ψ) Δ • ψ)) :2 - applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, simp_env ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)))) :3 - applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, simp_env (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃)) :3 - (* remove redundant assumptions *) - applicable_strong_weakening Δ ? λ φ _, simp_env (rm φ Δ) :1 - Δ -. -Next Obligation. apply Wf.measure_wf, wf_env_order. Qed. - -Definition simp_form φ:= ⋀ (simp_env [φ]). - -Definition pointed_env_order_refl pe1 pe2 := env_order_refl (pe1.2 :: pe1.2 :: pe1.1) (pe2.2 :: pe2.2 :: pe2.1). - -Lemma simp_env_order pe : env_order_refl (simp_env pe) pe. -Proof. -Admitted. - -Variable p : variable. - -Obligation Tactic := solve[simpl; intros; simpl in *; intuition; order_tac] || (try (apply Wf.measure_wf, wf_pointed_order)). -Program Fixpoint EA (b : bool) (pe : pointed_env) {wf pointed_env_order pe} : form := - let Δ := simp_env pe.1 in (* because pattern matchin in lets is not broken *) - let ϕ := pe.2 in - let E pe H := EA true pe in - let A pe H := EA false pe in - (* E *) - if b then - ⋀ (in_map Δ (@e_rule p Δ ϕ E A)) - (* A *) - else - ⋁ (in_map Δ (@a_rule_env p Δ ϕ E A)) ⊻ @a_rule_form p Δ ϕ E A. -Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. -order_tac. do 2 apply env_order_eq_add. apply simp_env_order. Qed. -Next Obligation. intros. subst Δ ϕ. eapply env_order_lt_le_trans; [exact H|]. -order_tac. do 2 apply env_order_eq_add. apply simp_env_order. Qed. - -Definition E Δ := EA true (Δ, ⊥). -Definition A := EA false. - -Lemma EA_eq Δ ϕ: let Δ' := simp_env Δ in - (E Δ = ⋀ (in_map Δ' (@e_rule p Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) /\ - (A (Δ, ϕ) = (⋁ (in_map Δ' (@a_rule_env p Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)))) ⊻ - @a_rule_form p Δ' ϕ (λ pe _, E pe.1) (λ pe _, A pe)). -Proof. -simpl. unfold E, A, EA. simpl. -unfold EA_func at 1. -rewrite -> Wf.Fix_eq; simpl. -- split; trivial. - unfold EA_func at 1. rewrite -> Wf.Fix_eq; simpl; trivial. - intros [[|] Δ'] f g Heq; simpl; f_equal. - + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. - + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; - now rewrite Heq. - + apply a_rule_form_cong; intros; now rewrite Heq. -- intros [[|] Δ'] f g Heq; simpl; f_equal. - + apply in_map_ext. intros. trivial. apply e_rule_cong; intros; now rewrite Heq. - + f_equal. apply in_map_ext. intros. apply a_rule_env_cong; intros; - now rewrite Heq. - + apply a_rule_form_cong; intros; now rewrite Heq. -Qed. - -Definition Ef (ψ : form) := E [ψ]. -Definition Af (ψ : form) := A ([], ψ). - -End UniformInterpolation. - -Section Equivalence. - -Definition equiv_form φ ψ : Type := (φ ≼ ψ) * (ψ ≼ φ). - -Definition equiv_env Δ Γ: Set := (list_to_set_disj Δ ⊢ ⋀ Γ) * (list_to_set_disj Γ ⊢ ⋀ Δ). - -Lemma symmetric_equiv_env Δ Γ: equiv_env Δ Γ -> equiv_env Γ Δ . -Proof. intros [H1 H2]. split; trivial. Qed. - -Lemma equiv_env_L1 Γ Δ Δ' φ: (equiv_env Δ Δ') -> - Γ ⊎ list_to_set_disj Δ ⊢ φ -> Γ ⊎ list_to_set_disj Δ' ⊢ φ. -Proof. -intros [H1 H2]. -Admitted. - -Lemma equiv_env_L2 Γ Δ Δ' φ: (equiv_env Δ Δ') -> - list_to_set_disj Δ ⊎ Γ ⊢ φ -> list_to_set_disj Δ' ⊎ Γ ⊢ φ. -Proof. -Admitted. - -End Equivalence. - -Infix "≡f" := equiv_form (at level 120). - -Global Infix "≡e" := equiv_env (at level 120). -(* -Section Soundness. - -Variable p : variable. - -Proposition entail_correct_equiv_env Δ Δ' ϕ : (Δ ≡e Δ') -> - (list_to_set_disj Δ ⊢ PropQuantifiers.E p Δ') * - (list_to_set_disj Δ • PropQuantifiers.A p (Δ', ϕ) ⊢ ϕ). -Proof. -intro Heq. split. -- peapply (equiv_env_L1 gmultiset_empty Δ' Δ); eauto. now apply symmetric_equiv_env. - peapply (entail_correct p Δ' ϕ).1 . -- peapply (equiv_env_L2 {[PropQuantifiers.A p (Δ', ϕ)]} Δ' Δ); eauto. now apply symmetric_equiv_env. - peapply (entail_correct p Δ' ϕ).2. -Qed. - -Lemma pq_correct_equiv_env Γ Δ ϕ Δ': - (∀ φ0, φ0 ∈ Γ -> ¬ occurs_in p φ0) -> - (Γ ⊎ list_to_set_disj Δ ⊢ ϕ) -> - (Δ ≡e Δ') -> - (¬ occurs_in p ϕ -> Γ• PropQuantifiers.E p Δ' ⊢ ϕ) * - (Γ• PropQuantifiers.E p Δ' ⊢ PropQuantifiers.A p (Δ', ϕ)). -Proof. -intros Hocc Hp Heq. split. -- intro Hocc'. apply pq_correct; trivial. eapply equiv_env_L1; eauto. -- apply pq_correct; trivial. eapply equiv_env_L1; eauto. -Qed. - - - -End Soundness. -*) \ No newline at end of file diff --git a/theories/iSL/Order.v b/theories/iSL/Order.v index acbf8cb..c52accc 100644 --- a/theories/iSL/Order.v +++ b/theories/iSL/Order.v @@ -36,6 +36,14 @@ Definition env_order_refl Δ Δ' := (env_weight Δ) ≤(env_weight Δ'). Global Notation "Δ ≼ Δ'" := (env_order_refl Δ Δ') (at level 150). +Lemma env_order_env_order_refl Δ Δ' : env_order Δ Δ' -> env_order_refl Δ Δ'. +Proof. unfold env_order, env_order_refl, ltof. lia. Qed. + +Global Hint Resolve env_order_env_order_refl: order. + +Lemma env_order_self Δ : Δ ≼ Δ. +Proof. unfold env_order_refl. trivial. Qed. + Global Instance Proper_env_weight: Proper ((≡ₚ) ==> (=)) env_weight. Proof. intros Γ Δ Heq. unfold env_weight. now rewrite Heq. @@ -396,12 +404,22 @@ Hint Extern 5 (?a ≺· ?b) => order_tac : proof. Lemma pointed_env_order_bot_R pe Δ φ: (pe ≺· (Δ, ⊥)) -> pe ≺· (Δ, φ). Proof. -Admitted. +intro Hlt. destruct pe as (Γ, ψ). unfold pointed_env_order. simpl. +eapply env_order_lt_le_trans. exact Hlt. simpl. +unfold env_order_refl. repeat rewrite env_weight_add. +assert(Hle: weight ⊥ ≤ weight φ) by (destruct φ; simpl; lia). +apply Nat.pow_le_mono_r with (a := 5) in Hle; lia. +Qed. Hint Resolve pointed_env_order_bot_R : order. Lemma pointed_env_order_bot_L pe Δ φ: ((Δ, φ) ≺· pe) -> (Δ, ⊥) ≺· pe. Proof. -Admitted. +intro Hlt. destruct pe as (Γ, ψ). unfold pointed_env_order. simpl. +eapply env_order_le_lt_trans; [|exact Hlt]. simpl. +unfold env_order_refl. repeat rewrite env_weight_add. +assert(Hle: weight ⊥ ≤ weight φ) by (destruct φ; simpl; lia). +apply Nat.pow_le_mono_r with (a := 5) in Hle; lia. +Qed. Hint Resolve pointed_env_order_bot_L : order. diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index b03fe6c..b86549d 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -74,6 +74,7 @@ match θ with | (□δ1 → δ2) => (□(E((□⁻¹ Δ') • δ2 • □δ1) _ ⇢ A((□⁻¹ Δ') • δ2 • □δ1, δ1) _)) ⇢ E(Δ' • δ2) _ end. + (** The implementation of the rules for defining A is separated into two pieces. Referring to Table 5 in Pitts, the definition a_rule_env handles A1-8 and A10, and the definition a_rule_form handles A9 and A11-13. *) @@ -448,7 +449,7 @@ split. { (* E *) apply conjunction_R1. intros φ Hin. apply in_in_map in Hin. destruct Hin as (ψ&Hin&Heq). subst φ. -apply (equiv_env_L0 Δ'); [subst Δ'; apply simp_env_equiv_env|]. +apply simp_env_equiv_env. rewrite <- HeqΔ' in *. clear HeqΔ' Δ. rename Δ' into Δ. assert(Hi : ψ ∈ list_to_set_disj Δ) by now apply elem_of_list_to_set_disj. destruct ψ; unfold e_rule; exhibit Hi 0; rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (list_to_set_disj_rm _ _)) _ _ eq_refl). @@ -479,10 +480,9 @@ destruct ψ; unfold e_rule; exhibit Hi 0; rewrite (proper_Provable _ _ (equiv_di ++ do 2 l_tac. apply HA. order_tac. -- exch 0. l_tac. apply weakening, HE. order_tac. - apply BoxR. apply weakening. box_tac. l_tac. apply HE. order_tac. - } (* A *) -apply (equiv_env_L2 _ Δ'); [subst Δ'; apply simp_env_equiv_env|]. +apply ImpR_rev, simp_env_equiv_env, ImpR. rewrite <- HeqΔ' in *. clear HeqΔ' Δ. rename Δ' into Δ. apply make_disj_sound_L, OrL. - apply disjunction_L. intros φ Hin. @@ -598,6 +598,7 @@ assert(Ho' : pointed_env_order_refl (elements Γ ++ simp_env Δ, ϕ) (elements assert(Hind' := λ Γ0 Δ0 ψ Ho, Hind (elements Γ0 ++ Δ0, ψ) (env_order_lt_le_trans _ _ _ Ho Ho') _ _ _ eq_refl). simpl in Hind'. clear Ho'. clear Hind. rename Hind' into Hind. rewrite <- E_simp_env, <- A_simp_env. +(simp_env_equiv_env Δ).1 Hp assert(Hp' := equiv_env_L1 Γ _ _ _ (symmetric_equiv_env _ _ (simp_env_equiv_env Δ)) Hp). remember (simp_env Δ) as Δ'. clear Hp. rename Hp' into Hp. @@ -792,7 +793,8 @@ end; simpl. exch 1; exch 0. apply Hind. -- order_tac. (* cannot handle it. manual proof *) rewrite elements_open_boxes. - do 2 apply env_order_cancel_right. apply env_order_4; simpl; try lia. left. + do 2 apply env_order_cancel_right. apply env_order_4; simpl; try lia. + apply env_order_env_order_refl. apply env_order_disj_union_compat_strong_left; order_tac. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. @@ -829,7 +831,8 @@ end; simpl. exch 1; exch 0. apply Hind. -- order_tac. rewrite elements_open_boxes. - do 2 apply env_order_cancel_right. apply env_order_4; simpl; try lia. left. + do 2 apply env_order_cancel_right. apply env_order_4; simpl; try lia. + apply env_order_env_order_refl. apply env_order_disj_union_compat_strong_left; order_tac. -- occ. intro HF. destruct (occurs_in_open_boxes _ _ _ HF Hin1) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. @@ -918,7 +921,7 @@ end; simpl. repeat rewrite Permutation_middle. apply env_order_disj_union_compat_strong_left. order_tac. apply env_order_2. simpl; lia. simpl; lia. - apply env_order_eq_add. left. order_tac. + apply env_order_eq_add. apply env_order_env_order_refl. order_tac. -- intros φ0 Hφ0 HF. apply gmultiset_elem_of_disj_union in Hφ0. destruct Hφ0 as [Hφ0|]; [|ms]. destruct (occurs_in_open_boxes _ _ _ HF Hφ0) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. @@ -934,7 +937,7 @@ end; simpl. rewrite elements_open_boxes. repeat rewrite Permutation_middle. apply env_order_disj_union_compat_strong_left. order_tac. - apply env_order_2. simpl; lia. simpl; lia. now right. + apply env_order_2. simpl; lia. simpl; lia. unfold env_order_refl. trivial. -- intros φ0 Hφ0 HF. apply gmultiset_elem_of_disj_union in Hφ0. destruct Hφ0 as [Hφ0|]; [|ms]. destruct (occurs_in_open_boxes _ _ _ HF Hφ0) as (θ0 & Hθ0 & Hinθ). apply (Hnin θ0); ms. diff --git a/theories/iSL/Simp.v b/theories/iSL/Simp.v index 27f6434..0ad4293 100644 --- a/theories/iSL/Simp.v +++ b/theories/iSL/Simp.v @@ -527,7 +527,7 @@ match goal with end. Qed. -Require Import ISL.InvPropQuantifiers. +Require Import ISL.PropQuantifiers ISL.Simp_env. Definition E_simplified (p: variable) (ψ: form) := simp_form (E p [ψ]). Definition A_simplified (p: variable) (ψ: form) := simp_form (Af p (ψ)). diff --git a/theories/iSL/Simp_env.v b/theories/iSL/Simp_env.v index b69628c..beb196e 100644 --- a/theories/iSL/Simp_env.v +++ b/theories/iSL/Simp_env.v @@ -2,6 +2,7 @@ Require Import Coq.Program.Equality. Require Import ISL.Sequents ISL.SequentProps. Require Import ISL.Order ISL.DecisionProcedure. Require Import Coq.Classes.RelationClasses. +Require Import ISL.Cut ISL.Optimizations. Definition applicable_AndL (Γ : list form): {ψ1 & {ψ2 | (And ψ1 ψ2) ∈ Γ}} + (∀ ψ1 ψ2, (And ψ1 ψ2) ∈ Γ -> False). Proof. @@ -125,8 +126,8 @@ Lemma simp_env_eq (Δ : list form): simp_env Δ = applicable_ImpLAnd Δ ? λ δ₁ δ₂ δ₃ _, simp_env ((rm ((δ₁ ∧ δ₂)→ δ₃) Δ • (δ₁ → (δ₂ → δ₃)))) :3 applicable_ImpLOr Δ ? λ δ₁ δ₂ δ₃ _, simp_env (rm ((δ₁ ∨ δ₂)→ δ₃) Δ • (δ₁ → δ₃) • (δ₂ → δ₃)) :3 (* remove redundant assumptions *) - applicable_strong_weakening Δ ? λ φ _, simp_env (rm φ Δ) :1 - Δ. + (applicable_strong_weakening Δ ? λ φ _, simp_env (rm φ Δ) :1 + Δ). Proof. simpl. unfold simp_env. simpl. repeat rewrite Wf.WfExtensionality.fix_sub_eq_ext; simpl. split; trivial. @@ -134,6 +135,7 @@ Qed. Definition simp_form φ:= ⋀ (simp_env [φ]). + Definition pointed_env_order_refl pe1 pe2 := env_order_refl (pe1.2 :: pe1.2 :: pe1.1) (pe2.2 :: pe2.2 :: pe2.1). (* TODO: move *) @@ -145,6 +147,20 @@ match goal with x : _ ⊎ {[+?φ+]} = _ |- _ => exfalso; eapply (gmultiset_elem_of_empty φ); setoid_rewrite <- x; ms end. Qed. +Ltac simp_env_tac := +let Hf := fresh "Hf" in +match goal with +| |- context[sumor_bind2 ?cond ?A ?B] => + let φ1 := fresh "φ1" in let φ2 := fresh "φ2" in let Hc := fresh "Hc" in + destruct cond as [(φ1 & φ2 & Hc) | Hf] +| |- context[sumor_bind3 ?cond ?A ?B] => + let φ1 := fresh "φ1" in let φ2 := fresh "φ2" in let φ3 := fresh "φ3" in let Hc := fresh "Hc" in + destruct cond as [(φ1 & φ2 & φ3 & Hc) | Hf] +| |- context[sumor_bind1 ?cond ?A ?B] => + let φ1 := fresh "φ1" in let Hc := fresh "Hc" in + destruct cond as [(φ1 & Hc) | Hf] +end; simpl. + Lemma simp_env_order Δ : env_order_refl (simp_env Δ) Δ. Proof. revert Δ. apply (well_founded_induction_type wf_env_order). intros Δ Hind. @@ -152,24 +168,27 @@ rewrite simp_env_eq. case (Δ ⊢? ⊥). { intros [Hf _]. destruct Δ as [|φ Δ]. + now apply empty_entails_not_bot in Hf. - + rewrite env_weight_add. - unfold env_order_refl. destruct φ. - * admit. - * destruct Δ as [|φ Δ]. - -- now right. - -- left. unfold env_order, ltof, env_weight. simpl. pose (Order.pow5_gt_0 (weight φ)). lia. - * left. unfold env_order, ltof, env_weight. simpl. - replace 5 with (5 ^ 1) at 1 by trivial. lia. - apply Nat.lt_le_trans with (5 ^ weight φ); [|lia]. - apply Nat.pow_lt_mono_r. lia. destruct φ. simpl. lia. - nia. - pose (Order.pow5_gt_0 (weight φ)). lia. - intro. simpl. auto with proof. - -Admitted. + + unfold env_order_refl. do 2 rewrite env_weight_add. simpl. pose (weight_pos φ). + replace 5 with (0 + 5^1) at 1 by trivial. apply Nat.add_le_mono. lia. + apply Nat.pow_le_mono_r; lia. +} +intro Hbot. +repeat simp_env_tac. +all: try (apply env_order_env_order_refl; eapply env_order_le_lt_trans; [apply Hind|]; intuition; order_tac). +apply env_order_self. +Qed. Global Hint Resolve simp_env_order : order. +Lemma simp_env_nil: simp_env [] = []. +Proof. +assert(Ho := simp_env_order []). destruct (simp_env []) as [|φ Δ]. +trivial. +contradict Ho. unfold env_order_refl. rewrite env_weight_add. +apply Nat.lt_nge. +unfold env_weight. simpl. pose(Order.pow5_gt_0 (weight φ)). lia. +Qed. + Lemma simp_env_pointed_env_order_L pe Δ φ: (pe ≺· (simp_env Δ, φ)) -> pe ≺· (Δ, φ). Proof. intro Hl. eapply env_order_lt_le_trans; eauto. simpl. auto with order. @@ -187,57 +206,208 @@ Section Equivalence. Definition equiv_form φ ψ : Type := (φ ≼ ψ) * (ψ ≼ φ). -Definition equiv_env Δ Γ: Set := (list_to_set_disj Δ ⊢ ⋀ Γ) * (list_to_set_disj Γ ⊢ ⋀ Δ). + +Definition equiv_env Δ Δ': Set := + (∀ φ, list_to_set_disj Δ ⊢ φ -> list_to_set_disj Δ' ⊢ φ) * + (∀ φ, list_to_set_disj Δ' ⊢ φ -> list_to_set_disj Δ ⊢ φ). +(* slightly too general + (∀ Γ φ, Γ ⊎ list_to_set_disj Δ ⊢ φ -> Γ ⊎ list_to_set_disj Δ' ⊢ φ) * + (∀ Γ φ, Γ ⊎ list_to_set_disj Δ' ⊢ φ -> Γ ⊎ list_to_set_disj Δ ⊢ φ). +*) Lemma symmetric_equiv_env Δ Γ: equiv_env Δ Γ -> equiv_env Γ Δ . Proof. intros [H1 H2]. split; trivial. Qed. +(* TODO: move *) +Lemma conjunction_L' Γ Δ ϕ: (Γ ⊎ {[⋀ Δ]} ⊢ ϕ) -> Γ ⊎ list_to_set_disj Δ ⊢ ϕ. +Proof. +revert ϕ. unfold conjunction. +assert( Hstrong: ∀ θ ϕ : form, Γ ⊎ {[foldl make_conj θ (nodup form_eq_dec Δ)]} ⊢ ϕ + → (Γ ⊎ list_to_set_disj Δ) ⊎ {[θ]} ⊢ ϕ). +{ + induction Δ as [|δ Δ]; intros θ ϕ; simpl. + - intro Hp. peapply Hp. + - case in_dec; intros Hin Hp. + + peapply (weakening δ). apply IHΔ, Hp. ms. + + simpl in Hp. apply IHΔ in Hp. + peapply (AndL_rev (Γ ⊎ list_to_set_disj Δ) θ δ). apply make_conj_complete_L, Hp. +} + intros; apply additive_cut with (φ := ⊤); eauto with proof. +Qed. + +Lemma conjunction_R Δ: list_to_set_disj Δ ⊢ ⋀ Δ. +Proof. +apply conjunction_R1. intros φ Hφ. apply elem_of_list_to_set_disj in Hφ. +exhibit Hφ 0. apply generalised_axiom. +Qed. + +Lemma conjunction_L'' Γ Δ ϕ: Γ ⊎ list_to_set_disj Δ ⊢ ϕ -> (Γ ⊎ {[⋀ Δ]} ⊢ ϕ). +Proof. +revert ϕ. unfold conjunction. +assert( Hstrong: ∀ θ ϕ : form,(Γ ⊎ list_to_set_disj Δ) ⊎ {[θ]} ⊢ ϕ -> Γ ⊎ {[foldl make_conj θ (nodup form_eq_dec Δ)]} ⊢ ϕ). +{ + induction Δ as [|δ Δ]; intros θ ϕ; simpl. + - intro Hp. peapply Hp. + - case in_dec; intros Hin Hp. + + apply IHΔ. + assert(Hin' : δ ∈ (Γ ⊎ list_to_set_disj Δ)). + { apply gmultiset_elem_of_disj_union; right; apply elem_of_list_to_set_disj, elem_of_list_In, Hin. } + exhibit Hin' 1. exch 0. apply contraction. exch 1. exch 0. + rw (symmetry (difference_singleton _ _ Hin')) 2. peapply Hp. + + simpl. apply IHΔ. apply make_conj_sound_L, AndL. peapply Hp. +} +intros. apply Hstrong, weakening. assumption. +Qed. + +(* Lemma equiv_env_L0 Δ Δ' φ: (equiv_env Δ Δ') -> list_to_set_disj Δ ⊢ φ -> list_to_set_disj Δ' ⊢ φ. Proof. -intros [H1 H2]. -Admitted. +intros [Hp _] H. peapply (Hp ∅). peapply H. +Qed. +*) +Lemma equiv_env_refl Δ : equiv_env Δ Δ. +Proof. split; trivial. Qed. -Lemma equiv_env_L1 Γ Δ Δ' φ: (equiv_env Δ Δ') -> - Γ ⊎ list_to_set_disj Δ ⊢ φ -> Γ ⊎ list_to_set_disj Δ' ⊢ φ. +Lemma equiv_env_trans Δ Δ' Δ'' : equiv_env Δ Δ' -> equiv_env Δ' Δ'' -> equiv_env Δ Δ''. Proof. -intros [H1 H2]. -Admitted. +intros [H11 H12] [H21 H22]. split; intros; auto. +Qed. + +(* TODO: move *) +Hint Resolve elem_of_list_to_set_disj : proof. -Lemma equiv_env_L2 Γ Δ Δ' φ: (equiv_env Δ Δ') -> - list_to_set_disj Δ ⊎ Γ ⊢ φ -> list_to_set_disj Δ' ⊎ Γ ⊢ φ. + +Ltac l_tac' := repeat rewrite list_to_set_disj_open_boxes; + rewrite (proper_Provable _ _ (symmetry (list_to_set_disj_env_add _ _)) _ _ eq_refl) +|| rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (symmetry (list_to_set_disj_env_add _ _))) _ _ eq_refl) +|| rewrite (proper_Provable _ _ (equiv_disj_union_compat_r (equiv_disj_union_compat_r (symmetry (list_to_set_disj_env_add _ _)))) _ _ eq_refl) +|| rewrite (proper_Provable _ _ (equiv_disj_union_compat_r(equiv_disj_union_compat_r (equiv_disj_union_compat_r (symmetry (list_to_set_disj_env_add _ _))))) _ _ eq_refl). + +Lemma Provable_dec_of_Prop Γ φ: (∃ _ : list_to_set_disj Γ ⊢ φ, True) -> (list_to_set_disj Γ ⊢ φ). Proof. -Admitted. +destruct (Proof_tree_dec Γ φ) as [[Hφ1' _] | Hf']. tauto. +intros Hf. exfalso. destruct Hf as [Hf _]. tauto. +Qed. Lemma simp_env_equiv_env Δ: equiv_env (simp_env Δ) Δ. Proof. -Admitted. +Local Ltac equiv_tac := + repeat rewrite <- list_to_set_disj_rm; + repeat rewrite <- list_to_set_disj_env_add; + repeat (rewrite <- difference_singleton; trivial); + try rewrite <- list_to_set_disj_rm; + try (rewrite union_difference_L by trivial); + try (rewrite union_difference_R by trivial); + try ms. +Local Ltac peapply' th := (erewrite proper_Provable; [| |reflexivity]); [eapply th|equiv_tac]. +revert Δ. apply (well_founded_induction_type wf_env_order). intros Δ Hind. +rewrite simp_env_eq. +case (Δ ⊢? ⊥). +{ intro Hf. apply Provable_dec_of_Prop in Hf. + split; intros φ Hp. + + apply exfalso. assumption. + + simpl. peapply (ExFalso ∅). +} +intro Hbot. +repeat simp_env_tac; (try (eapply equiv_env_trans; [apply Hind; intuition; order_tac|])). +- apply elem_of_list_to_set_disj in Hc. + split; intros φ Hp. + + exhibit Hc 0. apply AndL. peapply' Hp. + + do 2 l_tac'. apply AndL_rev. peapply' Hp. +- destruct Hc as [Hc1 Hc2]. apply elem_of_list_to_set_disj in Hc1, Hc2. + assert(Hc3: Var φ1 ∈ (list_to_set_disj Δ ∖ {[φ1 → φ2]} : env)). + { apply in_difference. intro HF. discriminate HF. assumption. } + split; intros φ Hp. + + exhibit Hc2 0. exhibit Hc3 1. peapply ImpLVar. peapply' Hp. + + l_tac'. peapply' (ImpLVar_rev (list_to_set_disj (rm φ1 (rm (φ1 → φ2) Δ))) φ1 φ2). + peapply' Hp. +- apply elem_of_list_to_set_disj in Hc. + split; intros φ Hp. + + exhibit Hc 0. apply ImpLAnd. peapply' Hp. + + l_tac'. apply ImpLAnd_rev. peapply' Hp. +- apply elem_of_list_to_set_disj in Hc. + split; intros φ Hp. + + exhibit Hc 0. apply ImpLOr. peapply' Hp. + + do 2 l_tac'. apply ImpLOr_rev. peapply' Hp. +- destruct Hc as [Hc Hφ1]. apply elem_of_list_to_set_disj in Hc. + split; intros φ Hp. + + exhibit Hc 0. apply weakening. peapply' Hp. + + apply Provable_dec_of_Prop in Hφ1. + apply additive_cut with φ1. trivial. peapply' Hp. +- apply equiv_env_refl. +Qed. -Lemma simp_env_L1 Γ Δ φ: - Γ ⊎ list_to_set_disj (simp_env Δ) ⊢ φ -> Γ ⊎ list_to_set_disj Δ ⊢ φ. +Lemma equiv_env_L1 Γ Δ Δ' φ: (equiv_env Δ Δ') -> + Γ ⊎ list_to_set_disj Δ ⊢ φ -> Γ ⊎ list_to_set_disj Δ' ⊢ φ. Proof. -apply equiv_env_L1, simp_env_equiv_env. +intros [H1 _] Hp. +revert φ Hp. +induction Γ as [|x Γ] using gmultiset_rec; intros φ Hp. +- peapply H1. peapply Hp. +- peapply (ImpR_rev (Γ ⊎ list_to_set_disj Δ') x). + apply IHΓ, ImpR. peapply Hp. Qed. End Equivalence. -Global Hint Resolve simp_env_L1 : proof. +(* Global Hint Resolve simp_env_L1 : proof. *) Infix "≡f" := equiv_form (at level 120). Global Infix "≡e" := equiv_env (at level 120). Section Variables. + +Local Ltac occ p := simpl; tauto || +match goal with +| Hnin : ∀ φ0 : form, φ0 ∈ ?Γ → ¬ occurs_in p φ0 |- _ => + + let Hin := fresh "Hin" in + try (match goal with |Hocc : ?a ∈ ?Γ |- _ => let Hocc' := fresh "Hocc" in assert (Hocc' := Hnin _ Hocc); simpl in Hocc'; + let Hin' := fresh "Hinx" in assert(Hin' := In_open_boxes _ _ Hocc); simpl in *; + try rewrite open_boxes_remove in * by trivial end); + intro; repeat rewrite env_in_add; repeat rewrite difference_include; simpl; + intro Hin; try tauto; + try (destruct Hin as [Hin|[Hin|Hin]] ||destruct Hin as [Hin|Hin]); + subst; simpl; try tauto; + repeat (apply difference_include in Hin; [|assumption]); + try (now apply Hnin) +end. + Lemma equiv_env_vars Δ x: (∃ θ : form, ((θ ∈ simp_env Δ) /\ occurs_in x θ)) -> ∃ θ : form, ((θ ∈ Δ) /\ occurs_in x θ). Proof. -intros [θ [Hin Hocc]]. -Admitted. +revert Δ. refine (well_founded_induction_type wf_env_order _ _). intros Δ Hind. +rewrite (simp_env_eq Δ). +case (Δ ⊢? ⊥). +{ intros _ [θ [Hin Hocc]]. apply elem_of_list_singleton in Hin. subst. contradict Hocc. } +intro Hbot. +repeat simp_env_tac; trivial. +all:(decompose record Hc; + intro Hocc; apply Hind in Hocc; [|order_tac]; destruct Hocc as [θ [Hin Hocc]]; + apply elem_of_list_to_set_disj in Hin; + repeat rewrite <- ?list_to_set_disj_env_add, gmultiset_elem_of_disj_union in Hin; + repeat ((try apply elem_of_list_to_set_disj, elem_of_list_In, in_rm, elem_of_list_In in Hin; + try apply gmultiset_elem_of_singleton in Hin; + subst; try solve[eexists; split; simpl in *; eauto; simpl in *; eauto || tauto]) || destruct Hin as [Hin | Hin])). +Qed. End Variables. Lemma simp_env_idempotent Δ: simp_env (simp_env Δ) = simp_env Δ. Proof. -Admitted. \ No newline at end of file +revert Δ. apply (well_founded_induction_type wf_env_order). intros Δ Hind. +rewrite (simp_env_eq Δ). +case (Δ ⊢? ⊥). +{ intro e. rewrite simp_env_eq. case ([⊥] ⊢? ⊥). trivial. intro Hf. contradict Hf. + simpl. peapply (ExFalso ∅). } +intro Hbot. +repeat simp_env_tac; (try (intuition; apply Hind; order_tac)). +rewrite simp_env_eq. +case (Δ ⊢? ⊥); [intros [He _]; tauto| intros _]. +repeat simp_env_tac; try (try (apply Provable_dec_of_Prop in Hc); contradict Hc; intuition; eauto); trivial. +eapply Hf3. eauto. now apply Provable_dec_of_Prop. +Qed. From 2989ca9ac1f75274c990d62a2ac727f8590c25dd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hugo=20F=C3=A9r=C3=A9e?= Date: Wed, 11 Sep 2024 15:39:53 +0200 Subject: [PATCH 12/12] typo --- theories/iSL/PropQuantifiers.v | 1 - 1 file changed, 1 deletion(-) diff --git a/theories/iSL/PropQuantifiers.v b/theories/iSL/PropQuantifiers.v index b86549d..037e12e 100644 --- a/theories/iSL/PropQuantifiers.v +++ b/theories/iSL/PropQuantifiers.v @@ -598,7 +598,6 @@ assert(Ho' : pointed_env_order_refl (elements Γ ++ simp_env Δ, ϕ) (elements assert(Hind' := λ Γ0 Δ0 ψ Ho, Hind (elements Γ0 ++ Δ0, ψ) (env_order_lt_le_trans _ _ _ Ho Ho') _ _ _ eq_refl). simpl in Hind'. clear Ho'. clear Hind. rename Hind' into Hind. rewrite <- E_simp_env, <- A_simp_env. -(simp_env_equiv_env Δ).1 Hp assert(Hp' := equiv_env_L1 Γ _ _ _ (symmetric_equiv_env _ _ (simp_env_equiv_env Δ)) Hp). remember (simp_env Δ) as Δ'. clear Hp. rename Hp' into Hp.