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

Oops, forgot `as`...

parent 75b7a0b4
No related branches found
No related tags found
No related merge requests found
Pipeline #78842 passed
...@@ -179,7 +179,7 @@ Global Instance proto_map_aux_contractive {V} ...@@ -179,7 +179,7 @@ Global Instance proto_map_aux_contractive {V}
Proof. Proof.
intros n rec1 rec2 Hrec p. simpl. apply proto_elim_ne=> // a m1 m2 Hm. intros n rec1 rec2 Hrec p. simpl. apply proto_elim_ne=> // a m1 m2 Hm.
f_equiv=> v p' /=. do 2 f_equiv; [done|]. f_equiv=> v p' /=. do 2 f_equiv; [done|].
apply Next_contractive; by dist_later_intro n' Hn'. apply Next_contractive; by dist_later_intro as n' Hn'.
Qed. Qed.
Definition proto_map_aux_2 {V} Definition proto_map_aux_2 {V}
...@@ -239,7 +239,7 @@ Proof. ...@@ -239,7 +239,7 @@ Proof.
destruct (proto_case p) as [->|(a & m & ->)]; [by rewrite !proto_map_end|]. destruct (proto_case p) as [->|(a & m & ->)]; [by rewrite !proto_map_end|].
rewrite !proto_map_message /=. rewrite !proto_map_message /=.
apply proto_message_ne=> // v p' /=. f_equiv; [done|]. f_equiv. apply proto_message_ne=> // v p' /=. f_equiv; [done|]. f_equiv.
apply Next_contractive; dist_later_intro n' Hn'; eauto using dist_le with lia. apply Next_contractive; dist_later_intro as n' Hn'; eauto using dist_le with lia.
Qed. Qed.
Lemma proto_map_ext {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} Lemma proto_map_ext {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'}
(gn1 gn2 : PROPn' -n> PROPn) (g1 g2 : PROP -n> PROP') p : (gn1 gn2 : PROPn' -n> PROPn) (g1 g2 : PROP -n> PROP') p :
...@@ -255,7 +255,7 @@ Proof. ...@@ -255,7 +255,7 @@ Proof.
induction (lt_wf n) as [n _ IH]=> PROPn ? PROP ? p /=. induction (lt_wf n) as [n _ IH]=> PROPn ? PROP ? p /=.
destruct (proto_case p) as [->|(a & m & ->)]; [by rewrite !proto_map_end|]. destruct (proto_case p) as [->|(a & m & ->)]; [by rewrite !proto_map_end|].
rewrite !proto_map_message /=. apply proto_message_ne=> // v p' /=. f_equiv. rewrite !proto_map_message /=. apply proto_message_ne=> // v p' /=. f_equiv.
apply Next_contractive; dist_later_intro n' Hn'; auto. apply Next_contractive; dist_later_intro as n' Hn'; auto.
Qed. Qed.
Lemma proto_map_compose {V} Lemma proto_map_compose {V}
`{Hcn:!Cofe PROPn, Hcn':!Cofe PROPn', Hcn'':!Cofe PROPn'', `{Hcn:!Cofe PROPn, Hcn':!Cofe PROPn', Hcn'':!Cofe PROPn'',
...@@ -270,7 +270,7 @@ Proof. ...@@ -270,7 +270,7 @@ Proof.
PROP ? PROP' ? PROP'' ? gn1 gn2 g1 g2 p /=. PROP ? PROP' ? PROP'' ? gn1 gn2 g1 g2 p /=.
destruct (proto_case p) as [->|(a & c & ->)]; [by rewrite !proto_map_end|]. destruct (proto_case p) as [->|(a & c & ->)]; [by rewrite !proto_map_end|].
rewrite !proto_map_message /=. apply proto_message_ne=> // v p' /=. rewrite !proto_map_message /=. apply proto_message_ne=> // v p' /=.
do 3 f_equiv. apply Next_contractive; dist_later_intro n' Hn'; simpl; auto. do 3 f_equiv. apply Next_contractive; dist_later_intro as n' Hn'; simpl; auto.
Qed. Qed.
Program Definition protoOF (V : Type) (Fn F : oFunctor) Program Definition protoOF (V : Type) (Fn F : oFunctor)
...@@ -302,6 +302,6 @@ Global Instance protoOF_contractive (V : Type) (Fn F : oFunctor) ...@@ -302,6 +302,6 @@ Global Instance protoOF_contractive (V : Type) (Fn F : oFunctor)
Proof. Proof.
intros HFn HF A1 ? A2 ? B1 ? B2 ? n f g Hfg p; simpl in *. intros HFn HF A1 ? A2 ? B1 ? B2 ? n f g Hfg p; simpl in *.
apply proto_map_ne=> y //=. apply proto_map_ne=> y //=.
+ apply HFn. dist_later_intro n' Hn'. f_equiv; apply Hfg. + apply HFn. dist_later_intro as n' Hn'. f_equiv; apply Hfg.
+ apply HF. by dist_later_intro n' Hn'. + apply HF. by dist_later_intro as n' Hn'.
Qed. Qed.
...@@ -115,7 +115,7 @@ Section term_types. ...@@ -115,7 +115,7 @@ Section term_types.
Proof. Proof.
intros A A' ? B B' ?. apply Ltty_ne=> v. f_equiv=> w. intros A A' ? B B' ?. apply Ltty_ne=> v. f_equiv=> w.
f_equiv; [by f_contractive|]. f_equiv; [by f_contractive|].
apply (wp_contractive _ _ _ _ _)=> v'. dist_later_intro n' Hn'. by f_equiv. apply (wp_contractive _ _ _ _ _)=> v'. dist_later_intro as n' Hn'. by f_equiv.
Qed. Qed.
Global Instance lty_arr_ne `{heapGS Σ} : NonExpansive2 lty_arr. Global Instance lty_arr_ne `{heapGS Σ} : NonExpansive2 lty_arr.
Proof. solve_proper. Qed. Proof. solve_proper. Qed.
...@@ -135,7 +135,7 @@ Section term_types. ...@@ -135,7 +135,7 @@ Section term_types.
Proof. Proof.
intros F F' A. apply Ltty_ne=> w. f_equiv=> B. intros F F' A. apply Ltty_ne=> w. f_equiv=> B.
apply (wp_contractive _ _ _ _ _)=> u. specialize (A B). apply (wp_contractive _ _ _ _ _)=> u. specialize (A B).
dist_later_intro n' Hn'. by f_equiv. dist_later_intro as n' Hn'. by f_equiv.
Qed. Qed.
Global Instance lty_forall_ne `{heapGS Σ} k n : Global Instance lty_forall_ne `{heapGS Σ} k n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@lty_forall Σ _ k). Proper (pointwise_relation _ (dist n) ==> dist n) (@lty_forall Σ _ k).
......
...@@ -72,14 +72,14 @@ Section cofe_solver_2. ...@@ -72,14 +72,14 @@ Section cofe_solver_2.
rewrite -{2}(ofe_iso_12 T_result y). f_equiv. rewrite -{2}(ofe_iso_12 T_result y). f_equiv.
rewrite -(oFunctor_map_id (F _) (ofe_iso_2 T_result y)) rewrite -(oFunctor_map_id (F _) (ofe_iso_2 T_result y))
-!oFunctor_map_compose. -!oFunctor_map_compose.
f_equiv; apply Fcontr; dist_later_intro n' Hn'; split=> x /=; f_equiv; apply Fcontr; dist_later_intro as n' Hn'; split=> x /=;
rewrite ofe_iso_21 -{2}(oFunctor_map_id (F _) x) rewrite ofe_iso_21 -{2}(oFunctor_map_id (F _) x)
-!oFunctor_map_compose; do 2 f_equiv; split=> z /=; auto. -!oFunctor_map_compose; do 2 f_equiv; split=> z /=; auto.
- intros y. apply equiv_dist=> n. revert An Hcn A Hc y. - intros y. apply equiv_dist=> n. revert An Hcn A Hc y.
induction (lt_wf n) as [n _ IH]; intros An ? A ? y. induction (lt_wf n) as [n _ IH]; intros An ? A ? y.
rewrite T_iso_fun_unfold_1 T_iso_fun_unfold_2 /= ofe_iso_21. rewrite T_iso_fun_unfold_1 T_iso_fun_unfold_2 /= ofe_iso_21.
rewrite -(oFunctor_map_id (F _) y) -!oFunctor_map_compose. rewrite -(oFunctor_map_id (F _) y) -!oFunctor_map_compose.
f_equiv; apply Fcontr; dist_later_intro n' Hn'; split=> x /=; f_equiv; apply Fcontr; dist_later_intro as n' Hn'; split=> x /=;
rewrite -{2}(ofe_iso_12 T_result x); f_equiv; rewrite -{2}(ofe_iso_12 T_result x); f_equiv;
rewrite -{2}(oFunctor_map_id (F _) (ofe_iso_2 T_result x)) rewrite -{2}(oFunctor_map_id (F _) (ofe_iso_2 T_result x))
-!oFunctor_map_compose; -!oFunctor_map_compose;
......
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