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

Improve `iSpecialize ("H" $! x1 .. xn)`.

Instead of doing all the instantiations by invoking a single type
class search, it now performs the instantiations by invoking
individual type class searches. This a.) gives better error messages
and b.) works when `xj` depends on `xi`.
parent 211c2363
No related branches found
No related tags found
No related merge requests found
...@@ -601,6 +601,13 @@ Global Instance into_exist_laterN {A} n P (Φ : A → uPred M) : ...@@ -601,6 +601,13 @@ Global Instance into_exist_laterN {A} n P (Φ : A → uPred M) :
IntoExist P Φ Inhabited A IntoExist (▷^n P) (λ a, ▷^n (Φ a))%I. IntoExist P Φ Inhabited A IntoExist (▷^n P) (λ a, ▷^n (Φ a))%I.
Proof. rewrite /IntoExist=> HP ?. by rewrite HP laterN_exist. Qed. Proof. rewrite /IntoExist=> HP ?. by rewrite HP laterN_exist. Qed.
(* IntoForall *)
Global Instance into_forall_forall {A} (Φ : A uPred M) : IntoForall ( a, Φ a) Φ.
Proof. done. Qed.
Global Instance into_forall_always {A} P (Φ : A uPred M) :
IntoForall P Φ IntoForall ( P) (λ a, (Φ a))%I.
Proof. rewrite /IntoForall=> HP. by rewrite HP always_forall. Qed.
(* FromModal *) (* FromModal *)
Global Instance from_modal_later P : FromModal ( P) P. Global Instance from_modal_later P : FromModal ( P) P.
Proof. apply later_intro. Qed. Proof. apply later_intro. Qed.
......
...@@ -88,6 +88,11 @@ Class IntoExist {M A} (P : uPred M) (Φ : A → uPred M) := ...@@ -88,6 +88,11 @@ Class IntoExist {M A} (P : uPred M) (Φ : A → uPred M) :=
Arguments into_exist {_ _} _ _ {_}. Arguments into_exist {_ _} _ _ {_}.
Hint Mode IntoExist + - ! - : typeclass_instances. Hint Mode IntoExist + - ! - : typeclass_instances.
Class IntoForall {M A} (P : uPred M) (Φ : A uPred M) :=
into_forall : P x, Φ x.
Arguments into_forall {_ _} _ _ {_}.
Hint Mode IntoForall + - ! - : typeclass_instances.
Class FromModal {M} (P Q : uPred M) := from_modal : Q P. Class FromModal {M} (P Q : uPred M) := from_modal : Q P.
Arguments from_modal {_} _ _ {_}. Arguments from_modal {_} _ _ {_}.
Hint Mode FromModal + ! - : typeclass_instances. Hint Mode FromModal + ! - : typeclass_instances.
......
...@@ -791,27 +791,13 @@ Lemma tac_pure_forall_intro {A} Δ (φ : A → Prop) : ...@@ -791,27 +791,13 @@ Lemma tac_pure_forall_intro {A} Δ (φ : A → Prop) :
( a, Δ φ a) Δ ⌜∀ a, φ a⌝. ( a, Δ φ a) Δ ⌜∀ a, φ a⌝.
Proof. intros. rewrite pure_forall. by apply forall_intro. Qed. Proof. intros. rewrite pure_forall. by apply forall_intro. Qed.
Class ForallSpecialize {As} (xs : hlist As) Lemma tac_forall_specialize {A} Δ Δ' i p P (Φ : A uPred M) Q x :
(P : uPred M) (Φ : himpl As (uPred M)) := forall_specialize : P Φ xs. envs_lookup i Δ = Some (p, P) IntoForall P Φ
Arguments forall_specialize {_} _ _ _ {_}. envs_simple_replace i p (Esnoc Enil i (Φ x)) Δ = Some Δ'
Global Instance forall_specialize_nil P : ForallSpecialize hnil P P | 100.
Proof. done. Qed.
Global Instance forall_specialize_cons A As x xs Φ (Ψ : A himpl As (uPred M)) :
( x, ForallSpecialize xs (Φ x) (Ψ x))
ForallSpecialize (hcons x xs) ( x : A, Φ x) Ψ.
Proof. rewrite /ForallSpecialize /= => <-. by rewrite (forall_elim x). Qed.
Global Instance forall_specialize_always As xs P (Ψ : himpl As (uPred M)) :
ForallSpecialize xs P Ψ ForallSpecialize xs ( P) Ψ | 101.
Proof. rewrite /ForallSpecialize=> ->. by rewrite always_elim. Qed.
Lemma tac_forall_specialize {As} Δ Δ' i p P (Φ : himpl As (uPred M)) Q xs :
envs_lookup i Δ = Some (p, P) ForallSpecialize xs P Φ
envs_simple_replace i p (Esnoc Enil i (Φ xs)) Δ = Some Δ'
(Δ' Q) Δ Q. (Δ' Q) Δ Q.
Proof. Proof.
intros. rewrite envs_simple_replace_sound //; simpl. intros. rewrite envs_simple_replace_sound //; simpl.
by rewrite right_id (forall_specialize _ P) wand_elim_r. by rewrite right_id (into_forall P) (forall_elim x) wand_elim_r.
Qed. Qed.
Lemma tac_forall_revert {A} Δ (Φ : A uPred M) : Lemma tac_forall_revert {A} Δ (Φ : A uPred M) :
...@@ -849,5 +835,3 @@ Proof. ...@@ -849,5 +835,3 @@ Proof.
rewrite right_id always_if_elim. by apply elim_modal. rewrite right_id always_if_elim. by apply elim_modal.
Qed. Qed.
End tactics. End tactics.
Hint Mode ForallSpecialize + - - ! - : typeclass_instances.
...@@ -268,15 +268,17 @@ Notation "( H $! x1 .. xn 'with' pat )" := ...@@ -268,15 +268,17 @@ Notation "( H $! x1 .. xn 'with' pat )" :=
Notation "( H 'with' pat )" := (ITrm H hnil pat) (at level 0). Notation "( H 'with' pat )" := (ITrm H hnil pat) (at level 0).
Local Tactic Notation "iSpecializeArgs" constr(H) open_constr(xs) := Local Tactic Notation "iSpecializeArgs" constr(H) open_constr(xs) :=
match xs with let rec go xs :=
| hnil => idtac match xs with
| _ => | hnil => idtac
eapply tac_forall_specialize with _ H _ _ _ xs; (* (i:=H) (a:=x) *) | hcons ?x ?xs =>
[env_cbv; reflexivity || fail 1 "iSpecialize:" H "not found" eapply tac_forall_specialize with _ H _ _ _ x; (* (i:=H) (a:=x) *)
|let P := match goal with |- ForallSpecialize _ ?P _ => P end in [env_cbv; reflexivity || fail 1 "iSpecialize:" H "not found"
apply _ || fail 1 "iSpecialize:" P "not a forall of the right arity or type" |let P := match goal with |- IntoForall ?P _ => P end in
|cbn [himpl hcurry]; reflexivity|] apply _ || fail 1 "iSpecialize:" P "not a forall or not a forall of the right type"
end. |env_cbv; reflexivity|go xs]
end in
go xs.
Local Tactic Notation "iSpecializePat" constr(H) constr(pat) := Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
let solve_to_wand H1 := let solve_to_wand H1 :=
......
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