Skip to content
Snippets Groups Projects
Commit 29a97bb2 authored by Ralf Jung's avatar Ralf Jung
Browse files

use solve_proper_core for saved_pred_own_contractive

parent 8c13c02f
No related branches found
No related tags found
No related merge requests found
......@@ -620,9 +620,9 @@ Definition ccompose {A B C}
(f : B -n> C) (g : A -n> B) : A -n> C := CofeMor (f g).
Instance: Params (@ccompose) 3.
Infix "◎" := ccompose (at level 40, left associativity).
Lemma ccompose_ne {A B C} (f1 f2 : B -n> C) (g1 g2 : A -n> B) n :
f1 {n} f2 g1 {n} g2 f1 g1 {n} f2 g2.
Proof. by intros Hf Hg x; rewrite /= (Hg x) (Hf (g2 x)). Qed.
Global Instance ccompose_ne {A B C} :
NonExpansive2 (@ccompose A B C).
Proof. intros n ?? Hf g1 g2 Hg x. rewrite /= (Hg x) (Hf (g2 x)) //. Qed.
(* Function space maps *)
Definition ofe_mor_map {A A' B B'} (f : A' -n> A) (g : B -n> B')
......
......@@ -95,8 +95,7 @@ Definition saved_pred_own `{savedPredG Σ A} (γ : gname) (Φ : A -n> iProp Σ)
Instance saved_pred_own_contractive `{savedPredG Σ A} γ : Contractive (saved_pred_own γ).
Proof.
intros n Φ Φ' . rewrite /saved_pred_own /saved_anything_own /=.
do 3 f_equiv. intros x. rewrite /=. by f_contractive.
solve_proper_core ltac:(fun _ => first [ intros ?; progress simpl | f_contractive | f_equiv ]).
Qed.
Lemma saved_pred_alloc_strong `{savedPredG Σ A} (G : gset gname) (Φ : A -n> iProp Σ) :
......
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