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

Change positions of unfold/fold to make things consistent with iProp.

parent 316a5108
Branches robbert/fold_iProto
No related tags found
1 merge request!40Change positions of unfold/fold to make things consistent with iProp.
...@@ -53,48 +53,63 @@ Definition proto_auxO (V : Type) (PROP : ofe) (A : ofe) : ofe := ...@@ -53,48 +53,63 @@ Definition proto_auxO (V : Type) (PROP : ofe) (A : ofe) : ofe :=
Definition proto_auxOF (V : Type) (PROP : ofe) : oFunctor := Definition proto_auxOF (V : Type) (PROP : ofe) : oFunctor :=
optionOF (actionO * (V -d> -n> PROP)). optionOF (actionO * (V -d> -n> PROP)).
Definition proto_result (V : Type) := result_2 (proto_auxOF V). Definition pre_proto_result (V : Type) := result_2 (proto_auxOF V).
Definition pre_proto (V : Type) (PROPn PROP : ofe) `{!Cofe PROPn, !Cofe PROP} : ofe :=
solution_2_car (pre_proto_result V) PROPn _ PROP _.
Global Instance pre_proto_cofe {V} `{!Cofe PROPn, !Cofe PROP} :
Cofe (pre_proto V PROPn PROP).
Proof. apply _. Qed.
Definition proto (V : Type) (PROPn PROP : ofe) `{!Cofe PROPn, !Cofe PROP} : ofe := Definition proto (V : Type) (PROPn PROP : ofe) `{!Cofe PROPn, !Cofe PROP} : ofe :=
solution_2_car (proto_result V) PROPn _ PROP _. proto_auxO V PROP (pre_proto V PROP PROPn).
Global Instance proto_cofe {V} `{!Cofe PROPn, !Cofe PROP} : Cofe (proto V PROPn PROP). Global Instance proto_cofe {V} `{!Cofe PROPn, !Cofe PROP} :
Cofe (proto V PROPn PROP).
Proof. apply _. Qed. Proof. apply _. Qed.
Lemma proto_iso {V} `{!Cofe PROPn, !Cofe PROP} : Lemma proto_iso {V} `{!Cofe PROPn, !Cofe PROP} :
ofe_iso (proto_auxO V PROP (proto V PROP PROPn)) (proto V PROPn PROP). ofe_iso (proto V PROPn PROP) (pre_proto V PROPn PROP).
Proof. apply proto_result. Qed. Proof. apply pre_proto_result. Qed.
Definition proto_fold {V} `{!Cofe PROPn, !Cofe PROP} : Definition proto_fold {V} `{!Cofe PROPn, !Cofe PROP} :
proto_auxO V PROP (proto V PROP PROPn) -n> proto V PROPn PROP := ofe_iso_1 proto_iso. pre_proto V PROPn PROP -n> proto V PROPn PROP := ofe_iso_2 proto_iso.
Definition proto_unfold {V} `{!Cofe PROPn, !Cofe PROP} : Definition proto_unfold {V} `{!Cofe PROPn, !Cofe PROP} :
proto V PROPn PROP -n> proto_auxO V PROP (proto V PROP PROPn) := ofe_iso_2 proto_iso. proto V PROPn PROP -n> pre_proto V PROPn PROP := ofe_iso_1 proto_iso.
Lemma proto_fold_unfold {V} `{!Cofe PROPn, !Cofe PROP} (p : proto V PROPn PROP) : Lemma proto_fold_unfold {V} `{!Cofe PROPn, !Cofe PROP} (p : proto V PROPn PROP) :
proto_fold (proto_unfold p) p. proto_fold (proto_unfold p) p.
Proof. apply (ofe_iso_12 proto_iso). Qed.
Lemma proto_unfold_fold {V} `{!Cofe PROPn, !Cofe PROP}
(p : proto_auxO V PROP (proto V PROP PROPn)) :
proto_unfold (proto_fold p) p.
Proof. apply (ofe_iso_21 proto_iso). Qed. Proof. apply (ofe_iso_21 proto_iso). Qed.
Lemma proto_unfold_fold {V} `{!Cofe PROPn, !Cofe PROP} (p : pre_proto V PROPn PROP) :
proto_unfold (proto_fold p) p.
Proof. apply (ofe_iso_12 proto_iso). Qed.
Definition proto_end {V} `{!Cofe PROPn, !Cofe PROP} : proto V PROPn PROP := Definition proto_end {V} `{!Cofe PROPn, !Cofe PROP} : proto V PROPn PROP :=
proto_fold None. None.
Definition proto_message {V} `{!Cofe PROPn, !Cofe PROP} (a : action) Definition proto_message {V} `{!Cofe PROPn, !Cofe PROP} (a : action)
(m : V laterO (proto V PROP PROPn) -n> PROP) : proto V PROPn PROP := (m : V laterO (proto V PROP PROPn) -n> PROP) : proto V PROPn PROP :=
proto_fold (Some (a, m)). Some (a, λ v, m v laterO_map proto_fold).
Global Instance proto_message_ne {V} `{!Cofe PROPn, !Cofe PROP} a n : Global Instance proto_message_ne {V} `{!Cofe PROPn, !Cofe PROP} a n :
Proper (pointwise_relation V (dist n) ==> dist n) Proper (pointwise_relation V (dist n) ==> dist n)
(proto_message (PROPn:=PROPn) (PROP:=PROP) a). (proto_message (PROPn:=PROPn) (PROP:=PROP) a).
Proof. intros c1 c2 Hc. rewrite /proto_message. f_equiv. by repeat constructor. Qed. Proof.
intros c1 c2 Hc. rewrite /proto_message.
(repeat constructor)=> //= v. by f_equiv.
Qed.
Global Instance proto_message_proper {V} `{!Cofe PROPn, !Cofe PROP} a : Global Instance proto_message_proper {V} `{!Cofe PROPn, !Cofe PROP} a :
Proper (pointwise_relation V () ==> ()) Proper (pointwise_relation V () ==> ())
(proto_message (PROPn:=PROPn) (PROP:=PROP) a). (proto_message (PROPn:=PROPn) (PROP:=PROP) a).
Proof. intros c1 c2 Hc. rewrite /proto_message. f_equiv. by repeat constructor. Qed. Proof.
intros c1 c2 Hc. rewrite /proto_message.
(repeat constructor)=> //= v. by f_equiv.
Qed.
Lemma proto_case {V} `{!Cofe PROPn, !Cofe PROP} (p : proto V PROPn PROP) : Lemma proto_case {V} `{!Cofe PROPn, !Cofe PROP} (p : proto V PROPn PROP) :
p proto_end a m, p proto_message a m. p proto_end a m, p proto_message a m.
Proof. Proof.
destruct (proto_unfold p) as [[a m]|] eqn:E; simpl in *; last first. destruct p as [[a m]|]; [|by left].
- left. by rewrite -(proto_fold_unfold p) E. right. exists a, (λ v, m v laterO_map proto_unfold).
- right. exists a, m. by rewrite /proto_message -E proto_fold_unfold. rewrite /proto_message. do 2 f_equiv. intros v p; simpl. f_equiv.
rewrite -later_map_compose -{1}(later_map_id p).
apply later_map_ext=> p' /=. by rewrite proto_unfold_fold.
Qed. Qed.
Global Instance proto_inhabited {V} `{!Cofe PROPn, !Cofe PROP} : Global Instance proto_inhabited {V} `{!Cofe PROPn, !Cofe PROP} :
Inhabited (proto V PROPn PROP) := populate proto_end. Inhabited (proto V PROPn PROP) := populate proto_end.
...@@ -103,22 +118,16 @@ Lemma proto_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} ...@@ -103,22 +118,16 @@ Lemma proto_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP}
proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a1 m1 proto_message a2 m2 proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a1 m1 proto_message a2 m2
⊣⊢@{SPROP} a1 = a2 ( v p', m1 v p' m2 v p'). ⊣⊢@{SPROP} a1 = a2 ( v p', m1 v p' m2 v p').
Proof. Proof.
trans (proto_unfold (proto_message a1 m1) proto_unfold (proto_message a2 m2) : SPROP)%I. rewrite /proto_message option_equivI prod_equivI /=.
{ iSplit. rewrite discrete_eq discrete_fun_equivI. f_equiv; [done|]. f_equiv=> x.
- iIntros "Heq". by iRewrite "Heq". rewrite ofe_morO_equivI /=. iSplit; iIntros "H %p //".
- iIntros "Heq". rewrite -{2}(proto_fold_unfold (proto_message _ _)). assert (p later_map proto_fold (later_map proto_unfold p)) as ->; last done.
iRewrite "Heq". by rewrite proto_fold_unfold. } rewrite -later_map_compose -{1}(later_map_id p).
rewrite /proto_message !proto_unfold_fold option_equivI prod_equivI /=. apply later_map_ext=> p' /=. by rewrite proto_fold_unfold.
rewrite discrete_eq discrete_fun_equivI.
by setoid_rewrite ofe_morO_equivI.
Qed. Qed.
Lemma proto_message_end_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a m : Lemma proto_message_end_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a m :
proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m proto_end ⊢@{SPROP} False. proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m proto_end ⊢@{SPROP} False.
Proof. Proof. by rewrite option_equivI. Qed.
trans (proto_unfold (proto_message a m) proto_unfold proto_end : SPROP)%I.
{ iIntros "Heq". by iRewrite "Heq". }
by rewrite /proto_message !proto_unfold_fold option_equivI.
Qed.
Lemma proto_end_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a m : Lemma proto_end_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a m :
proto_end proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m ⊢@{SPROP} False. proto_end proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m ⊢@{SPROP} False.
Proof. by rewrite internal_eq_sym proto_message_end_equivI. Qed. Proof. by rewrite internal_eq_sym proto_message_end_equivI. Qed.
...@@ -126,7 +135,11 @@ Proof. by rewrite internal_eq_sym proto_message_end_equivI. Qed. ...@@ -126,7 +135,11 @@ Proof. by rewrite internal_eq_sym proto_message_end_equivI. Qed.
Definition proto_elim {V} `{!Cofe PROPn, !Cofe PROP} {A} Definition proto_elim {V} `{!Cofe PROPn, !Cofe PROP} {A}
(x : A) (f : action (V laterO (proto V PROP PROPn) -n> PROP) A) (x : A) (f : action (V laterO (proto V PROP PROPn) -n> PROP) A)
(p : proto V PROPn PROP) : A := (p : proto V PROPn PROP) : A :=
match proto_unfold p with None => x | Some (a, m) => f a m end. match p with
| None => x
| Some (a, m) => f a (λ v, m v laterO_map proto_unfold)
end.
Global Arguments proto_elim : simpl never.
Lemma proto_elim_ne {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} Lemma proto_elim_ne {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe}
(x : A) (f1 f2 : action (V laterO (proto V PROP PROPn) -n> PROP) A) p1 p2 n : (x : A) (f1 f2 : action (V laterO (proto V PROP PROPn) -n> PROP) A) p1 p2 n :
...@@ -134,31 +147,22 @@ Lemma proto_elim_ne {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} ...@@ -134,31 +147,22 @@ Lemma proto_elim_ne {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe}
p1 {n} p2 p1 {n} p2
proto_elim x f1 p1 {n} proto_elim x f2 p2. proto_elim x f1 p1 {n} proto_elim x f2 p2.
Proof. Proof.
intros Hf Hp. rewrite /proto_elim. intros Hf [[a1 m1] [a2 m2] [[=->] ?]|]; rewrite /proto_elim //=.
apply (_ : NonExpansive proto_unfold) in Hp apply Hf=> v. by f_equiv.
as [[a1 m1] [a2 m2] [-> ?]|]; simplify_eq/=; [|done].
apply Hf. destruct n; by simpl.
Qed. Qed.
Lemma proto_elim_end {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} Lemma proto_elim_end {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe}
(x : A) (f : action (V laterO (proto V PROP PROPn) -n> PROP) A) : (x : A) (f : action (V laterO (proto V PROP PROPn) -n> PROP) A) :
proto_elim x f proto_end x. proto_elim x f proto_end x.
Proof. Proof. done. Qed.
rewrite /proto_elim /proto_end.
pose proof (proto_unfold_fold (V:=V) (PROPn:=PROPn) (PROP:=PROP) None) as Hfold.
by destruct (proto_unfold (proto_fold None)) as [[??]|] eqn:E; inversion Hfold.
Qed.
Lemma proto_elim_message {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} Lemma proto_elim_message {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe}
(x : A) (f : action (V laterO (proto V PROP PROPn) -n> PROP) A) a m : (x : A) (f : action (V laterO (proto V PROP PROPn) -n> PROP) A) a m :
( a, Proper (pointwise_relation _ () ==> ()) (f a)) ( a, Proper (pointwise_relation _ () ==> ()) (f a))
proto_elim x f (proto_message a m) f a m. proto_elim x f (proto_message a m) f a m.
Proof. Proof.
intros. rewrite /proto_elim /proto_message /=. intros. rewrite /proto_elim /proto_message /=. f_equiv=> v p /=. f_equiv.
pose proof (proto_unfold_fold (V:=V) (PROPn:=PROPn) rewrite -later_map_compose -{2}(later_map_id p).
(PROP:=PROP) (Some (a, m))) as Hfold. apply later_map_ext=> p' /=. by rewrite proto_fold_unfold.
destruct (proto_unfold (proto_fold (Some (a, m))))
as [[??]|] eqn:E; inversion Hfold as [?? [Ha Hc]|]; simplify_eq/=.
by f_equiv=> v.
Qed. Qed.
(** Functor *) (** Functor *)
...@@ -214,7 +218,7 @@ Qed. ...@@ -214,7 +218,7 @@ Qed.
Lemma proto_map_end {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} Lemma proto_map_end {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'}
(gn : PROPn' -n> PROPn) (g : PROP -n> PROP') : (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') :
proto_map (V:=V) gn g proto_end proto_end. proto_map (V:=V) gn g proto_end proto_end.
Proof. by rewrite proto_map_unfold /proto_map_aux /= proto_elim_end. Qed. Proof. by rewrite proto_map_unfold /proto_map_aux. Qed.
Lemma proto_map_message {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} Lemma proto_map_message {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'}
(gn : PROPn' -n> PROPn) (g : PROP -n> PROP') a m : (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') a m :
proto_map (V:=V) gn g (proto_message a m) proto_map (V:=V) gn g (proto_message a m)
......
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