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

Simplify big_opM again (revert e760dfb5).

parent 785b2175
No related branches found
No related tags found
No related merge requests found
......@@ -55,9 +55,9 @@ Fixpoint big_op `{Op A, Empty A} (xs : list A) : A :=
Arguments big_op _ _ _ !_ /.
Instance: Params (@big_op) 3.
Definition big_opM `{FinMapToList K A M, Op B, Empty B}
(f : K A list B) (m : M) : B := big_op (map_to_list m ≫= curry f).
Instance: Params (@big_opM) 4.
Definition big_opM `{FinMapToList K A M, Op A, Empty A} (m : M) : A :=
big_op (snd <$> map_to_list m).
Instance: Params (@big_opM) 6.
(** Updates *)
Definition ra_update_set `{Op A, Valid A} (x : A) (P : A Prop) :=
......@@ -141,26 +141,21 @@ Proof.
Qed.
Context `{FinMap K M}.
Context `{Equiv B} `{!Equivalence (() : relation B)} (f : K B list A).
Lemma big_opM_empty : big_opM f ( : M B) ∅.
Proof. by unfold big_opM; rewrite map_to_list_empty. Qed.
Lemma big_opM_insert (m : M B) i (y : B) :
m !! i = None big_opM f (<[i:=y]> m) big_op (f i y) big_opM f m.
Proof.
intros ?; unfold big_opM.
by rewrite map_to_list_insert, bind_cons, big_op_app by done.
Qed.
Lemma big_opM_singleton i (y : B) : big_opM f ({[i,y]} : M B) big_op (f i y).
Lemma big_opM_empty : big_opM ( : M A) ∅.
Proof. unfold big_opM. by rewrite map_to_list_empty. Qed.
Lemma big_opM_insert (m : M A) i x :
m !! i = None big_opM (<[i:=x]> m) x big_opM m.
Proof. intros ?; unfold big_opM. by rewrite map_to_list_insert by done. Qed.
Lemma big_opM_singleton i x : big_opM ({[i,x]} : M A) x.
Proof.
unfold singleton, map_singleton.
rewrite big_opM_insert by auto using lookup_empty; simpl.
by rewrite big_opM_empty, (right_id _ _).
Qed.
Global Instance big_opM_proper :
( i, Proper (() ==> ()) (f i)) Proper (() ==> ()) (big_opM f: M B A).
Global Instance big_opM_proper : Proper (() ==> ()) (big_opM : M A _).
Proof.
intros Hf m1; induction m1 as [|i x m1 ? IH] using map_ind.
{ by intros m2; rewrite (symmetry_iff () ), map_equiv_empty; intros ->. }
intros m1; induction m1 as [|i x m1 ? IH] using map_ind.
{ by intros m2; rewrite (symmetry_iff ()), map_equiv_empty; intros ->. }
intros m2 Hm2; rewrite big_opM_insert by done.
rewrite (IH (delete i m2)) by (by rewrite <-Hm2, delete_insert).
destruct (map_equiv_lookup (<[i:=x]> m1) m2 i x)
......
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