From 1cfbdb17b4bcc57404740b4582753dfbed77b758 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Thu, 4 Feb 2016 15:05:45 +0100 Subject: [PATCH] More stuff for indexed products. * Insert and singleton operation. * Identity element. * Non-expansiveness and properness of insert and singleton. * Frame preserving updates. * Functoriality. --- algebra/cmra.v | 80 ++++++++++++++++++++++++++++++++++++++++------ algebra/cofe.v | 50 +++++++++++++++++++++++++++++ algebra/fin_maps.v | 2 +- 3 files changed, 122 insertions(+), 10 deletions(-) diff --git a/algebra/cmra.v b/algebra/cmra.v index 8e1bbdff9..6296a2714 100644 --- a/algebra/cmra.v +++ b/algebra/cmra.v @@ -531,10 +531,15 @@ Instance prodRA_map_ne {A A' B B'} n : (** ** Indexed product *) Section iprod_cmra. Context {A} {B : A → cmraT}. + Implicit Types f g : iprod B. Instance iprod_op : Op (iprod B) := λ f g x, f x ⋅ g x. + Definition iprod_lookup_op f g x : (f ⋅ g) x = f x ⋅ g x := eq_refl. Instance iprod_unit : Unit (iprod B) := λ f x, unit (f x). + Definition iprod_lookup_unit f x : (unit f) x = unit (f x) := eq_refl. + Global Instance iprod_empty `{∀ x, Empty (B x)} : Empty (iprod B) := λ x, ∅. Instance iprod_validN : ValidN (iprod B) := λ n f, ∀ x, ✓{n} (f x). Instance iprod_minus : Minus (iprod B) := λ f g x, f x ⩪ g x. + Definition iprod_lookup_minus f g x : (f ⩪ g) x = f x ⩪ g x := eq_refl. Lemma iprod_includedN_spec (f g : iprod B) n : f ≼{n} g ↔ ∀ x, f x ≼{n} g x. Proof. split. @@ -545,21 +550,21 @@ Section iprod_cmra. Definition iprod_cmra_mixin : CMRAMixin (iprod B). Proof. split. - * by intros n f1 f2 f3 Hf x; rewrite /op /iprod_op (Hf x). - * by intros n f1 f2 Hf x; rewrite /unit /iprod_unit (Hf x). + * by intros n f1 f2 f3 Hf x; rewrite iprod_lookup_op (Hf x). + * by intros n f1 f2 Hf x; rewrite iprod_lookup_unit (Hf x). * by intros n f1 f2 Hf ? x; rewrite -(Hf x). - * by intros n f f' Hf g g' Hg i; rewrite /minus /iprod_minus (Hf i) (Hg i). + * by intros n f f' Hf g g' Hg i; rewrite iprod_lookup_minus (Hf i) (Hg i). * by intros f x. * intros n f Hf x; apply cmra_validN_S, Hf. - * by intros f1 f2 f3 x; rewrite /op /iprod_op associative. - * by intros f1 f2 x; rewrite /op /iprod_op commutative. - * by intros f x; rewrite /op /iprod_op /unit /iprod_unit cmra_unit_l. - * by intros f x; rewrite /unit /iprod_unit cmra_unit_idempotent. + * by intros f1 f2 f3 x; rewrite iprod_lookup_op associative. + * by intros f1 f2 x; rewrite iprod_lookup_op commutative. + * 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. * intros n f1 f2; rewrite !iprod_includedN_spec=> Hf x. - by rewrite /unit /iprod_unit; apply cmra_unit_preservingN, Hf. + 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. * intros n f1 f2; rewrite iprod_includedN_spec=> Hf x. - by rewrite /op /iprod_op /minus /iprod_minus cmra_op_minus; try apply Hf. + by rewrite iprod_lookup_op iprod_lookup_minus cmra_op_minus; try apply Hf. Qed. Definition iprod_cmra_extend_mixin : CMRAExtendMixin (iprod B). Proof. @@ -570,6 +575,63 @@ Section iprod_cmra. Qed. Canonical Structure iprodRA : cmraT := CMRAT iprod_cofe_mixin iprod_cmra_mixin iprod_cmra_extend_mixin. + Global Instance iprod_cmra_identity `{∀ x, Empty (B x)} : + (∀ x, CMRAIdentity (B x)) → CMRAIdentity iprodRA. + Proof. + intros ?; split. + * intros n x; apply cmra_empty_valid. + * by intros f x; rewrite iprod_lookup_op left_id. + * by intros f Hf x; apply (timeless _). + Qed. + + Context `{∀ x x' : A, Decision (x = x')}. + Lemma iprod_insert_updateP x (P : B x → Prop) (Q : iprod B → Prop) g y1 : + y1 ~~>: P → (∀ y2, P y2 → Q (iprod_insert x y2 g)) → + iprod_insert x y1 g ~~>: Q. + Proof. + intros Hy1 HP gf n Hg. destruct (Hy1 (gf x) n) as (y2&?&?). + { move: (Hg x). by rewrite iprod_lookup_op iprod_lookup_insert. } + exists (iprod_insert x y2 g); split; [auto|]. + intros x'; destruct (decide (x' = x)) as [->|]; + rewrite iprod_lookup_op ?iprod_lookup_insert //. + move: (Hg x'). by rewrite iprod_lookup_op !iprod_lookup_insert_ne. + Qed. + Lemma iprod_insert_updateP' x (P : B x → Prop) g y1 : + y1 ~~>: P → + iprod_insert x y1 g ~~>: λ g', ∃ y2, g' = iprod_insert x y2 g ∧ P y2. + Proof. eauto using iprod_insert_updateP. Qed. + Lemma iprod_insert_update g x y1 y2 : + y1 ~~> y2 → iprod_insert x y1 g ~~> iprod_insert x y2 g. + Proof. + rewrite !cmra_update_updateP; + eauto using iprod_insert_updateP with congruence. + Qed. + + Context `{∀ x, Empty (B x)}. + Lemma iprod_singleton_updateP x (P : B x → Prop) (Q : iprod B → Prop) y1 : + y1 ~~>: P → (∀ y2, P y2 → Q (iprod_singleton x y2)) → + iprod_singleton x y1 ~~>: Q. + Proof. rewrite /iprod_singleton; eauto using iprod_insert_updateP. Qed. + Lemma iprod_singleton_updateP' x (P : B x → Prop) y1 : + y1 ~~>: P → + iprod_singleton x y1 ~~>: λ g', ∃ y2, g' = iprod_singleton x y2 ∧ P y2. + Proof. eauto using iprod_singleton_updateP. Qed. + Lemma iprod_singleton_update x y1 y2 : + y1 ~~> y2 → iprod_singleton x y1 ~~> iprod_singleton x y2. + Proof. by intros; apply iprod_insert_update. Qed. End iprod_cmra. Arguments iprodRA {_} _. + +Instance iprod_map_cmra_monotone {A} {B1 B2: A → cmraT} (f : ∀ x, B1 x → B2 x) : + (∀ x, CMRAMonotone (f x)) → CMRAMonotone (iprod_map f). +Proof. + split. + * intros n g1 g2; rewrite !iprod_includedN_spec=> Hf x. + rewrite /iprod_map; apply includedN_preserving, Hf. + * intros n g Hg x; rewrite /iprod_map; apply validN_preserving, Hg. +Qed. +Definition iprodRA_map {A} {B1 B2: A → cmraT} (f : iprod (λ x, B1 x -n> B2 x)) : + iprodRA B1 -n> iprodRA B2 := CofeMor (iprod_map f). +Instance laterRA_map_ne {A} {B1 B2 : A → cmraT} n : + Proper (dist n ==> dist n) (@iprodRA_map A B1 B2) := _. diff --git a/algebra/cofe.v b/algebra/cofe.v index 2b48fcc6a..e2e0fd437 100644 --- a/algebra/cofe.v +++ b/algebra/cofe.v @@ -371,9 +371,17 @@ Proof. intros n f g Hf n'; apply Hf. Qed. (** Indexed product *) (** Need to put this in a definition to make canonical structures to work. *) Definition iprod {A} (B : A → cofeT) := ∀ x, B x. +Definition iprod_insert `{∀ x x' : A, Decision (x = x')} {B : A → cofeT} + (x : A) (y : B x) (f : iprod B) : iprod B := λ x', + match decide (x = x') with left H => eq_rect _ B y _ H | right _ => f x' end. +Definition iprod_singleton + `{∀ x x' : A, Decision (x = x')} {B : A → cofeT} `{∀ x : A, Empty (B x)} + (x : A) (y : B x) : iprod B := iprod_insert x y (λ _, ∅). Section iprod_cofe. Context {A} {B : A → cofeT}. + Implicit Types x : A. + Implicit Types f g : iprod B. Instance iprod_equiv : Equiv (iprod B) := λ f g, ∀ x, f x ≡ g x. Instance iprod_dist : Dist (iprod B) := λ n f g, ∀ x, f x ={n}= g x. Program Definition iprod_chain (c : chain (iprod B)) (x : A) : chain (B x) := @@ -397,6 +405,48 @@ Section iprod_cofe. apply (chain_cauchy c); lia. Qed. Canonical Structure iprodC : cofeT := CofeT iprod_cofe_mixin. + + Context `{∀ x x' : A, Decision (x = x')}. + Global Instance iprod_insert_ne x n : + Proper (dist n ==> dist n ==> dist n) (iprod_insert x). + Proof. + intros y1 y2 ? f1 f2 ? x'; rewrite /iprod_insert. + by destruct (decide _) as [[]|]. + Qed. + Global Instance iprod_insert_proper x : + Proper ((≡) ==> (≡) ==> (≡)) (iprod_insert x) := ne_proper_2 _. + Lemma iprod_lookup_insert f x y : (iprod_insert x y f) x = y. + Proof. + rewrite /iprod_insert; destruct (decide _) as [Hx|]; last done. + by rewrite (proof_irrel Hx eq_refl). + Qed. + Lemma iprod_lookup_insert_ne f x x' y : + x ≠x' → (iprod_insert x y f) x' = f x'. + Proof. by rewrite /iprod_insert; destruct (decide _). Qed. + + Context `{∀ x : A, Empty (B x)}. + Global Instance iprod_singleton_ne x n : + Proper (dist n ==> dist n) (iprod_singleton x). + Proof. by intros y1 y2 Hy; rewrite /iprod_singleton Hy. Qed. + Global Instance iprod_singleton_proper x : + Proper ((≡) ==> (≡)) (iprod_singleton x) := ne_proper _. + Lemma iprod_lookup_singleton x y : (iprod_singleton x y) x = y. + Proof. by rewrite /iprod_singleton iprod_lookup_insert. Qed. + Lemma iprod_lookup_singleton_ne x x' y : + x ≠x' → (iprod_singleton x y) x' = ∅. + Proof. intros; by rewrite /iprod_singleton iprod_lookup_insert_ne. Qed. End iprod_cofe. Arguments iprodC {_} _. + +Definition iprod_map {A} {B1 B2 : A → cofeT} (f : ∀ x, B1 x → B2 x) + (g : iprod B1) : iprod B2 := λ x, f _ (g x). +Instance iprod_map_ne {A} {B1 B2 : A → cofeT} (f : ∀ x, B1 x → B2 x) n : + (∀ x, Proper (dist n ==> dist n) (f x)) → + Proper (dist n ==> dist n) (iprod_map f). +Proof. by intros ? y1 y2 Hy x; rewrite /iprod_map (Hy x). Qed. +Definition iprodC_map {A} {B1 B2 : A → cofeT} (f : iprod (λ x, B1 x -n> B2 x)) : + iprodC B1 -n> iprodC B2 := CofeMor (iprod_map f). +Instance laterC_map_ne {A} {B1 B2 : A → cofeT} n : + Proper (dist n ==> dist n) (@iprodC_map A B1 B2). +Proof. intros f1 f2 Hf g x; apply Hf. Qed. diff --git a/algebra/fin_maps.v b/algebra/fin_maps.v index d9062626f..1ee9030b4 100644 --- a/algebra/fin_maps.v +++ b/algebra/fin_maps.v @@ -202,7 +202,7 @@ Proof. intros j; move: (Hm j)=>{Hm}; rewrite !lookup_op=>Hm. destruct (decide (i = j)); simplify_map_equality'; auto. Qed. -Lemma map_insert_updateP' (P : A → Prop) (Q : gmap K A → Prop) m i x : +Lemma map_insert_updateP' (P : A → Prop) m i x : x ~~>: P → <[i:=x]>m ~~>: λ m', ∃ y, m' = <[i:=y]>m ∧ P y. Proof. eauto using map_insert_updateP. Qed. Lemma map_insert_update m i x y : x ~~> y → <[i:=x]>m ~~> <[i:=y]>m. -- GitLab