Skip to content
Snippets Groups Projects
Commit 7fe3cfc4 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Fix inconsistent lemmas for conjunction/disjunction.

parent 96778f37
No related branches found
No related tags found
No related merge requests found
...@@ -332,10 +332,10 @@ Lemma and_elim_r P Q : (P ∧ Q)%I ⊆ Q. ...@@ -332,10 +332,10 @@ Lemma and_elim_r P Q : (P ∧ Q)%I ⊆ Q.
Proof. by intros x n ? [??]. Qed. Proof. by intros x n ? [??]. Qed.
Lemma and_intro P Q R : P Q P R P (Q R)%I. Lemma and_intro P Q R : P Q P R P (Q R)%I.
Proof. intros HQ HR x n ??; split; auto. Qed. Proof. intros HQ HR x n ??; split; auto. Qed.
Lemma or_intro_l P Q R : P Q P (Q R)%I. Lemma or_intro_l P Q : P (P Q)%I.
Proof. intros HQ x n ??; left; auto. Qed. Proof. intros x n ??; left; auto. Qed.
Lemma or_intro_r P Q R : P R P (Q R)%I. Lemma or_intro_r P Q : Q (P Q)%I.
Proof. intros HR x n ??; right; auto. Qed. Proof. intros x n ??; right; auto. Qed.
Lemma or_elim R P Q : P R Q R (P Q)%I R. Lemma or_elim R P Q : P R Q R (P Q)%I R.
Proof. intros HP HQ x n ? [?|?]. by apply HP. by apply HQ. Qed. Proof. intros HP HQ x n ? [?|?]. by apply HP. by apply HQ. Qed.
Lemma impl_intro P Q R : (R P)%I Q R (P Q)%I. Lemma impl_intro P Q R : (R P)%I Q R (P Q)%I.
...@@ -375,8 +375,14 @@ Lemma and_elim_l' P Q R : P ⊆ R → (P ∧ Q)%I ⊆ R. ...@@ -375,8 +375,14 @@ Lemma and_elim_l' P Q R : P ⊆ R → (P ∧ Q)%I ⊆ R.
Proof. by rewrite and_elim_l. Qed. Proof. by rewrite and_elim_l. Qed.
Lemma and_elim_r' P Q R : Q R (P Q)%I R. Lemma and_elim_r' P Q R : Q R (P Q)%I R.
Proof. by rewrite and_elim_r. Qed. Proof. by rewrite and_elim_r. Qed.
Lemma or_intro_l' P Q R : P Q P (Q R)%I.
Hint Resolve or_elim or_intro_l or_intro_r. Proof. intros ->; apply or_intro_l. Qed.
Lemma or_intro_r' P Q R : P R P (Q R)%I.
Proof. intros ->; apply or_intro_r. Qed.
Lemma exist_intro' P `(Q : A uPred M) a : P Q a P ( a, Q a)%I.
Proof. intros ->; apply exist_intro. Qed.
Hint Resolve or_elim or_intro_l' or_intro_r'.
Hint Resolve and_intro and_elim_l' and_elim_r'. Hint Resolve and_intro and_elim_l' and_elim_r'.
Hint Immediate True_intro False_elim. Hint Immediate True_intro False_elim.
...@@ -523,11 +529,15 @@ Proof. by intros P P' HP Q Q' HQ; apply wand_mono. Qed. ...@@ -523,11 +529,15 @@ Proof. by intros P P' HP Q Q' HQ; apply wand_mono. Qed.
Global Instance sep_True : RightId () True%I (@uPred_sep M). Global Instance sep_True : RightId () True%I (@uPred_sep M).
Proof. by intros P; rewrite (commutative _), (left_id _ _). Qed. Proof. by intros P; rewrite (commutative _), (left_id _ _). Qed.
Lemma sep_elim_l P Q R : P R (P Q)%I R. Lemma sep_elim_l P Q : (P Q)%I P.
Proof. by intros HR; rewrite <-(right_id _ () R)%I, HR, (True_intro Q). Qed. Proof. by rewrite (True_intro Q), (right_id _ _). Qed.
Lemma sep_elim_r P Q : (P Q)%I Q. Lemma sep_elim_r P Q : (P Q)%I Q.
Proof. by rewrite (commutative ())%I; apply sep_elim_l. Qed. Proof. by rewrite (commutative ())%I; apply sep_elim_l. Qed.
Hint Resolve sep_elim_l sep_elim_r. Lemma sep_elim_l' P Q R : P R (P Q)%I R.
Proof. intros ->; apply sep_elim_l. Qed.
Lemma sep_elim_r' P Q R : Q R (P Q)%I R.
Proof. intros ->; apply sep_elim_r. Qed.
Hint Resolve sep_elim_l' sep_elim_r'.
Lemma sep_and P Q : (P Q)%I (P Q)%I. Lemma sep_and P Q : (P Q)%I (P Q)%I.
Proof. auto. Qed. Proof. auto. Qed.
Global Instance sep_False : LeftAbsorb () False%I (@uPred_sep M). Global Instance sep_False : LeftAbsorb () False%I (@uPred_sep M).
......
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