From 7c714af0e89f999048b9b3c188218ce95d8f560d Mon Sep 17 00:00:00 2001 From: Jacques-Henri Jourdan <jacques-henri.jourdan@normalesup.org> Date: Tue, 19 Dec 2017 15:52:25 +0100 Subject: [PATCH] Monpred : try to fix intowand instances. --- _CoqProject | 2 +- theories/proofmode/monpred.v | 90 ++++++++++++------- ...roofmode_monPred.v => proofmode_monpred.v} | 25 +++++- 3 files changed, 85 insertions(+), 32 deletions(-) rename theories/tests/{proofmode_monPred.v => proofmode_monpred.v} (57%) diff --git a/_CoqProject b/_CoqProject index e3a98170f..6b739e43e 100644 --- a/_CoqProject +++ b/_CoqProject @@ -97,7 +97,7 @@ theories/tests/heap_lang.v theories/tests/one_shot.v theories/tests/proofmode.v theories/tests/proofmode_iris.v -theories/tests/proofmode_monPred.v +theories/tests/proofmode_monpred.v theories/tests/list_reverse.v theories/tests/tree_sum.v theories/tests/ipm_paper.v diff --git a/theories/proofmode/monpred.v b/theories/proofmode/monpred.v index 0dabf37cd..f3842e4b8 100644 --- a/theories/proofmode/monpred.v +++ b/theories/proofmode/monpred.v @@ -2,18 +2,21 @@ From iris.bi Require Export monpred. From iris.proofmode Require Import tactics. Import MonPred. +Class MakeMonPredCar {I : bi_index} {PROP : bi} (i : I) + (P : monPred I PROP) (ð“Ÿ : PROP) := + make_monPred_car : P i ⊣⊢ ð“Ÿ. +Arguments MakeMonPredCar {_ _} _ _%I _%I. +Hint Mode MakeMonPredCar + + - ! -. + Section bi. Context {I : bi_index} {PROP : bi}. Local Notation monPred := (monPred I PROP). +Local Notation MakeMonPredCar := (@MakeMonPredCar I PROP). Implicit Types P Q R : monPred. Implicit Types ð“Ÿ ð“ ð“¡ : PROP. Implicit Types φ : Prop. Implicit Types i j : I. -Class MakeMonPredCar i P ð“Ÿ := - make_monPred_car : P i ⊣⊢ ð“Ÿ. -Arguments MakeMonPredCar _ _%I _%I. - Global Instance make_monPred_car_pure φ i : MakeMonPredCar i ⌜φ⌠⌜φâŒ. Proof. rewrite /MakeMonPredCar. by unseal. Qed. Global Instance make_monPred_car_internal_eq {A : ofeT} (x y : A) i : @@ -141,17 +144,21 @@ Proof. by destruct a, pe=><-; try unseal. Qed. -(* FIXME : P and Q do not go through MakeMonPredCar, which prevent - lazilly unfolding of monPred connectives there. The reason we - cannot do that is that IntoWand is sometimes called with known - values, which we do not wnat to replace with existentials. *) -Global Instance into_wand_monPred_car p q R P Q i : - IntoWand p q R P Q → IntoWand p q (R i) (P i) (Q i). +Lemma into_wand_monPred_car_unknown_unknown p q R P ð“Ÿ Q ð“ i : + IntoWand p q R P Q → MakeMonPredCar i P 𓟠→ MakeMonPredCar i Q ð“ → + IntoWand p q (R i) ð“Ÿ ð“ . Proof. rewrite /IntoWand /MakeMonPredCar /bi_affinely_if /bi_persistently_if. - destruct p, q=> /bi.wand_elim_l' [/(_ i) H]; apply bi.wand_intro_r; + destruct p, q=> /bi.wand_elim_l' [/(_ i) H] <- <-; apply bi.wand_intro_r; revert H; unseal; done. Qed. +Lemma into_wand_monPred_car_unknown_known p q R P ð“Ÿ Q i : + IntoWand p q R P Q → MakeMonPredCar i P 𓟠→ IntoWand p q (R i) ð“Ÿ (Q i). +Proof. intros. eapply into_wand_monPred_car_unknown_unknown=>//. apply _. Qed. +Lemma into_wand_monPred_car_known_unknown p q R P Q ð“ i : + IntoWand p q R P Q → MakeMonPredCar i Q ð“ → IntoWand p q (R i) (P i) ð“ . +Proof. intros. eapply into_wand_monPred_car_unknown_unknown=>//. apply _. Qed. + Global Instance into_wand_wand'_monPred p q P Q ð“Ÿ ð“ i : IntoWand' p q ((P -∗ Q) i) ð“Ÿ ð“ → IntoWand p q ((P -∗ Q) i) ð“Ÿ ð“ | 100. Proof. done. Qed. @@ -163,31 +170,37 @@ Global Instance from_forall_monPred_car_wand P Q (Φ Ψ : I → PROP) i : (∀ j, MakeMonPredCar j P (Φ j)) → (∀ j, MakeMonPredCar j Q (Ψ j)) → FromForall ((P -∗ Q) i)%I (λ j, ⌜i ⊑ j⌠→ Φ j -∗ Ψ j)%I. Proof. - rewrite /FromForall /MakeMonPredCar. unseal=> H1 H2. - setoid_rewrite H1. setoid_rewrite H2. done. -Qed. -Global Instance into_forall_monPred_car_wand P Q (Φ Ψ : I → PROP) i : - (∀ j, MakeMonPredCar j P (Φ j)) → (∀ j, MakeMonPredCar j Q (Ψ j)) → - IntoForall ((P -∗ Q) i) (λ j, ⌜i ⊑ j⌠→ Φ j -∗ Ψ j)%I. -Proof. - rewrite /IntoForall /MakeMonPredCar. unseal=> H1 H2. - setoid_rewrite H1. setoid_rewrite H2. done. + rewrite /FromForall /MakeMonPredCar. unseal=> H1 H2. do 2 f_equiv. + by rewrite H1 H2. Qed. Global Instance from_forall_monPred_car_impl P Q (Φ Ψ : I → PROP) i : (∀ j, MakeMonPredCar j P (Φ j)) → (∀ j, MakeMonPredCar j Q (Ψ j)) → FromForall ((P → Q) i)%I (λ j, ⌜i ⊑ j⌠→ Φ j → Ψ j)%I. Proof. - rewrite /FromForall /MakeMonPredCar. unseal=> H1 H2. - setoid_rewrite H1. setoid_rewrite H2. done. -Qed. -Global Instance into_forall_monPred_car_impl P Q (Φ Ψ : I → PROP) i : - (∀ j, MakeMonPredCar j P (Φ j)) → (∀ j, MakeMonPredCar j Q (Ψ j)) → - IntoForall ((P → Q) i) (λ j, ⌜i ⊑ j⌠→ Φ j → Ψ j)%I. -Proof. - rewrite /IntoForall /MakeMonPredCar. unseal=> H1 H2. - setoid_rewrite H1. setoid_rewrite H2. done. + rewrite /FromForall /MakeMonPredCar. unseal=> H1 H2. do 2 f_equiv. + by rewrite H1 H2 bi.pure_impl_forall. Qed. +(* FIXME : I disable these instances, because I prefer not adding the + corresponding IntoWand instance. Adding the IntoWand instance would + be ambiguous, because we want to be abble to iApply without proving + view inclusion when they match (i.e., we want to use + [into_wand_monPred_car]). *) +(* Global Instance into_forall_monPred_car_wand P Q (Φ Ψ : I → PROP) i : *) +(* (∀ j, MakeMonPredCar j P (Φ j)) → (∀ j, MakeMonPredCar j Q (Ψ j)) → *) +(* IntoForall ((P -∗ Q) i) (λ j, ⌜i ⊑ j⌠→ Φ j -∗ Ψ j)%I. *) +(* Proof. *) +(* rewrite /IntoForall /MakeMonPredCar. unseal=> H1 H2. do 2 f_equiv. *) +(* by rewrite H1 H2. *) +(* Qed. *) +(* Global Instance into_forall_monPred_car_impl P Q (Φ Ψ : I → PROP) i : *) +(* (∀ j, MakeMonPredCar j P (Φ j)) → (∀ j, MakeMonPredCar j Q (Ψ j)) → *) +(* IntoForall ((P → Q) i) (λ j, ⌜i ⊑ j⌠→ Φ j → Ψ j)%I. *) +(* Proof. *) +(* rewrite /IntoForall /MakeMonPredCar. unseal=> H1 H2. do 2 f_equiv. *) +(* by rewrite H1 H2 bi.pure_impl_forall. *) +(* Qed. *) + Global Instance from_and_monPred_car P Q1 ð“ 1 Q2 ð“ 2 i : FromAnd P Q1 Q2 → MakeMonPredCar i Q1 ð“ 1 → MakeMonPredCar i Q2 ð“ 2 → FromAnd (P i) ð“ 1 ð“ 2. @@ -272,7 +285,24 @@ Global Instance from_modal_monPred_car i P Q ð“ : Proof. by rewrite /FromModal /MakeMonPredCar=> <- <-. Qed. End bi. -Hint Mode MakeMonPredCar + + - ! -. +(* When P and/or Q are evars when doing typeclass search on [IntoWand + (R i) P Q], we use [MakeMonPredCar] in order to normalize the + result of unification. However, when they are not evars, we want to + propagate the known information through typeclass search. Hence, we + do not want to use [MakeMonPredCar]. + + As a result, depending on P and Q being evars, we use a different + version of [into_wand_monPred_car_xx_xx]. *) +Hint Extern 3 (IntoWand _ _ (monPred_car _ _) ?P ?Q) => + is_evar P; is_evar Q; + eapply @into_wand_monPred_car_unknown_unknown + : typeclass_instances. +Hint Extern 2 (IntoWand _ _ (monPred_car _ _) ?P (monPred_car ?Q _)) => + eapply @into_wand_monPred_car_unknown_known + : typeclass_instances. +Hint Extern 2 (IntoWand _ _ (monPred_car _ _) (monPred_car ?P _) ?Q) => + eapply @into_wand_monPred_car_known_unknown + : typeclass_instances. Section sbi. Context {I : bi_index} {PROP : sbi}. diff --git a/theories/tests/proofmode_monPred.v b/theories/tests/proofmode_monpred.v similarity index 57% rename from theories/tests/proofmode_monPred.v rename to theories/tests/proofmode_monpred.v index e9f2053ea..a18ce8ef0 100644 --- a/theories/tests/proofmode_monPred.v +++ b/theories/tests/proofmode_monpred.v @@ -27,4 +27,27 @@ Section tests. Lemma test_iStartProof_7 P : ((P ≡ P)%I : monPredI). Proof. iStartProof PROP. done. Qed. -End tests. \ No newline at end of file + Lemma test_intowand_1 P Q : (P -∗ Q) -∗ P -∗ Q. + Proof. + iStartProof PROP. iIntros (i) "HW". iIntros (j ->) "HP". by iApply "HW". + Qed. + Lemma test_intowand_2 P Q : (P -∗ Q) -∗ P -∗ Q. + Proof. + iStartProof PROP. iIntros (i) "HW". iIntros (j ->) "HP". + iSpecialize ("HW" with "[HP //]"). done. + Qed. + Lemma test_intowand_3 P Q : (P -∗ Q) -∗ P -∗ Q. + Proof. + iStartProof PROP. iIntros (i) "HW". iIntros (j ->) "HP". + iSpecialize ("HW" with "HP"). done. + Qed. + Lemma test_intowand_4 P Q : (P -∗ Q) -∗ â–· P -∗ â–· Q. + Proof. + iStartProof PROP. iIntros (i) "HW". iIntros (j ->) "HP". by iApply "HW". + Qed. + Lemma test_intowand_5 P Q : (P -∗ Q) -∗ â–· P -∗ â–· Q. + Proof. + iStartProof PROP. iIntros (i) "HW". iIntros (j ->) "HP". + iSpecialize ("HW" with "HP"). done. + Qed. +End tests. -- GitLab