Skip to content
Snippets Groups Projects
Commit 0c77ba26 authored by Dan Frumin's avatar Dan Frumin
Browse files

solve_proper_from_ne tactic

parent 80ec05b7
No related branches found
No related tags found
No related merge requests found
-Q theories reloc -Q theories reloc
-arg -w -arg -notation-overridden,-redundant-canonical-projection,-several-object-files -arg -w -arg -notation-overridden,-redundant-canonical-projection,-several-object-files
theories/prelude/ctx_subst.v theories/prelude/ctx_subst.v
theories/prelude/tactics.v
theories/logic/spec_ra.v theories/logic/spec_ra.v
theories/logic/spec_rules.v theories/logic/spec_rules.v
theories/logic/model.v theories/logic/model.v
......
...@@ -7,7 +7,7 @@ From iris.heap_lang Require Export lifting metatheory. ...@@ -7,7 +7,7 @@ From iris.heap_lang Require Export lifting metatheory.
From iris.base_logic.lib Require Import invariants. From iris.base_logic.lib Require Import invariants.
From iris.algebra Require Import list gmap. From iris.algebra Require Import list gmap.
From iris.heap_lang Require Import notation proofmode. From iris.heap_lang Require Import notation proofmode.
From reloc Require Import logic.spec_rules prelude.ctx_subst. From reloc Require Import prelude.tactics logic.spec_rules prelude.ctx_subst.
From reloc Require Export logic.spec_ra. From reloc Require Export logic.spec_ra.
(** Semantic intepretation of types *) (** Semantic intepretation of types *)
...@@ -46,7 +46,7 @@ Section lty2_ofe. ...@@ -46,7 +46,7 @@ Section lty2_ofe.
Global Instance lty2_car_ne n : Proper (dist n ==> (=) ==> (=) ==> dist n) lty2_car. Global Instance lty2_car_ne n : Proper (dist n ==> (=) ==> (=) ==> dist n) lty2_car.
Proof. by intros A A' ? w1 w2 <- ? ? <-. Qed. Proof. by intros A A' ? w1 w2 <- ? ? <-. Qed.
Global Instance lty2_car_proper : Proper (() ==> (=) ==> (=) ==> ()) lty2_car. Global Instance lty2_car_proper : Proper (() ==> (=) ==> (=) ==> ()) lty2_car.
Proof. by intros A A' ? w1 w2 <- ? ? <-. Qed. Proof. solve_proper_from_ne. Qed.
End lty2_ofe. End lty2_ofe.
Section semtypes. Section semtypes.
...@@ -65,7 +65,7 @@ Section semtypes. ...@@ -65,7 +65,7 @@ Section semtypes.
Global Instance interp_expr_proper E e e' : Global Instance interp_expr_proper E e e' :
Proper (() ==> ()) (interp_expr E e e'). Proper (() ==> ()) (interp_expr E e e').
Proof. apply ne_proper=>n. by apply interp_expr_ne. Qed. Proof. apply ne_proper=>n. apply _. Qed.
Definition lty2_unit : lty2 := Lty2 (λ w1 w2, w1 = #() w2 = #() ⌝%I). Definition lty2_unit : lty2 := Lty2 (λ w1 w2, w1 = #() w2 = #() ⌝%I).
Definition lty2_bool : lty2 := Lty2 (λ w1 w2, b : bool, w1 = #b w2 = #b )%I. Definition lty2_bool : lty2 := Lty2 (λ w1 w2, b : bool, w1 = #b w2 = #b )%I.
...@@ -185,7 +185,7 @@ Instance option_mbind_ne {A B : ofeT} n : ...@@ -185,7 +185,7 @@ Instance option_mbind_ne {A B : ofeT} n :
Proper (((dist n) ==> (dist n)) ==> (dist n) ==> (dist n)) (@option_bind A B). Proper (((dist n) ==> (dist n)) ==> (dist n) ==> (dist n)) (@option_bind A B).
Proof. destruct 2; simpl; try constructor; auto. Qed. Proof. destruct 2; simpl; try constructor; auto. Qed.
(** todo generalize *) (** todo move somewhere *)
Instance map_zip_with_ne {K} {A B C : ofeT} (f : A B C) Instance map_zip_with_ne {K} {A B C : ofeT} (f : A B C)
`{Countable K} `{!EqDecision K} n : `{Countable K} `{!EqDecision K} n :
Proper (dist n ==> dist n ==> dist n) f Proper (dist n ==> dist n ==> dist n) f
...@@ -204,42 +204,18 @@ Proof. ...@@ -204,42 +204,18 @@ Proof.
intros Γ Γ' ? vvs ->. intros Γ Γ' ? vvs ->.
rewrite /env_ltyped2. rewrite /env_ltyped2.
f_equiv. f_equiv.
- repeat f_equiv. split. - f_equiv.
{ intros Hvvs x. split; intros HH. split;
- apply Hvvs. destruct (Γ !! x) as [?|] eqn:lawl; eauto. intros Hvvs x; specialize ( x); rewrite -(Hvvs x);
specialize ( x). revert . by apply (is_Some_ne n).
destruct HH as [? HH].
rewrite HH lawl. inversion 1.
- destruct (Γ' !! x) as [?|] eqn:lawl; eauto.
specialize ( x). revert .
apply (Hvvs x) in HH.
destruct HH as [? HH].
rewrite HH lawl. inversion 1. }
(* MM TASTY COPYPASTA WITH CARBONARYA SAUCE *)
{ intros Hvvs x. split; intros HH.
- apply Hvvs. destruct (Γ' !! x) as [?|] eqn:lawl; eauto.
specialize ( x). revert .
destruct HH as [? HH].
rewrite HH lawl. inversion 1.
- destruct (Γ !! x) as [?|] eqn:lawl; eauto.
specialize ( x). revert .
apply (Hvvs x) in HH.
destruct HH as [? HH].
rewrite HH lawl. inversion 1. }
- apply big_opM_ne2; first apply _. - apply big_opM_ne2; first apply _.
+ by intros x A B ->. + by intros x A B ->.
+ apply map_zip_with_ne=>//. apply _. + apply map_zip_with_ne=>//. apply _.
Qed. Qed.
(* Hmmm *)
Instance env_ltyped2_proper `{relocG Σ} : Instance env_ltyped2_proper `{relocG Σ} :
Proper (() ==> (=) ==> ()) env_ltyped2. Proper (() ==> (=) ==> ()) env_ltyped2.
Proof. Proof. solve_proper_from_ne. Qed.
intros Γ Γ' ? vvs ->.
apply equiv_dist=>n.
setoid_rewrite equiv_dist in .
by rewrite .
Qed.
Section refinement. Section refinement.
Context `{relocG Σ}. Context `{relocG Σ}.
...@@ -264,17 +240,9 @@ Section refinement. ...@@ -264,17 +240,9 @@ Section refinement.
solve_proper. solve_proper.
Qed. Qed.
(* TODO: this is a bit icky *)
Global Instance refines_proper E : Global Instance refines_proper E :
Proper (() ==> (=) ==> (=) ==> () ==> ()) (refines E). Proper (() ==> (=) ==> (=) ==> () ==> ()) (refines E).
Proof. Proof. solve_proper_from_ne. Qed.
intros Γ Γ' ? e -> ? e' -> A A' HA.
apply equiv_dist=>n.
setoid_rewrite equiv_dist in HA.
setoid_rewrite equiv_dist in .
by rewrite HA .
Defined.
End refinement. End refinement.
Notation "⟦ A ⟧ₑ" := (λ e e', interp_expr e e' A). Notation "⟦ A ⟧ₑ" := (λ e e', interp_expr e e' A).
......
From stdpp Require Export tactics.
From iris.algebra Require Export ofe.
(* Hmmm *)
Ltac my_prepare :=
intros;
repeat lazymatch goal with
| |- Proper _ _ => intros ???
| |- (_ ==> _)%signature _ _ => intros ???
| |- pointwise_relation _ _ _ _ => intros ?
end; simplify_eq.
Ltac solve_proper_from_ne :=
my_prepare;
solve [repeat first [done | eassumption | apply equiv_dist=>? |
match goal with
| [H : _ _ |- _] => setoid_rewrite equiv_dist in H; rewrite H
end] ].
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