diff --git a/coq-gpfsl.opam b/coq-gpfsl.opam index bcce9cd5c02d17c1b9cbb807906367a917b1f665..cd3946f5fa1f9f4988d06fb966734fb4525f46a2 100644 --- a/coq-gpfsl.opam +++ b/coq-gpfsl.opam @@ -10,7 +10,7 @@ version: "dev" synopsis: "A combination of GPS and FSL in the promising semantics WITHOUT promises" depends: [ - "coq-iris" { (= "dev.2025-03-28.0.fa344cbe") | (= "dev") } + "coq-iris" { (= "dev.2025-03-31.0.017605dd") | (= "dev") } ] build: ["./make-package" "gpfsl" "-j%{jobs}%"] diff --git a/gpfsl/orc11/tview.v b/gpfsl/orc11/tview.v index 4e329ceb9df7b966cf7197503db89ef444725577..3ef508b3381821a272daeec481800cb7c5954571 100644 --- a/gpfsl/orc11/tview.v +++ b/gpfsl/orc11/tview.v @@ -13,7 +13,7 @@ Section ThreadView. Decision (∀ l, M !! l ⊑ V). Proof. assert (IFF : (∀ l, M !! l ⊑ V) ↔ (Forall (λ lV', Some lV'.2 ⊑ V) (map_to_list M))). - { rewrite list.Forall_forall. split. + { rewrite Forall_forall. split. - intros ? [l V'] Eq%elem_of_map_to_list. by rewrite /= -Eq. - intros HM l. destruct (M!!l) eqn:Eq; [|done]. apply elem_of_map_to_list in Eq. eapply (HM (_,_)). eauto. } @@ -26,7 +26,7 @@ Section ThreadView. Proof. assert (IFF : (∀ (l : loc) (V' : view), M !! l = Some V' → V ⊑ V') ↔ (Forall (λ lV', V ⊑ lV'.2) (map_to_list M))). - { rewrite list.Forall_forall. split. + { rewrite Forall_forall. split. - intros ? [l V'] Eq%elem_of_map_to_list. eauto. - intros HM l V' ?. by eapply (HM (_, _)), elem_of_map_to_list. } destruct (decide (Forall (λ lV' : loc * view, V ⊑ lV'.2) (map_to_list M)));