diff --git a/opam b/opam index 790563b1e03442add1370dd746edc460934dd127..a65068707896f833688e6e3c0c38eadb45004813 100644 --- a/opam +++ b/opam @@ -12,5 +12,5 @@ install: [make "install"] remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris"] depends: [ "coq" { (= "8.7.2") | (= "8.8.2") | (>= "8.9" & < "8.11~") | (= "dev") } - "coq-stdpp" { (= "dev.2019-07-08.0.2e0bf441") | (= "dev") } + "coq-stdpp" { (= "dev.2019-08-13.1.6827b242") | (= "dev") } ] diff --git a/theories/algebra/gmap.v b/theories/algebra/gmap.v index b41d6458a13b1cedbf814f7d541a25d35d4e8adc..38fb0a48a19c329eb3ba5ce6efb0f53cf1faa98a 100644 --- a/theories/algebra/gmap.v +++ b/theories/algebra/gmap.v @@ -34,7 +34,7 @@ Definition gmap_compl `{Cofe A} : Compl gmapO := λ c, Global Program Instance gmap_cofe `{Cofe A} : Cofe gmapO := {| compl := gmap_compl |}. Next Obligation. - intros ? n c k. rewrite /compl /gmap_compl lookup_imap. + intros ? n c k. rewrite /compl /gmap_compl map_lookup_imap. feed inversion (λ H, chain_cauchy c 0 n H k);simplify_option_eq;auto with lia. by rewrite conv_compl /=; apply reflexive_eq. Qed. @@ -156,7 +156,7 @@ Proof. last by rewrite -lookup_op. exists (map_imap (λ i _, projT1 (FUN i)) y1). exists (map_imap (λ i _, proj1_sig (projT2 (FUN i))) y2). - split; [|split]=>i; rewrite ?lookup_op !lookup_imap; + split; [|split]=>i; rewrite ?lookup_op !map_lookup_imap; destruct (FUN i) as (z1i&z2i&Hmi&Hz1i&Hz2i)=>/=. + destruct (y1 !! i), (y2 !! i); inversion Hz1i; inversion Hz2i; subst=>//. + revert Hz1i. case: (y1!!i)=>[?|] //.