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

Shorter names for common math notions.

Also do some minor clean up.
parent ba6d9390
No related branches found
No related tags found
No related merge requests found
......@@ -59,9 +59,9 @@ Program Instance agree_op : Op (agree A) := λ x y,
Next Obligation. naive_solver eauto using agree_valid_S, dist_S. Qed.
Instance agree_unit : Unit (agree A) := id.
Instance agree_minus : Minus (agree A) := λ x y, x.
Instance: Commutative () (@op (agree A) _).
Instance: Comm () (@op (agree A) _).
Proof. intros x y; split; [naive_solver|by intros n (?&?&Hxy); apply Hxy]. Qed.
Definition agree_idempotent (x : agree A) : x x x.
Definition agree_idemp (x : agree A) : x x x.
Proof. split; naive_solver. Qed.
Instance: n : nat, Proper (dist n ==> impl) (@validN (agree A) _ n).
Proof.
......@@ -79,18 +79,18 @@ Proof.
eauto using agree_valid_le.
Qed.
Instance: Proper (dist n ==> dist n ==> dist n) (@op (agree A) _).
Proof. by intros n x1 x2 Hx y1 y2 Hy; rewrite Hy !(commutative _ _ y2) Hx. Qed.
Proof. by intros n x1 x2 Hx y1 y2 Hy; rewrite Hy !(comm _ _ y2) Hx. Qed.
Instance: Proper (() ==> () ==> ()) op := ne_proper_2 _.
Instance: Associative () (@op (agree A) _).
Instance: Assoc () (@op (agree A) _).
Proof.
intros x y z; split; simpl; intuition;
repeat match goal with H : agree_is_valid _ _ |- _ => clear H end;
by cofe_subst; rewrite !agree_idempotent.
by cofe_subst; rewrite !agree_idemp.
Qed.
Lemma agree_includedN (x y : agree A) n : x {n} y y {n} x y.
Proof.
split; [|by intros ?; exists y].
by intros [z Hz]; rewrite Hz (associative _) agree_idempotent.
by intros [z Hz]; rewrite Hz assoc agree_idemp.
Qed.
Definition agree_cmra_mixin : CMRAMixin (agree A).
Proof.
......@@ -99,7 +99,7 @@ Proof.
* intros n x [? Hx]; split; [by apply agree_valid_S|intros n' ?].
rewrite (Hx n'); last auto.
symmetry; apply dist_le with n; try apply Hx; auto.
* intros x; apply agree_idempotent.
* intros x; apply agree_idemp.
* by intros x y n [(?&?&?) ?].
* by intros x y n; rewrite agree_includedN.
Qed.
......@@ -108,13 +108,13 @@ Proof. intros Hxy; apply Hxy. Qed.
Lemma agree_valid_includedN (x y : agree A) n : {n} y x {n} y x {n} y.
Proof.
move=> Hval [z Hy]; move: Hval; rewrite Hy.
by move=> /agree_op_inv->; rewrite agree_idempotent.
by move=> /agree_op_inv->; rewrite agree_idemp.
Qed.
Definition agree_cmra_extend_mixin : CMRAExtendMixin (agree A).
Proof.
intros n x y1 y2 Hval Hx; exists (x,x); simpl; split.
* by rewrite agree_idempotent.
* by move: Hval; rewrite Hx; move=> /agree_op_inv->; rewrite agree_idempotent.
* by rewrite agree_idemp.
* by move: Hval; rewrite Hx; move=> /agree_op_inv->; rewrite agree_idemp.
Qed.
Canonical Structure agreeRA : cmraT :=
CMRAT agree_cofe_mixin agree_cmra_mixin agree_cmra_extend_mixin.
......@@ -125,7 +125,7 @@ Solve Obligations with done.
Global Instance to_agree_ne n : Proper (dist n ==> dist n) to_agree.
Proof. intros x1 x2 Hx; split; naive_solver eauto using @dist_le. Qed.
Global Instance to_agree_proper : Proper (() ==> ()) to_agree := ne_proper _.
Global Instance to_agree_inj n : Injective (dist n) (dist n) (to_agree).
Global Instance to_agree_inj n : Inj (dist n) (dist n) (to_agree).
Proof. by intros x y [_ Hxy]; apply Hxy. Qed.
Lemma to_agree_car n (x : agree A) : {n} x to_agree (x n) {n} x.
Proof. intros [??]; split; naive_solver eauto using agree_valid_le. Qed.
......
......@@ -106,10 +106,10 @@ Proof.
* by intros n x1 x2 [Hx Hx'] y1 y2 [Hy Hy'];
split; simpl; rewrite ?Hy ?Hy' ?Hx ?Hx'.
* intros n [[] ?] ?; naive_solver eauto using cmra_includedN_S, cmra_validN_S.
* by split; simpl; rewrite associative.
* by split; simpl; rewrite commutative.
* by split; simpl; rewrite assoc.
* by split; simpl; rewrite comm.
* by split; simpl; rewrite ?cmra_unit_l.
* by split; simpl; rewrite ?cmra_unit_idempotent.
* by split; simpl; rewrite ?cmra_unit_idemp.
* intros n ??; rewrite! auth_includedN; intros [??].
by split; simpl; apply cmra_unit_preservingN.
* assert ( n (a b1 b2 : A), b1 b2 {n} a b1 {n} a).
......@@ -153,8 +153,8 @@ Lemma auth_update a a' b b' :
Proof.
move=> Hab [[?| |] bf1] n // =>-[[bf2 Ha] ?]; do 2 red; simpl in *.
destruct (Hab n (bf1 bf2)) as [Ha' ?]; auto.
{ by rewrite Ha left_id associative. }
split; [by rewrite Ha' left_id associative; apply cmra_includedN_l|done].
{ by rewrite Ha left_id assoc. }
split; [by rewrite Ha' left_id assoc; apply cmra_includedN_l|done].
Qed.
Lemma auth_local_update L `{!LocalUpdate Lv L} a a' :
......@@ -170,7 +170,7 @@ Lemma auth_update_op_l a a' b :
Proof. by intros; apply (auth_local_update _). Qed.
Lemma auth_update_op_r a a' b :
(a b) a a' ~~> (a b) (a' b).
Proof. rewrite -!(commutative _ b); apply auth_update_op_l. Qed.
Proof. rewrite -!(comm _ b); apply auth_update_op_l. Qed.
(* This does not seem to follow from auth_local_update.
The trouble is that given ✓ (L a ⋅ a'), Lv a
......
......@@ -43,10 +43,10 @@ Record CMRAMixin A `{Dist A, Equiv A, Unit A, Op A, ValidN A, Minus A} := {
(* valid *)
mixin_cmra_validN_S n x : {S n} x {n} x;
(* monoid *)
mixin_cmra_associative : Associative () ();
mixin_cmra_commutative : Commutative () ();
mixin_cmra_assoc : Assoc () ();
mixin_cmra_comm : Comm () ();
mixin_cmra_unit_l x : unit x x x;
mixin_cmra_unit_idempotent x : unit (unit x) unit x;
mixin_cmra_unit_idemp x : unit (unit x) unit x;
mixin_cmra_unit_preservingN n x y : x {n} y unit x {n} unit y;
mixin_cmra_validN_op_l n x y : {n} (x y) {n} x;
mixin_cmra_op_minus n x y : x {n} y x y x {n} y
......@@ -101,14 +101,14 @@ Section cmra_mixin.
Proof. apply (mixin_cmra_minus_ne _ (cmra_mixin A)). Qed.
Lemma cmra_validN_S n x : {S n} x {n} x.
Proof. apply (mixin_cmra_validN_S _ (cmra_mixin A)). Qed.
Global Instance cmra_associative : Associative () (@op A _).
Proof. apply (mixin_cmra_associative _ (cmra_mixin A)). Qed.
Global Instance cmra_commutative : Commutative () (@op A _).
Proof. apply (mixin_cmra_commutative _ (cmra_mixin A)). Qed.
Global Instance cmra_assoc : Assoc () (@op A _).
Proof. apply (mixin_cmra_assoc _ (cmra_mixin A)). Qed.
Global Instance cmra_comm : Comm () (@op A _).
Proof. apply (mixin_cmra_comm _ (cmra_mixin A)). Qed.
Lemma cmra_unit_l x : unit x x x.
Proof. apply (mixin_cmra_unit_l _ (cmra_mixin A)). Qed.
Lemma cmra_unit_idempotent x : unit (unit x) unit x.
Proof. apply (mixin_cmra_unit_idempotent _ (cmra_mixin A)). Qed.
Lemma cmra_unit_idemp x : unit (unit x) unit x.
Proof. apply (mixin_cmra_unit_idemp _ (cmra_mixin A)). Qed.
Lemma cmra_unit_preservingN n x y : x {n} y unit x {n} unit y.
Proof. apply (mixin_cmra_unit_preservingN _ (cmra_mixin A)). Qed.
Lemma cmra_validN_op_l n x y : {n} (x y) {n} x.
......@@ -166,7 +166,7 @@ Proof. apply (ne_proper _). Qed.
Global Instance cmra_op_ne' n : Proper (dist n ==> dist n ==> dist n) (@op A _).
Proof.
intros x1 x2 Hx y1 y2 Hy.
by rewrite Hy (commutative _ x1) Hx (commutative _ y2).
by rewrite Hy (comm _ x1) Hx (comm _ y2).
Qed.
Global Instance ra_op_proper' : Proper (() ==> () ==> ()) (@op A _).
Proof. apply (ne_proper_2 _). Qed.
......@@ -217,15 +217,15 @@ Proof. induction 2; eauto using cmra_validN_S. Qed.
Lemma cmra_valid_op_l x y : (x y) x.
Proof. rewrite !cmra_valid_validN; eauto using cmra_validN_op_l. Qed.
Lemma cmra_validN_op_r x y n : {n} (x y) {n} y.
Proof. rewrite (commutative _ x); apply cmra_validN_op_l. Qed.
Proof. rewrite (comm _ x); apply cmra_validN_op_l. Qed.
Lemma cmra_valid_op_r x y : (x y) y.
Proof. rewrite !cmra_valid_validN; eauto using cmra_validN_op_r. Qed.
(** ** Units *)
Lemma cmra_unit_r x : x unit x x.
Proof. by rewrite (commutative _ x) cmra_unit_l. Qed.
Proof. by rewrite (comm _ x) cmra_unit_l. Qed.
Lemma cmra_unit_unit x : unit x unit x unit x.
Proof. by rewrite -{2}(cmra_unit_idempotent x) cmra_unit_r. Qed.
Proof. by rewrite -{2}(cmra_unit_idemp x) cmra_unit_r. Qed.
Lemma cmra_unit_validN x n : {n} x {n} unit x.
Proof. rewrite -{1}(cmra_unit_l x); apply cmra_validN_op_l. Qed.
Lemma cmra_unit_valid x : x unit x.
......@@ -243,7 +243,7 @@ Proof.
split.
* by intros x; exists (unit x); rewrite cmra_unit_r.
* intros x y z [z1 Hy] [z2 Hz]; exists (z1 z2).
by rewrite (associative _) -Hy -Hz.
by rewrite assoc -Hy -Hz.
Qed.
Global Instance cmra_included_preorder: PreOrder (@included A _ _).
Proof.
......@@ -265,22 +265,22 @@ Proof. by exists y. Qed.
Lemma cmra_included_l x y : x x y.
Proof. by exists y. Qed.
Lemma cmra_includedN_r n x y : y {n} x y.
Proof. rewrite (commutative op); apply cmra_includedN_l. Qed.
Proof. rewrite (comm op); apply cmra_includedN_l. Qed.
Lemma cmra_included_r x y : y x y.
Proof. rewrite (commutative op); apply cmra_included_l. Qed.
Proof. rewrite (comm op); apply cmra_included_l. Qed.
Lemma cmra_unit_preserving x y : x y unit x unit y.
Proof. rewrite !cmra_included_includedN; eauto using cmra_unit_preservingN. Qed.
Lemma cmra_included_unit x : unit x x.
Proof. by exists x; rewrite cmra_unit_l. Qed.
Lemma cmra_preservingN_l n x y z : x {n} y z x {n} z y.
Proof. by intros [z1 Hz1]; exists z1; rewrite Hz1 (associative op). Qed.
Proof. by intros [z1 Hz1]; exists z1; rewrite Hz1 (assoc op). Qed.
Lemma cmra_preserving_l x y z : x y z x z y.
Proof. by intros [z1 Hz1]; exists z1; rewrite Hz1 (associative op). Qed.
Proof. by intros [z1 Hz1]; exists z1; rewrite Hz1 (assoc op). Qed.
Lemma cmra_preservingN_r n x y z : x {n} y x z {n} y z.
Proof. by intros; rewrite -!(commutative _ z); apply cmra_preservingN_l. Qed.
Proof. by intros; rewrite -!(comm _ z); apply cmra_preservingN_l. Qed.
Lemma cmra_preserving_r x y z : x y x z y z.
Proof. by intros; rewrite -!(commutative _ z); apply cmra_preserving_l. Qed.
Proof. by intros; rewrite -!(comm _ z); apply cmra_preserving_l. Qed.
Lemma cmra_included_dist_l x1 x2 x1' n :
x1 x2 x1' {n} x1 x2', x1' x2' x2' {n} x2.
......@@ -321,7 +321,7 @@ Section identity.
Lemma cmra_empty_least x : x.
Proof. by exists x; rewrite left_id. Qed.
Global Instance cmra_empty_right_id : RightId () ().
Proof. by intros x; rewrite (commutative op) left_id. Qed.
Proof. by intros x; rewrite (comm op) left_id. Qed.
Lemma cmra_unit_empty : unit ∅.
Proof. by rewrite -{2}(cmra_unit_l ) right_id. Qed.
End identity.
......@@ -336,7 +336,7 @@ Lemma local_update L `{!LocalUpdate Lv L} x y :
Proof. by rewrite equiv_dist=>?? n; apply (local_updateN L). Qed.
Global Instance local_update_op x : LocalUpdate (λ _, True) (op x).
Proof. split. apply _. by intros n y1 y2 _ _; rewrite associative. Qed.
Proof. split. apply _. by intros n y1 y2 _ _; rewrite assoc. Qed.
(** ** Updates *)
Global Instance cmra_update_preorder : PreOrder (@cmra_update A).
......@@ -366,10 +366,10 @@ Lemma cmra_updateP_op (P1 P2 Q : A → Prop) x1 x2 :
x1 ~~>: P1 x2 ~~>: P2 ( y1 y2, P1 y1 P2 y2 Q (y1 y2)) x1 x2 ~~>: Q.
Proof.
intros Hx1 Hx2 Hy z n ?.
destruct (Hx1 (x2 z) n) as (y1&?&?); first by rewrite associative.
destruct (Hx1 (x2 z) n) as (y1&?&?); first by rewrite assoc.
destruct (Hx2 (y1 z) n) as (y2&?&?);
first by rewrite associative (commutative _ x2) -associative.
exists (y1 y2); split; last rewrite (commutative _ y1) -associative; auto.
first by rewrite assoc (comm _ x2) -assoc.
exists (y1 y2); split; last rewrite (comm _ y1) -assoc; auto.
Qed.
Lemma cmra_updateP_op' (P1 P2 : A Prop) x1 x2 :
x1 ~~>: P1 x2 ~~>: P2 x1 x2 ~~>: λ y, y1 y2, y = y1 y2 P1 y1 P2 y2.
......@@ -448,10 +448,10 @@ Class RA A `{Equiv A, Unit A, Op A, Valid A, Minus A} := {
ra_validN_ne :> Proper (() ==> impl) valid;
ra_minus_ne :> Proper (() ==> () ==> ()) minus;
(* monoid *)
ra_associative :> Associative () ();
ra_commutative :> Commutative () ();
ra_assoc :> Assoc () ();
ra_comm :> Comm () ();
ra_unit_l x : unit x x x;
ra_unit_idempotent x : unit (unit x) unit x;
ra_unit_idemp x : unit (unit x) unit x;
ra_unit_preserving x y : x y unit x unit y;
ra_valid_op_l x y : (x y) x;
ra_op_minus x y : x y x y x y
......@@ -524,10 +524,10 @@ Section prod.
* by intros n x1 x2 [Hx1 Hx2] y1 y2 [Hy1 Hy2];
split; rewrite /= ?Hx1 ?Hx2 ?Hy1 ?Hy2.
* by intros n x [??]; split; apply cmra_validN_S.
* split; simpl; apply (associative _).
* split; simpl; apply (commutative _).
* split; simpl; apply cmra_unit_l.
* split; simpl; apply cmra_unit_idempotent.
* by split; rewrite /= assoc.
* by split; rewrite /= comm.
* by split; rewrite /= cmra_unit_l.
* by split; rewrite /= cmra_unit_idemp.
* intros n x y; rewrite !prod_includedN.
by intros [??]; split; apply cmra_unit_preservingN.
* intros n x y [??]; split; simpl in *; eauto using cmra_validN_op_l.
......
......@@ -22,21 +22,21 @@ Global Instance big_op_permutation : Proper ((≡ₚ) ==> (≡)) big_op.
Proof.
induction 1 as [|x xs1 xs2 ? IH|x y xs|xs1 xs2 xs3]; simpl; auto.
* by rewrite IH.
* by rewrite !(associative _) (commutative _ x).
* by rewrite !assoc (comm _ x).
* by transitivity (big_op xs2).
Qed.
Global Instance big_op_proper : Proper (() ==> ()) big_op.
Proof. by induction 1; simpl; repeat apply (_ : Proper (_ ==> _ ==> _) op). Qed.
Lemma big_op_app xs ys : big_op (xs ++ ys) big_op xs big_op ys.
Proof.
induction xs as [|x xs IH]; simpl; first by rewrite ?(left_id _ _).
by rewrite IH (associative _).
induction xs as [|x xs IH]; simpl; first by rewrite ?left_id.
by rewrite IH assoc.
Qed.
Lemma big_op_contains xs ys : xs `contains` ys big_op xs big_op ys.
Proof.
induction 1 as [|x xs ys|x y xs|x xs ys|xs ys zs]; rewrite //=.
* by apply cmra_preserving_l.
* by rewrite !associative (commutative _ y).
* by rewrite !assoc (comm _ y).
* by transitivity (big_op ys); last apply cmra_included_r.
* by transitivity (big_op ys).
Qed.
......@@ -58,7 +58,7 @@ Qed.
Lemma big_opM_singleton i x : big_opM ({[i x]} : M A) x.
Proof.
rewrite -insert_empty big_opM_insert /=; last auto using lookup_empty.
by rewrite big_opM_empty (right_id _ _).
by rewrite big_opM_empty right_id.
Qed.
Global Instance big_opM_proper : Proper (() ==> ()) (big_opM : M A _).
Proof.
......
......@@ -25,7 +25,7 @@ Module ra_reflection. Section ra_reflection.
eval Σ e big_op ((λ n, from_option (Σ !! n)) <$> flatten e).
Proof.
by induction e as [| |e1 IH1 e2 IH2];
rewrite /= ?(right_id _ _) ?fmap_app ?big_op_app ?IH1 ?IH2.
rewrite /= ?right_id ?fmap_app ?big_op_app ?IH1 ?IH2.
Qed.
Lemma flatten_correct Σ e1 e2 :
flatten e1 `contains` flatten e2 eval Σ e1 eval Σ e2.
......
......@@ -337,7 +337,7 @@ Section later.
Canonical Structure laterC : cofeT := CofeT later_cofe_mixin.
Global Instance Next_contractive : Contractive (@Next A).
Proof. intros [|n] x y Hxy; [done|]; apply Hxy; lia. Qed.
Global Instance Later_inj n : Injective (dist n) (dist (S n)) (@Next A).
Global Instance Later_inj n : Inj (dist n) (dist (S n)) (@Next A).
Proof. by intros x y. Qed.
End later.
......
......@@ -46,14 +46,14 @@ Class DRA A `{Equiv A, Valid A, Unit A, Disjoint A, Op A, Minus A} := {
dra_unit_valid x : x unit x;
dra_minus_valid x y : x y x y (y x);
(* monoid *)
dra_associative :> Associative () ();
dra_assoc :> Assoc () ();
dra_disjoint_ll x y z : x y z x y x y z x z;
dra_disjoint_move_l x y z : x y z x y x y z x y z;
dra_symmetric :> Symmetric (@disjoint A _);
dra_commutative x y : x y x y x y y x;
dra_comm x y : x y x y x y y x;
dra_unit_disjoint_l x : x unit x x;
dra_unit_l x : x unit x x x;
dra_unit_idempotent x : x unit (unit x) unit x;
dra_unit_idemp x : x unit (unit x) unit x;
dra_unit_preserving x y : x y x y unit x unit y;
dra_disjoint_minus x y : x y x y x y x;
dra_op_minus x y : x y x y x y x y
......@@ -73,12 +73,12 @@ Qed.
Lemma dra_disjoint_rl x y z : x y z y z x y z x y.
Proof. intros ???. rewrite !(symmetry_iff _ x). by apply dra_disjoint_ll. Qed.
Lemma dra_disjoint_lr x y z : x y z x y x y z y z.
Proof. intros ????. rewrite dra_commutative //. by apply dra_disjoint_ll. Qed.
Proof. intros ????. rewrite dra_comm //. by apply dra_disjoint_ll. Qed.
Lemma dra_disjoint_move_r x y z :
x y z y z x y z x y z.
Proof.
intros; symmetry; rewrite dra_commutative; eauto using dra_disjoint_rl.
apply dra_disjoint_move_l; auto; by rewrite dra_commutative.
intros; symmetry; rewrite dra_comm; eauto using dra_disjoint_rl.
apply dra_disjoint_move_l; auto; by rewrite dra_comm.
Qed.
Hint Immediate dra_disjoint_move_l dra_disjoint_move_r.
Hint Unfold dra_included.
......@@ -114,11 +114,11 @@ Proof.
+ exists z. by rewrite Hx ?Hy; tauto.
* intros [x px ?] [y py ?] [z pz ?]; split; simpl;
[intuition eauto 2 using dra_disjoint_lr, dra_disjoint_rl
|intros; apply (associative _)].
* intros [x px ?] [y py ?]; split; naive_solver eauto using dra_commutative.
|by intros; rewrite assoc].
* intros [x px ?] [y py ?]; split; naive_solver eauto using dra_comm.
* intros [x px ?]; split;
naive_solver eauto using dra_unit_l, dra_unit_disjoint_l.
* intros [x px ?]; split; naive_solver eauto using dra_unit_idempotent.
* intros [x px ?]; split; naive_solver eauto using dra_unit_idemp.
* intros x y Hxy; exists (unit y unit x).
destruct x as [x px ?], y as [y py ?], Hxy as [[z pz ?] [??]]; simpl in *.
assert (py unit x unit y)
......
......@@ -31,9 +31,9 @@ Global Instance Excl_ne : Proper (dist n ==> dist n) (@Excl A).
Proof. by constructor. Qed.
Global Instance Excl_proper : Proper (() ==> ()) (@Excl A).
Proof. by constructor. Qed.
Global Instance Excl_inj : Injective () () (@Excl A).
Global Instance Excl_inj : Inj () () (@Excl A).
Proof. by inversion_clear 1. Qed.
Global Instance Excl_dist_inj n : Injective (dist n) (dist n) (@Excl A).
Global Instance Excl_dist_inj n : Inj (dist n) (dist n) (@Excl A).
Proof. by inversion_clear 1. Qed.
Program Definition excl_chain
(c : chain (excl A)) (x : A) (H : maybe Excl (c 1) = Some x) : chain A :=
......
......@@ -121,10 +121,10 @@ Proof.
* by intros n m1 m2 Hm ? i; rewrite -(Hm i).
* by intros n m1 m1' Hm1 m2 m2' Hm2 i; rewrite !lookup_minus (Hm1 i) (Hm2 i).
* intros n m Hm i; apply cmra_validN_S, Hm.
* by intros m1 m2 m3 i; rewrite !lookup_op associative.
* by intros m1 m2 i; rewrite !lookup_op commutative.
* by intros m1 m2 m3 i; rewrite !lookup_op assoc.
* by intros m1 m2 i; rewrite !lookup_op comm.
* by intros m i; rewrite lookup_op !lookup_unit cmra_unit_l.
* by intros m i; rewrite !lookup_unit cmra_unit_idempotent.
* by intros m i; rewrite !lookup_unit cmra_unit_idemp.
* intros n x y; rewrite !map_includedN_spec; intros Hm i.
by rewrite !lookup_unit; apply cmra_unit_preservingN.
* intros n m1 m2 Hm i; apply cmra_validN_op_l with (m2 !! i).
......
......@@ -141,10 +141,10 @@ Section iprod_cmra.
* by intros n f1 f2 Hf ? x; rewrite -(Hf x).
* by intros n f f' Hf g g' Hg i; rewrite iprod_lookup_minus (Hf i) (Hg i).
* intros n f Hf x; apply cmra_validN_S, Hf.
* by intros f1 f2 f3 x; rewrite iprod_lookup_op associative.
* by intros f1 f2 x; rewrite iprod_lookup_op commutative.
* by intros f1 f2 f3 x; rewrite iprod_lookup_op assoc.
* by intros f1 f2 x; rewrite iprod_lookup_op comm.
* by intros f x; rewrite iprod_lookup_op iprod_lookup_unit cmra_unit_l.
* by intros f x; rewrite iprod_lookup_unit cmra_unit_idempotent.
* by intros f x; rewrite iprod_lookup_unit cmra_unit_idemp.
* intros n f1 f2; rewrite !iprod_includedN_spec=> Hf x.
by rewrite iprod_lookup_unit; apply cmra_unit_preservingN, Hf.
* intros n f1 f2 Hf x; apply cmra_validN_op_l with (f2 x), Hf.
......
......@@ -45,7 +45,7 @@ Global Instance Some_ne : Proper (dist n ==> dist n) (@Some A).
Proof. by constructor. Qed.
Global Instance is_Some_ne n : Proper (dist n ==> iff) (@is_Some A).
Proof. inversion_clear 1; split; eauto. Qed.
Global Instance Some_dist_inj : Injective (dist n) (dist n) (@Some A).
Global Instance Some_dist_inj : Inj (dist n) (dist n) (@Some A).
Proof. by inversion_clear 1. Qed.
Global Instance None_timeless : Timeless (@None A).
Proof. inversion_clear 1; constructor. Qed.
......@@ -92,10 +92,10 @@ Proof.
* by destruct 1; rewrite /validN /option_validN //=; cofe_subst.
* by destruct 1; inversion_clear 1; constructor; cofe_subst.
* intros n [x|]; unfold validN, option_validN; eauto using cmra_validN_S.
* intros [x|] [y|] [z|]; constructor; rewrite ?associative; auto.
* intros [x|] [y|]; constructor; rewrite 1?commutative; auto.
* intros [x|] [y|] [z|]; constructor; rewrite ?assoc; auto.
* intros [x|] [y|]; constructor; rewrite 1?comm; auto.
* by intros [x|]; constructor; rewrite cmra_unit_l.
* by intros [x|]; constructor; rewrite cmra_unit_idempotent.
* by intros [x|]; constructor; rewrite cmra_unit_idemp.
* intros n mx my; rewrite !option_includedN;intros [->|(x&y&->&->&?)]; auto.
right; exists (unit x), (unit y); eauto using cmra_unit_preservingN.
* intros n [x|] [y|]; rewrite /validN /option_validN /=;
......
......@@ -151,7 +151,7 @@ Proof.
* intros ???? (z&Hy&?&Hxz); destruct Hxz; inversion Hy;clear Hy; setoid_subst;
rewrite ?disjoint_union_difference; auto using closed_up with sts.
eapply closed_up_set; eauto 2 using closed_disjoint with sts.
* intros [] [] []; constructor; rewrite ?(associative _); auto with sts.
* intros [] [] []; constructor; rewrite ?assoc; auto with sts.
* destruct 4; inversion_clear 1; constructor; auto with sts.
* destruct 4; inversion_clear 1; constructor; auto with sts.
* destruct 1; constructor; auto with sts.
......
......@@ -140,7 +140,7 @@ Next Obligation.
intros M P Q x y n1 n2 (x1&x2&Hx&?&?) Hxy ??.
assert ( x2', y {n2} x1 x2' x2 x2') as (x2'&Hy&?).
{ destruct Hxy as [z Hy]; exists (x2 z); split; eauto using cmra_included_l.
apply dist_le with n1; auto. by rewrite (associative op) -Hx Hy. }
apply dist_le with n1; auto. by rewrite (assoc op) -Hx Hy. }
clear Hxy; cofe_subst y; exists x1, x2'; split_ands; [done| |].
* apply uPred_weaken with x1 n1; eauto using cmra_validN_op_l.
* apply uPred_weaken with x2 n1; eauto using cmra_validN_op_r.
......@@ -179,7 +179,7 @@ Program Definition uPred_ownM {M : cmraT} (a : M) : uPred M :=
Next Obligation. by intros M a x1 x2 n [a' ?] Hx; exists a'; rewrite -Hx. Qed.
Next Obligation.
intros M a x1 x n1 n2 [a' Hx1] [x2 Hx] ??.
exists (a' x2). by rewrite (associative op) -(dist_le _ _ _ _ Hx1) // Hx.
exists (a' x2). by rewrite (assoc op) -(dist_le _ _ _ _ Hx1) // Hx.
Qed.
Program Definition uPred_valid {M A : cmraT} (a : A) : uPred M :=
{| uPred_holds n x := {n} a |}.
......@@ -245,11 +245,11 @@ Arguments uPred_holds {_} !_ _ _ /.
Global Instance: PreOrder (@uPred_entails M).
Proof. split. by intros P x i. by intros P Q Q' HP HQ x i ??; apply HQ, HP. Qed.
Global Instance: AntiSymmetric () (@uPred_entails M).
Global Instance: AntiSymm () (@uPred_entails M).
Proof. intros P Q HPQ HQP; split; auto. Qed.
Lemma equiv_spec P Q : P Q P Q Q P.
Proof.
split; [|by intros [??]; apply (anti_symmetric ())].
split; [|by intros [??]; apply (anti_symm ())].
intros HPQ; split; intros x i; apply HPQ.
Qed.
Global Instance entails_proper :
......@@ -434,7 +434,7 @@ Proof. intros; apply const_elim with φ; eauto. Qed.
Lemma const_elim_r φ Q R : (φ Q R) (Q φ) R.
Proof. intros; apply const_elim with φ; eauto. Qed.
Lemma const_equiv (φ : Prop) : φ ( φ : uPred M)%I True%I.
Proof. intros; apply (anti_symmetric _); auto using const_intro. Qed.
Proof. intros; apply (anti_symm _); auto using const_intro. Qed.
Lemma equiv_eq {A : cofeT} P (a b : A) : a b P (a b).
Proof. intros ->; apply eq_refl. Qed.
Lemma eq_sym {A : cofeT} (a b : A) : (a b) (b a).
......@@ -484,57 +484,57 @@ Global Instance exist_mono' A :
Proper (pointwise_relation _ () ==> ()) (@uPred_exist M A).
Proof. intros P1 P2; apply exist_mono. Qed.
Global Instance and_idem : Idempotent () (@uPred_and M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Global Instance or_idem : Idempotent () (@uPred_or M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Global Instance and_comm : Commutative () (@uPred_and M).
Proof. intros P Q; apply (anti_symmetric ()); auto. Qed.
Global Instance and_idem : IdemP () (@uPred_and M).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance or_idem : IdemP () (@uPred_or M).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_comm : Comm () (@uPred_and M).
Proof. intros P Q; apply (anti_symm ()); auto. Qed.
Global Instance True_and : LeftId () True%I (@uPred_and M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_True : RightId () True%I (@uPred_and M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance False_and : LeftAbsorb () False%I (@uPred_and M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_False : RightAbsorb () False%I (@uPred_and M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance True_or : LeftAbsorb () True%I (@uPred_or M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance or_True : RightAbsorb () True%I (@uPred_or M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance False_or : LeftId () False%I (@uPred_or M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance or_False : RightId () False%I (@uPred_or M).
Proof. intros P; apply (anti_symmetric ()); auto. Qed.
Global Instance and_assoc : Associative () (@uPred_and M).
Proof. intros P Q R; apply (anti_symmetric ()); auto. Qed.
Global Instance or_comm : Commutative () (@uPred_or M).
Proof. intros P Q; apply (anti_symmetric ()); auto. Qed.
Global Instance or_assoc : Associative () (@uPred_or M).
Proof. intros P Q R; apply (anti_symmetric ()); auto. Qed.
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_assoc : Assoc () (@uPred_and M).
Proof. intros P Q R; apply (anti_symm ()); auto. Qed.
Global Instance or_comm : Comm () (@uPred_or M).
Proof. intros P Q; apply (anti_symm ()); auto. Qed.
Global Instance or_assoc : Assoc () (@uPred_or M).
Proof. intros P Q R; apply (anti_symm ()); auto. Qed.
Global Instance True_impl : LeftId () True%I (@uPred_impl M).
Proof.
intros P; apply (anti_symmetric ()).
intros P; apply (anti_symm ()).
* by rewrite -(left_id True%I uPred_and (_ _)%I) impl_elim_r.
* by apply impl_intro_l; rewrite left_id.
Qed.
Lemma or_and_l P Q R : (P Q R)%I ((P Q) (P R))%I.
Proof.
apply (anti_symmetric ()); first auto.
apply (anti_symm ()); first auto.
do 2 (apply impl_elim_l', or_elim; apply impl_intro_l); auto.
Qed.
Lemma or_and_r P Q R : (P Q R)%I ((P R) (Q R))%I.
Proof. by rewrite -!(commutative _ R) or_and_l. Qed.
Proof. by rewrite -!(comm _ R) or_and_l. Qed.
Lemma and_or_l P Q R : (P (Q R))%I (P Q P R)%I.
Proof.
apply (anti_symmetric ()); last auto.
apply (anti_symm ()); last auto.
apply impl_elim_r', or_elim; apply impl_intro_l; auto.
Qed.
Lemma and_or_r P Q R : ((P Q) R)%I (P R Q R)%I.
Proof. by rewrite -!(commutative _ R) and_or_l. Qed.
Proof. by rewrite -!(comm _ R) and_or_l. Qed.
Lemma and_exist_l {A} P (Q : A uPred M) : (P a, Q a)%I ( a, P Q a)%I.
Proof.
apply (anti_symmetric ()).
apply (anti_symm ()).
- apply impl_elim_r'. apply exist_elim=>a. apply impl_intro_l.
by rewrite -(exist_intro a).
- apply exist_elim=>a. apply and_intro; first by rewrite and_elim_l.
......@@ -542,8 +542,8 @@ Proof.
Qed.
Lemma and_exist_r {A} P (Q : A uPred M) : (( a, Q a) P)%I ( a, Q a P)%I.
Proof.
rewrite -(commutative _ P) and_exist_l.
apply exist_proper=>a. by rewrite commutative.
rewrite -(comm _ P) and_exist_l.
apply exist_proper=>a. by rewrite comm.
Qed.
(* BI connectives *)
......@@ -558,19 +558,19 @@ Proof.
* intros (x1&x2&?&_&?); cofe_subst; eauto using uPred_weaken, cmra_included_r.
* by intros ?; exists (unit x), x; rewrite cmra_unit_l.
Qed.
Global Instance sep_commutative : Commutative () (@uPred_sep M).
Global Instance sep_comm : Comm () (@uPred_sep M).
Proof.
by intros P Q x n ?; split;
intros (x1&x2&?&?&?); exists x2, x1; rewrite (commutative op).
intros (x1&x2&?&?&?); exists x2, x1; rewrite (comm op).
Qed.
Global Instance sep_associative : Associative () (@uPred_sep M).
Global Instance sep_assoc : Assoc () (@uPred_sep M).
Proof.
intros P Q R x n ?; split.
* intros (x1&x2&Hx&?&y1&y2&Hy&?&?); exists (x1 y1), y2; split_ands; auto.
+ by rewrite -(associative op) -Hy -Hx.
+ by rewrite -(assoc op) -Hy -Hx.
+ by exists x1, y1.
* intros (x1&x2&Hx&(y1&y2&Hy&?&?)&?); exists y1, (y2 x2); split_ands; auto.
+ by rewrite (associative op) -Hy -Hx.
+ by rewrite (assoc op) -Hy -Hx.
+ by exists y2, x2.
Qed.
Lemma wand_intro_r P Q R : (P Q) R P (Q -★ R).
......@@ -595,11 +595,11 @@ Global Instance wand_mono' : Proper (flip (⊑) ==> (⊑) ==> (⊑)) (@uPred_wan
Proof. by intros P P' HP Q Q' HQ; apply wand_mono. Qed.
Global Instance sep_True : RightId () True%I (@uPred_sep M).
Proof. by intros P; rewrite (commutative _) (left_id _ _). Qed.
Proof. by intros P; rewrite comm left_id. Qed.
Lemma sep_elim_l P Q : (P Q) P.
Proof. by rewrite (True_intro Q) (right_id _ _). Qed.
Proof. by rewrite (True_intro Q) right_id. Qed.
Lemma sep_elim_r P Q : (P Q) Q.
Proof. by rewrite (commutative ())%I; apply sep_elim_l. Qed.
Proof. by rewrite (comm ())%I; apply sep_elim_l. Qed.
Lemma sep_elim_l' P Q R : P R (P Q) R.
Proof. intros ->; apply sep_elim_l. Qed.
Lemma sep_elim_r' P Q R : Q R (P Q) R.
......@@ -610,9 +610,9 @@ Proof. by intros; rewrite -(left_id True%I uPred_sep R); apply sep_mono. Qed.
Lemma sep_intro_True_r P Q R : R P True Q R (P Q).
Proof. by intros; rewrite -(right_id True%I uPred_sep R); apply sep_mono. Qed.
Lemma wand_intro_l P Q R : (Q P) R P (Q -★ R).
Proof. rewrite (commutative _); apply wand_intro_r. Qed.
Proof. rewrite comm; apply wand_intro_r. Qed.
Lemma wand_elim_r P Q : (P (P -★ Q)) Q.
Proof. rewrite (commutative _ P); apply wand_elim_l. Qed.
Proof. rewrite (comm _ P); apply wand_elim_l. Qed.
Lemma wand_elim_l' P Q R : P (Q -★ R) (P Q) R.
Proof. intros ->; apply wand_elim_l. Qed.
Lemma wand_elim_r' P Q R : Q (P -★ R) (P Q) R.
......@@ -627,9 +627,9 @@ Lemma const_elim_sep_r φ Q R : (φ → Q ⊑ R) → (Q ★ ■ φ) ⊑ R.
Proof. intros; apply const_elim with φ; eauto. Qed.
Global Instance sep_False : LeftAbsorb () False%I (@uPred_sep M).
Proof. intros P; apply (anti_symmetric _); auto. Qed.
Proof. intros P; apply (anti_symm _); auto. Qed.
Global Instance False_sep : RightAbsorb () False%I (@uPred_sep M).
Proof. intros P; apply (anti_symmetric _); auto. Qed.
Proof. intros P; apply (anti_symm _); auto. Qed.
Lemma sep_and_l P Q R : (P (Q R)) ((P Q) (P R)).
Proof. auto. Qed.
......@@ -637,22 +637,22 @@ Lemma sep_and_r P Q R : ((P ∧ Q) ★ R) ⊑ ((P ★ R) ∧ (Q ★ R)).
Proof. auto. Qed.
Lemma sep_or_l P Q R : (P (Q R))%I ((P Q) (P R))%I.
Proof.
apply (anti_symmetric ()); last by eauto 8.
apply (anti_symm ()); last by eauto 8.
apply wand_elim_r', or_elim; apply wand_intro_l.
- by apply or_intro_l.
- by apply or_intro_r.
Qed.
Lemma sep_or_r P Q R : ((P Q) R)%I ((P R) (Q R))%I.
Proof. by rewrite -!(commutative _ R) sep_or_l. Qed.
Proof. by rewrite -!(comm _ R) sep_or_l. Qed.
Lemma sep_exist_l {A} P (Q : A uPred M) : (P a, Q a)%I ( a, P Q a)%I.
Proof.
intros; apply (anti_symmetric ()).
intros; apply (anti_symm ()).
- apply wand_elim_r', exist_elim=>a. apply wand_intro_l.
by rewrite -(exist_intro a).
- apply exist_elim=> a; apply sep_mono; auto using exist_intro.
Qed.
Lemma sep_exist_r {A} (P: A uPred M) Q: (( a, P a) Q)%I ( a, P a Q)%I.
Proof. setoid_rewrite (commutative _ _ Q); apply sep_exist_l. Qed.
Proof. setoid_rewrite (comm _ _ Q); apply sep_exist_l. Qed.
Lemma sep_forall_l {A} P (Q : A uPred M) : (P a, Q a) ( a, P Q a).
Proof. by apply forall_intro=> a; rewrite forall_elim. Qed.
Lemma sep_forall_r {A} (P : A uPred M) Q : (( a, P a) Q) ( a, P a Q).
......@@ -724,7 +724,7 @@ Qed.
Lemma always_intro P Q : P Q P Q.
Proof.
intros HPQ x n ??; apply HPQ; simpl in *; auto using cmra_unit_validN.
by rewrite cmra_unit_idempotent.
by rewrite cmra_unit_idemp.
Qed.
Lemma always_and P Q : ( (P Q))%I ( P Q)%I.
Proof. done. Qed.
......@@ -741,7 +741,7 @@ Qed.
Lemma always_and_sep_l_1 P Q : ( P Q) ( P Q).
Proof.
intros x n ? [??]; exists (unit x), x; simpl in *.
by rewrite cmra_unit_l cmra_unit_idempotent.
by rewrite cmra_unit_l cmra_unit_idemp.
Qed.
Lemma always_later P : ( P)%I ( P)%I.
Proof. done. Qed.
......@@ -759,26 +759,26 @@ Proof.
Qed.
Lemma always_eq {A:cofeT} (a b : A) : ( (a b))%I (a b : uPred M)%I.
Proof.
apply (anti_symmetric ()); auto using always_elim.
apply (anti_symm ()); auto using always_elim.
apply (eq_rewrite a b (λ b, (a b))%I); auto.
{ intros n; solve_proper. }
rewrite -(eq_refl _ True) always_const; auto.
Qed.
Lemma always_and_sep P Q : ( (P Q))%I ( (P Q))%I.
Proof. apply (anti_symmetric ()); auto using always_and_sep_1. Qed.
Proof. apply (anti_symm ()); auto using always_and_sep_1. Qed.
Lemma always_and_sep_l P Q : ( P Q)%I ( P Q)%I.
Proof. apply (anti_symmetric ()); auto using always_and_sep_l_1. Qed.
Proof. apply (anti_symm ()); auto using always_and_sep_l_1. Qed.
Lemma always_and_sep_r P Q : (P Q)%I (P Q)%I.
Proof. by rewrite !(commutative _ P) always_and_sep_l. Qed.
Proof. by rewrite !(comm _ P) always_and_sep_l. Qed.
Lemma always_sep P Q : ( (P Q))%I ( P Q)%I.
Proof. by rewrite -always_and_sep -always_and_sep_l always_and. Qed.
Lemma always_wand P Q : (P -★ Q) ( P -★ Q).
Proof. by apply wand_intro_r; rewrite -always_sep wand_elim_l. Qed.
Lemma always_sep_dup P : ( P)%I ( P P)%I.
Proof. by rewrite -always_sep -always_and_sep (idempotent _). Qed.
Proof. by rewrite -always_sep -always_and_sep (idemp _). Qed.
Lemma always_wand_impl P Q : ( (P -★ Q))%I ( (P Q))%I.
Proof.
apply (anti_symmetric ()); [|by rewrite -impl_wand].
apply (anti_symm ()); [|by rewrite -impl_wand].
apply always_intro, impl_intro_r.
by rewrite always_and_sep_l always_elim wand_elim_l.
Qed.
......@@ -792,16 +792,16 @@ Lemma ownM_op (a1 a2 : M) :
uPred_ownM (a1 a2) (uPred_ownM a1 uPred_ownM a2)%I.
Proof.
intros x n ?; split.
* intros [z ?]; exists a1, (a2 z); split; [by rewrite (associative op)|].
* intros [z ?]; exists a1, (a2 z); split; [by rewrite (assoc op)|].
split. by exists (unit a1); rewrite cmra_unit_r. by exists z.
* intros (y1&y2&Hx&[z1 Hy1]&[z2 Hy2]); exists (z1 z2).
by rewrite (associative op _ z1) -(commutative op z1) (associative op z1)
-(associative op _ a2) (commutative op z1) -Hy1 -Hy2.
by rewrite (assoc op _ z1) -(comm op z1) (assoc op z1)
-(assoc op _ a2) (comm op z1) -Hy1 -Hy2.
Qed.
Lemma always_ownM_unit (a : M) : ( uPred_ownM (unit a))%I uPred_ownM (unit a).
Proof.
intros x n; split; [by apply always_elim|intros [a' Hx]]; simpl.
rewrite -(cmra_unit_idempotent a) Hx.
rewrite -(cmra_unit_idemp a) Hx.
apply cmra_unit_preservingN, cmra_includedN_l.
Qed.
Lemma always_ownM (a : M) : unit a a ( uPred_ownM a)%I uPred_ownM a.
......@@ -809,7 +809,7 @@ Proof. by intros <-; rewrite always_ownM_unit. Qed.
Lemma ownM_something : True a, uPred_ownM a.
Proof. intros x n ??. by exists x; simpl. Qed.
Lemma ownM_empty `{Empty M, !CMRAIdentity M} : True uPred_ownM ∅.
Proof. intros x n ??; by exists x; rewrite (left_id _ _). Qed.
Proof. intros x n ??; by exists x; rewrite left_id. Qed.
Lemma ownM_valid (a : M) : uPred_ownM a a.
Proof. intros x n Hv [a' ?]; cofe_subst; eauto using cmra_validN_op_l. Qed.
Lemma valid_intro {A : cmraT} (a : A) : a True a.
......@@ -835,20 +835,20 @@ Global Instance big_and_perm : Proper ((≡ₚ) ==> (≡)) (@uPred_big_and M).
Proof.
induction 1 as [|P Ps Qs ? IH|P Q Ps|]; simpl; auto.
* by rewrite IH.
* by rewrite !associative (commutative _ P).
* by rewrite !assoc (comm _ P).
* etransitivity; eauto.
Qed.
Global Instance big_sep_perm : Proper (() ==> ()) (@uPred_big_sep M).
Proof.
induction 1 as [|P Ps Qs ? IH|P Q Ps|]; simpl; auto.
* by rewrite IH.
* by rewrite !associative (commutative _ P).
* by rewrite !assoc (comm _ P).
* etransitivity; eauto.
Qed.
Lemma big_and_app Ps Qs : (Π (Ps ++ Qs))%I (Π Ps Π Qs)%I.
Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?associative ?IH. Qed.
Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?assoc ?IH. Qed.
Lemma big_sep_app Ps Qs : (Π (Ps ++ Qs))%I (Π Ps Π Qs)%I.
Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?associative ?IH. Qed.
Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?assoc ?IH. Qed.
Lemma big_sep_and Ps : (Π Ps) (Π Ps).
Proof. by induction Ps as [|P Ps IH]; simpl; auto. Qed.
Lemma big_and_elem_of Ps P : P Ps (Π Ps) P.
......@@ -883,7 +883,7 @@ Proof.
intros; rewrite /TimelessP later_sep (timelessP P) (timelessP Q).
apply wand_elim_l', or_elim; apply wand_intro_r; auto.
apply wand_elim_r', or_elim; apply wand_intro_r; auto.
rewrite ?(commutative _ Q); auto.
rewrite ?(comm _ Q); auto.
Qed.
Global Instance wand_timeless P Q : TimelessP Q TimelessP (P -★ Q).
Proof.
......@@ -963,7 +963,7 @@ Proof. unfold ASL=> ?; revert ys; induction xs=> -[|??]; constructor; auto. Qed.
(* Derived lemmas for always stable *)
Lemma always_always P `{!AlwaysStable P} : ( P)%I P.
Proof. apply (anti_symmetric ()); auto using always_elim. Qed.
Proof. apply (anti_symm ()); auto using always_elim. Qed.
Lemma always_intro' P Q `{!AlwaysStable P} : P Q P Q.
Proof. rewrite -(always_always P); apply always_intro. Qed.
Lemma always_and_sep_l' P Q `{!AlwaysStable P} : (P Q)%I (P Q)%I.
......
......@@ -242,13 +242,13 @@ Proof.
revert v; induction e; intros; simplify_option_equality; auto with f_equal.
Qed.
Instance: Injective (=) (=) of_val.
Proof. by intros ?? Hv; apply (injective Some); rewrite -!to_of_val Hv. Qed.
Instance: Inj (=) (=) of_val.
Proof. by intros ?? Hv; apply (inj Some); rewrite -!to_of_val Hv. Qed.
Instance fill_item_inj Ki : Injective (=) (=) (fill_item Ki).
Instance fill_item_inj Ki : Inj (=) (=) (fill_item Ki).
Proof. destruct Ki; intros ???; simplify_equality'; auto with f_equal. Qed.
Instance ectx_fill_inj K : Injective (=) (=) (fill K).
Instance ectx_fill_inj K : Inj (=) (=) (fill K).
Proof. red; induction K as [|Ki K IH]; naive_solver. Qed.
Lemma fill_app K1 K2 e : fill (K1 ++ K2) e = fill K1 (fill K2 e).
......@@ -348,7 +348,7 @@ Proof.
* intros e1 σ1 e2 σ2 ? Hnval [K'' e1'' e2'' Heq1 -> Hstep].
destruct (heap_lang.step_by_val
K K'' e1 e1'' σ1 e2'' σ2 ef) as [K' ->]; eauto.
rewrite heap_lang.fill_app in Heq1; apply (injective _) in Heq1.
rewrite heap_lang.fill_app in Heq1; apply (inj _) in Heq1.
exists (heap_lang.fill K' e2''); rewrite heap_lang.fill_app; split; auto.
econstructor; eauto.
Qed.
......
......@@ -99,7 +99,7 @@ Proof.
apply sep_mono, later_mono; first done.
apply forall_intro=>e2; apply forall_intro=>σ2; apply forall_intro=>ef.
apply wand_intro_l.
rewrite always_and_sep_l' -associative -always_and_sep_l'.
rewrite always_and_sep_l' -assoc -always_and_sep_l'.
apply const_elim_l=>-[l [-> [-> [-> ?]]]].
by rewrite (forall_elim l) right_id const_equiv // left_id wand_elim_r.
Qed.
......
......@@ -68,7 +68,7 @@ Module LiftingTests.
rewrite /FindPred.
rewrite -(wp_bindi (AppLCtx _)) -wp_let //=.
revert n1. apply löb_all_1=>n1.
rewrite (commutative uPred_and ( _)%I) associative; apply const_elim_r=>?.
rewrite (comm uPred_and ( _)%I) assoc; apply const_elim_r=>?.
rewrite -wp_value' //.
rewrite -wp_rec' // =>-/=.
(* FIXME: ssr rewrite fails with "Error: _pattern_value_ is used in conclusion." *)
......
......@@ -514,71 +514,71 @@ Arguments insertE _ _ _ _ _ _ !_ _ !_ / : simpl nomatch.
(** ** Common properties *)
(** These operational type classes allow us to refer to common mathematical
properties in a generic way. For example, for injectivity of [(k ++)] it
allows us to write [injective (k ++)] instead of [app_inv_head k]. *)
Class Injective {A B} (R : relation A) (S : relation B) (f : A B) : Prop :=
injective: x y, S (f x) (f y) R x y.
Class Injective2 {A B C} (R1 : relation A) (R2 : relation B)
allows us to write [inj (k ++)] instead of [app_inv_head k]. *)
Class Inj {A B} (R : relation A) (S : relation B) (f : A B) : Prop :=
inj x y : S (f x) (f y) R x y.
Class Inj2 {A B C} (R1 : relation A) (R2 : relation B)
(S : relation C) (f : A B C) : Prop :=
injective2: x1 x2 y1 y2, S (f x1 x2) (f y1 y2) R1 x1 y1 R2 x2 y2.
inj2 x1 x2 y1 y2 : S (f x1 x2) (f y1 y2) R1 x1 y1 R2 x2 y2.
Class Cancel {A B} (S : relation B) (f : A B) (g : B A) : Prop :=
cancel: x, S (f (g x)) x.
Class Surjective {A B} (R : relation B) (f : A B) :=
surjective : y, x, R (f x) y.
Class Idempotent {A} (R : relation A) (f : A A A) : Prop :=
idempotent: x, R (f x x) x.
Class Commutative {A B} (R : relation A) (f : B B A) : Prop :=
commutative: x y, R (f x y) (f y x).
cancel : x, S (f (g x)) x.
Class Surj {A B} (R : relation B) (f : A B) :=
surj y : x, R (f x) y.
Class IdemP {A} (R : relation A) (f : A A A) : Prop :=
idemp x : R (f x x) x.
Class Comm {A B} (R : relation A) (f : B B A) : Prop :=
comm x y : R (f x y) (f y x).
Class LeftId {A} (R : relation A) (i : A) (f : A A A) : Prop :=
left_id: x, R (f i x) x.
left_id x : R (f i x) x.
Class RightId {A} (R : relation A) (i : A) (f : A A A) : Prop :=
right_id: x, R (f x i) x.
Class Associative {A} (R : relation A) (f : A A A) : Prop :=
associative: x y z, R (f x (f y z)) (f (f x y) z).
right_id x : R (f x i) x.
Class Assoc {A} (R : relation A) (f : A A A) : Prop :=
assoc x y z : R (f x (f y z)) (f (f x y) z).
Class LeftAbsorb {A} (R : relation A) (i : A) (f : A A A) : Prop :=
left_absorb: x, R (f i x) i.
left_absorb x : R (f i x) i.
Class RightAbsorb {A} (R : relation A) (i : A) (f : A A A) : Prop :=
right_absorb: x, R (f x i) i.
Class AntiSymmetric {A} (R S : relation A) : Prop :=
anti_symmetric: x y, S x y S y x R x y.
right_absorb x : R (f x i) i.
Class AntiSymm {A} (R S : relation A) : Prop :=
anti_symm x y : S x y S y x R x y.
Class Total {A} (R : relation A) := total x y : R x y R y x.
Class Trichotomy {A} (R : relation A) :=
trichotomy : x y, R x y x = y R y x.
trichotomy x y : R x y x = y R y x.
Class TrichotomyT {A} (R : relation A) :=
trichotomyT : x y, {R x y} + {x = y} + {R y x}.
trichotomyT x y : {R x y} + {x = y} + {R y x}.
Arguments irreflexivity {_} _ {_} _ _.
Arguments injective {_ _ _ _} _ {_} _ _ _.
Arguments injective2 {_ _ _ _ _ _} _ {_} _ _ _ _ _.
Arguments inj {_ _ _ _} _ {_} _ _ _.
Arguments inj2 {_ _ _ _ _ _} _ {_} _ _ _ _ _.
Arguments cancel {_ _ _} _ _ {_} _.
Arguments surjective {_ _ _} _ {_} _.
Arguments idempotent {_ _} _ {_} _.
Arguments commutative {_ _ _} _ {_} _ _.
Arguments surj {_ _ _} _ {_} _.
Arguments idemp {_ _} _ {_} _.
Arguments comm {_ _ _} _ {_} _ _.
Arguments left_id {_ _} _ _ {_} _.
Arguments right_id {_ _} _ _ {_} _.
Arguments associative {_ _} _ {_} _ _ _.
Arguments assoc {_ _} _ {_} _ _ _.
Arguments left_absorb {_ _} _ _ {_} _.
Arguments right_absorb {_ _} _ _ {_} _.
Arguments anti_symmetric {_ _} _ {_} _ _ _ _.
Arguments anti_symm {_ _} _ {_} _ _ _ _.
Arguments total {_} _ {_} _ _.
Arguments trichotomy {_} _ {_} _ _.
Arguments trichotomyT {_} _ {_} _ _.
Instance id_injective {A} : Injective (=) (=) (@id A).
Instance id_inj {A} : Inj (=) (=) (@id A).
Proof. intros ??; auto. Qed.
(** The following lemmas are specific versions of the projections of the above
type classes for Leibniz equality. These lemmas allow us to enforce Coq not to
use the setoid rewriting mechanism. *)
Lemma idempotent_L {A} (f : A A A) `{!Idempotent (=) f} x : f x x = x.
Lemma idemp_L {A} (f : A A A) `{!IdemP (=) f} x : f x x = x.
Proof. auto. Qed.
Lemma commutative_L {A B} (f : B B A) `{!Commutative (=) f} x y :
Lemma comm_L {A B} (f : B B A) `{!Comm (=) f} x y :
f x y = f y x.
Proof. auto. Qed.
Lemma left_id_L {A} (i : A) (f : A A A) `{!LeftId (=) i f} x : f i x = x.
Proof. auto. Qed.
Lemma right_id_L {A} (i : A) (f : A A A) `{!RightId (=) i f} x : f x i = x.
Proof. auto. Qed.
Lemma associative_L {A} (f : A A A) `{!Associative (=) f} x y z :
Lemma assoc_L {A} (f : A A A) `{!Assoc (=) f} x y z :
f x (f y z) = f (f x y) z.
Proof. auto. Qed.
Lemma left_absorb_L {A} (i : A) (f : A A A) `{!LeftAbsorb (=) i f} x :
......@@ -593,7 +593,7 @@ Proof. auto. Qed.
relation [R] instead of [⊆] to support multiple orders on the same type. *)
Class PartialOrder {A} (R : relation A) : Prop := {
partial_order_pre :> PreOrder R;
partial_order_anti_symmetric :> AntiSymmetric (=) R
partial_order_anti_symm :> AntiSymm (=) R
}.
Class TotalOrder {A} (R : relation A) : Prop := {
total_order_partial :> PartialOrder R;
......@@ -746,31 +746,17 @@ Proof. intuition. Qed.
Lemma symmetry_iff `(R : relation A) `{!Symmetric R} x y : R x y R y x.
Proof. intuition. Qed.
(** ** Pointwise relations *)
(** These instances are in Coq trunk since revision 15455, but are not in Coq
8.4 yet. *)
Instance pointwise_reflexive {A} `{R : relation B} :
Reflexive R Reflexive (pointwise_relation A R) | 9.
Proof. firstorder. Qed.
Instance pointwise_symmetric {A} `{R : relation B} :
Symmetric R Symmetric (pointwise_relation A R) | 9.
Proof. firstorder. Qed.
Instance pointwise_transitive {A} `{R : relation B} :
Transitive R Transitive (pointwise_relation A R) | 9.
Proof. firstorder. Qed.
(** ** Unit *)
Instance unit_equiv : Equiv unit := λ _ _, True.
Instance unit_equivalence : Equivalence (@equiv unit _).
Proof. repeat split. Qed.
(** ** Products *)
Instance prod_map_injective {A A' B B'} (f : A A') (g : B B') :
Injective (=) (=) f Injective (=) (=) g
Injective (=) (=) (prod_map f g).
Instance prod_map_inj {A A' B B'} (f : A A') (g : B B') :
Inj (=) (=) f Inj (=) (=) g Inj (=) (=) (prod_map f g).
Proof.
intros ?? [??] [??] ?; simpl in *; f_equal;
[apply (injective f)|apply (injective g)]; congruence.
[apply (inj f)|apply (inj g)]; congruence.
Qed.
Definition prod_relation {A B} (R1 : relation A) (R2 : relation B) :
......@@ -815,17 +801,17 @@ Lemma and_wlog_l (P Q : Prop) : (Q → P) → Q → (P ∧ Q).
Proof. tauto. Qed.
Lemma and_wlog_r (P Q : Prop) : P (P Q) (P Q).
Proof. tauto. Qed.
Instance: A B (x : B), Commutative (=) (λ _ _ : A, x).
Instance: A B (x : B), Comm (=) (λ _ _ : A, x).
Proof. red. trivial. Qed.
Instance: A (x : A), Associative (=) (λ _ _ : A, x).
Instance: A (x : A), Assoc (=) (λ _ _ : A, x).
Proof. red. trivial. Qed.
Instance: A, Associative (=) (λ x _ : A, x).
Instance: A, Assoc (=) (λ x _ : A, x).
Proof. red. trivial. Qed.
Instance: A, Associative (=) (λ _ x : A, x).
Instance: A, Assoc (=) (λ _ x : A, x).
Proof. red. trivial. Qed.
Instance: A, Idempotent (=) (λ x _ : A, x).
Instance: A, IdemP (=) (λ x _ : A, x).
Proof. red. trivial. Qed.
Instance: A, Idempotent (=) (λ _ x : A, x).
Instance: A, IdemP (=) (λ _ x : A, x).
Proof. red. trivial. Qed.
Instance left_id_propholds {A} (R : relation A) i f :
......@@ -841,7 +827,7 @@ Instance right_absorb_propholds {A} (R : relation A) i f :
RightAbsorb R i f x, PropHolds (R (f x i) i).
Proof. red. trivial. Qed.
Instance idem_propholds {A} (R : relation A) f :
Idempotent R f x, PropHolds (R (f x x) x).
IdemP R f x, PropHolds (R (f x x) x).
Proof. red. trivial. Qed.
Instance: `{R1 : relation A, R2 : relation B} (x : B),
......@@ -849,47 +835,47 @@ Instance: ∀ `{R1 : relation A, R2 : relation B} (x : B),
Proof. intros A R1 B R2 x ? y1 y2; reflexivity. Qed.
Instance: @PreOrder A (=).
Proof. split; repeat intro; congruence. Qed.
Lemma injective_iff {A B} {R : relation A} {S : relation B} (f : A B)
`{!Injective R S f} `{!Proper (R ==> S) f} x y : S (f x) (f y) R x y.
Lemma inj_iff {A B} {R : relation A} {S : relation B} (f : A B)
`{!Inj R S f} `{!Proper (R ==> S) f} x y : S (f x) (f y) R x y.
Proof. firstorder. Qed.
Instance: Injective (=) (=) (@inl A B).
Instance: Inj (=) (=) (@inl A B).
Proof. injection 1; auto. Qed.
Instance: Injective (=) (=) (@inr A B).
Instance: Inj (=) (=) (@inr A B).
Proof. injection 1; auto. Qed.
Instance: Injective2 (=) (=) (=) (@pair A B).
Instance: Inj2 (=) (=) (=) (@pair A B).
Proof. injection 1; auto. Qed.
Instance: `{Injective2 A B C R1 R2 R3 f} y, Injective R1 R3 (λ x, f x y).
Proof. repeat intro; edestruct (injective2 f); eauto. Qed.
Instance: `{Injective2 A B C R1 R2 R3 f} x, Injective R2 R3 (f x).
Proof. repeat intro; edestruct (injective2 f); eauto. Qed.
Instance: `{Inj2 A B C R1 R2 R3 f} y, Inj R1 R3 (λ x, f x y).
Proof. repeat intro; edestruct (inj2 f); eauto. Qed.
Instance: `{Inj2 A B C R1 R2 R3 f} x, Inj R2 R3 (f x).
Proof. repeat intro; edestruct (inj2 f); eauto. Qed.
Lemma cancel_injective `{Cancel A B R1 f g}
`{!Equivalence R1} `{!Proper (R2 ==> R1) f} : Injective R1 R2 g.
Lemma cancel_inj `{Cancel A B R1 f g}
`{!Equivalence R1} `{!Proper (R2 ==> R1) f} : Inj R1 R2 g.
Proof.
intros x y E. rewrite <-(cancel f g x), <-(cancel f g y), E. reflexivity.
Qed.
Lemma cancel_surjective `{Cancel A B R1 f g} : Surjective R1 f.
Lemma cancel_surj `{Cancel A B R1 f g} : Surj R1 f.
Proof. intros y. exists (g y). auto. Qed.
Lemma impl_transitive (P Q R : Prop) : (P Q) (Q R) (P R).
Proof. tauto. Qed.
Instance: Commutative () (@eq A).
Instance: Comm () (@eq A).
Proof. red; intuition. Qed.
Instance: Commutative () (λ x y, @eq A y x).
Instance: Comm () (λ x y, @eq A y x).
Proof. red; intuition. Qed.
Instance: Commutative () ().
Instance: Comm () ().
Proof. red; intuition. Qed.
Instance: Commutative () ().
Instance: Comm () ().
Proof. red; intuition. Qed.
Instance: Associative () ().
Instance: Assoc () ().
Proof. red; intuition. Qed.
Instance: Idempotent () ().
Instance: IdemP () ().
Proof. red; intuition. Qed.
Instance: Commutative () ().
Instance: Comm () ().
Proof. red; intuition. Qed.
Instance: Associative () ().
Instance: Assoc () ().
Proof. red; intuition. Qed.
Instance: Idempotent () ().
Instance: IdemP () ().
Proof. red; intuition. Qed.
Instance: LeftId () True ().
Proof. red; intuition. Qed.
......@@ -911,26 +897,26 @@ Instance: LeftId (↔) True impl.
Proof. unfold impl. red; intuition. Qed.
Instance: RightAbsorb () True impl.
Proof. unfold impl. red; intuition. Qed.
Lemma not_injective `{Injective A B R R' f} x y : ¬R x y ¬R' (f x) (f y).
Lemma not_inj `{Inj A B R R' f} x y : ¬R x y ¬R' (f x) (f y).
Proof. intuition. Qed.
Instance injective_compose {A B C} R1 R2 R3 (f : A B) (g : B C) :
Injective R1 R2 f Injective R2 R3 g Injective R1 R3 (g f).
Instance inj_compose {A B C} R1 R2 R3 (f : A B) (g : B C) :
Inj R1 R2 f Inj R2 R3 g Inj R1 R3 (g f).
Proof. red; intuition. Qed.
Instance surjective_compose {A B C} R (f : A B) (g : B C) :
Surjective (=) f Surjective R g Surjective R (g f).
Instance surj_compose {A B C} R (f : A B) (g : B C) :
Surj (=) f Surj R g Surj R (g f).
Proof.
intros ?? x. unfold compose. destruct (surjective g x) as [y ?].
destruct (surjective f y) as [z ?]. exists z. congruence.
intros ?? x. unfold compose. destruct (surj g x) as [y ?].
destruct (surj f y) as [z ?]. exists z. congruence.
Qed.
Section sig_map.
Context `{P : A Prop} `{Q : B Prop} (f : A B) (Hf : x, P x Q (f x)).
Definition sig_map (x : sig P) : sig Q := f (`x) Hf _ (proj2_sig x).
Global Instance sig_map_injective:
( x, ProofIrrel (P x)) Injective (=) (=) f Injective (=) (=) sig_map.
Global Instance sig_map_inj:
( x, ProofIrrel (P x)) Inj (=) (=) f Inj (=) (=) sig_map.
Proof.
intros ?? [x Hx] [y Hy]. injection 1. intros Hxy.
apply (injective f) in Hxy; subst. rewrite (proof_irrel _ Hy). auto.
apply (inj f) in Hxy; subst. rewrite (proof_irrel _ Hy). auto.
Qed.
End sig_map.
Arguments sig_map _ _ _ _ _ _ !_ /.
......@@ -15,13 +15,13 @@ Definition encode_nat `{Countable A} (x : A) : nat :=
pred (Pos.to_nat (encode x)).
Definition decode_nat `{Countable A} (i : nat) : option A :=
decode (Pos.of_nat (S i)).
Instance encode_injective `{Countable A} : Injective (=) (=) encode.
Instance encode_inj `{Countable A} : Inj (=) (=) encode.
Proof.
intros x y Hxy; apply (injective Some).
intros x y Hxy; apply (inj Some).
by rewrite <-(decode_encode x), Hxy, decode_encode.
Qed.
Instance encode_nat_injective `{Countable A} : Injective (=) (=) encode_nat.
Proof. unfold encode_nat; intros x y Hxy; apply (injective encode); lia. Qed.
Instance encode_nat_inj `{Countable A} : Inj (=) (=) encode_nat.
Proof. unfold encode_nat; intros x y Hxy; apply (inj encode); lia. Qed.
Lemma decode_encode_nat `{Countable A} x : decode_nat (encode_nat x) = Some x.
Proof.
pose proof (Pos2Nat.is_pos (encode x)).
......@@ -70,11 +70,11 @@ Section choice.
Definition choice (HA : x, P x) : { x | P x } := _choose_correct HA.
End choice.
Lemma surjective_cancel `{Countable A} `{ x y : B, Decision (x = y)}
(f : A B) `{!Surjective (=) f} : { g : B A & Cancel (=) f g }.
Lemma surj_cancel `{Countable A} `{ x y : B, Decision (x = y)}
(f : A B) `{!Surj (=) f} : { g : B A & Cancel (=) f g }.
Proof.
exists (λ y, choose (λ x, f x = y) (surjective f y)).
intros y. by rewrite (choose_correct (λ x, f x = y) (surjective f y)).
exists (λ y, choose (λ x, f x = y) (surj f y)).
intros y. by rewrite (choose_correct (λ x, f x = y) (surj f y)).
Qed.
(** * Instances *)
......@@ -197,7 +197,7 @@ Lemma list_encode_app' `{Countable A} (l1 l2 : list A) acc :
Proof.
revert acc; induction l1; simpl; auto.
induction l2 as [|x l IH]; intros acc; simpl; [by rewrite ?(left_id_L _ _)|].
by rewrite !(IH (Nat.iter _ _ _)), (associative_L _), x0_iter_x1.
by rewrite !(IH (Nat.iter _ _ _)), (assoc_L _), x0_iter_x1.
Qed.
Program Instance list_countable `{Countable A} : Countable (list A) :=
{| encode := list_encode 1; decode := list_decode [] 0 |}.
......@@ -211,7 +211,7 @@ Next Obligation.
{ by intros help l; rewrite help, (right_id_L _ _). }
induction l as [|x l IH] using @rev_ind; intros acc; [done|].
rewrite list_encode_app'; simpl; rewrite <-x0_iter_x1, decode_iter; simpl.
by rewrite decode_encode_nat; simpl; rewrite IH, <-(associative_L _).
by rewrite decode_encode_nat; simpl; rewrite IH, <-(assoc_L _).
Qed.
Lemma list_encode_app `{Countable A} (l1 l2 : list A) :
encode (l1 ++ l2)%list = encode l1 ++ encode l2.
......
......@@ -12,7 +12,7 @@ Proof. firstorder. Qed.
Lemma Is_true_reflect (b : bool) : reflect b b.
Proof. destruct b. by left. right. intros []. Qed.
Instance: Injective (=) () Is_true.
Instance: Inj (=) () Is_true.
Proof. intros [] []; simpl; intuition. Qed.
(** We introduce [decide_rel] to avoid inefficienct computation due to eager
......
......@@ -47,7 +47,7 @@ Lemma error_fmap_bind {S E A B C} (f : A → B) (g : B → error S E C) x s :
((f <$> x) ≫= g) s = (x ≫= g f) s.
Proof. by compute; destruct (x s) as [|[??]]. Qed.
Lemma error_associative {S E A B C} (f : A error S E B) (g : B error S E C) x s :
Lemma error_assoc {S E A B C} (f : A error S E B) (g : B error S E C) x s :
((x ≫= f) ≫= g) s = (a x; f a ≫= g) s.
Proof. by compute; destruct (x s) as [|[??]]. Qed.
Lemma error_of_option_bind {S E A B} (f : A option B) o e :
......@@ -114,7 +114,7 @@ Tactic Notation "error_proceed" :=
| H : (gets _ ≫= _) _ = _ |- _ => rewrite error_left_gets in H
| H : (modify _ ≫= _) _ = _ |- _ => rewrite error_left_modify in H
| H : ((_ <$> _) ≫= _) _ = _ |- _ => rewrite error_fmap_bind in H
| H : ((_ ≫= _) ≫= _) _ = _ |- _ => rewrite error_associative in H
| H : ((_ ≫= _) ≫= _) _ = _ |- _ => rewrite error_assoc in H
| H : (error_guard _ _ _) _ = _ |- _ =>
let H' := fresh in apply error_guard_ret in H; destruct H as [H' H]
| _ => progress simplify_equality'
......
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