Skip to content
Snippets Groups Projects
Commit 352292a8 authored by Jacques-Henri Jourdan's avatar Jacques-Henri Jourdan
Browse files

We usually fail faster to prove IntoPure than Absorbing/Affine.

parent 3b03bb6d
No related branches found
No related tags found
No related merge requests found
...@@ -613,9 +613,9 @@ Proof. ...@@ -613,9 +613,9 @@ Proof.
apply pure_elim_l=> . by rewrite -(exist_intro ). apply pure_elim_l=> . by rewrite -(exist_intro ).
Qed. Qed.
Global Instance into_exist_sep_pure P Q φ : Global Instance into_exist_sep_pure P Q φ :
TCOr (Affine P) (Absorbing Q) IntoPureT P φ IntoExist (P Q) (λ _ : φ, Q). IntoPureT P φ TCOr (Affine P) (Absorbing Q) IntoExist (P Q) (λ _ : φ, Q).
Proof. Proof.
intros ? (φ'&->&?). rewrite /IntoExist. intros (φ'&->&?) ?. rewrite /IntoExist.
eapply (pure_elim φ'); [by rewrite (into_pure P); apply sep_elim_l, _|]=>?. eapply (pure_elim φ'); [by rewrite (into_pure P); apply sep_elim_l, _|]=>?.
rewrite -exist_intro //. apply sep_elim_r, _. rewrite -exist_intro //. apply sep_elim_r, _.
Qed. Qed.
...@@ -658,10 +658,10 @@ Proof. ...@@ -658,10 +658,10 @@ Proof.
intros (φ'&->&?). by rewrite /FromForall -pure_impl_forall (into_pure P). intros (φ'&->&?). by rewrite /FromForall -pure_impl_forall (into_pure P).
Qed. Qed.
Global Instance from_forall_wand_pure P Q φ : Global Instance from_forall_wand_pure P Q φ :
TCOr (Affine P) (Absorbing Q) IntoPureT P φ IntoPureT P φ TCOr (Affine P) (Absorbing Q)
FromForall (P -∗ Q)%I (λ _ : φ, Q)%I. FromForall (P -∗ Q)%I (λ _ : φ, Q)%I.
Proof. Proof.
intros [|] (φ'&->&?); rewrite /FromForall; apply wand_intro_r. intros (φ'&->&?) [|]; rewrite /FromForall; apply wand_intro_r.
- rewrite -(affine_affinely P) (into_pure P) -persistent_and_affinely_sep_r. - rewrite -(affine_affinely P) (into_pure P) -persistent_and_affinely_sep_r.
apply pure_elim_r=>?. by rewrite forall_elim. apply pure_elim_r=>?. by rewrite forall_elim.
- by rewrite (into_pure P) -pure_wand_forall wand_elim_l. - by rewrite (into_pure P) -pure_wand_forall wand_elim_l.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment