diff --git a/opam b/opam index 92354b9360598c891f1778d36a2ca75cad49fda0..61eb7dd91befd7a3d0d2dcfe3a5f38d841029938 100644 --- a/opam +++ b/opam @@ -9,5 +9,5 @@ build: [make "-j%{jobs}%"] install: [make "install"] remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/actris" ] depends: [ - "coq-iris" { (= "dev.2020-05-26.1.d80e7abf") | (= "dev") } + "coq-iris" { (= "dev.2020-05-28.2.2643cb7d") | (= "dev") } ] diff --git a/theories/utils/compare.v b/theories/utils/compare.v index e4f7453d964939a7b50367284f67aadadece1f3d..d6c7187c6549e1edb03b058ac868d28eecfcab06 100644 --- a/theories/utils/compare.v +++ b/theories/utils/compare.v @@ -15,7 +15,7 @@ Definition cmp_spec `{!heapG Σ} {A} (I : A → val → iProp Σ) Definition IZ `{!heapG Σ} (x : Z) (v : val) : iProp Σ := ⌜v = #xâŒ%I. Definition cmpZ : val := λ: "x" "y", "x" ≤ "y". -Lemma cmpZ_spec `{!heapG Σ} : ⊢ cmp_spec IZ (≤) cmpZ. +Lemma cmpZ_spec `{!heapG Σ} : ⊢ cmp_spec IZ (≤)%Z cmpZ. Proof. rewrite /IZ. iIntros (x x' v v' Φ [-> ->]) "!> HΦ". wp_lam. wp_pures. by iApply "HΦ". diff --git a/theories/utils/switch.v b/theories/utils/switch.v index 588031079cf29f6b8d2b6c36b2e578e398f84511..1e2d61e6aa5fc48505a29d32371c9fb91389af8a 100644 --- a/theories/utils/switch.v +++ b/theories/utils/switch.v @@ -27,7 +27,7 @@ Fixpoint map_string_seq {A} end. Lemma lookup_map_string_seq_Some {A} (j i : nat) (xs : list A) : - map_string_seq "f" j xs !! ("f" +:+ pretty (i + j)%nat) = xs !! i. + map_string_seq "f" j xs !! ("f" +:+ pretty (i + j)) = xs !! i. Proof. revert i j. induction xs as [|x xs IH]=> -[|i] j //=. - by rewrite lookup_insert. @@ -44,7 +44,7 @@ Proof. Qed. Lemma lookup_map_string_seq_None_lt {A} y i j (xs : list A) : - (i < j)%nat → map_string_seq y j xs !! (y +:+ pretty i) = None. + i < j → map_string_seq y j xs !! (y +:+ pretty i) = None. Proof. revert j. induction xs as [|x xs IH]=> j ? //=. rewrite lookup_insert_ne; last (intros ?; simplify_eq/=; lia).