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

add some notation for writing cFunctors

parent 0a279595
No related branches found
No related tags found
No related merge requests found
...@@ -347,6 +347,9 @@ Structure cFunctor := CFunctor { ...@@ -347,6 +347,9 @@ Structure cFunctor := CFunctor {
Existing Instance cFunctor_ne. Existing Instance cFunctor_ne.
Instance: Params (@cFunctor_map) 5. Instance: Params (@cFunctor_map) 5.
Delimit Scope cFunctor_scope with CF.
Bind Scope cFunctor_scope with cFunctor.
Class cFunctorContractive (F : cFunctor) := Class cFunctorContractive (F : cFunctor) :=
cFunctor_contractive A1 A2 B1 B2 :> Contractive (@cFunctor_map F A1 A2 B1 B2). cFunctor_contractive A1 A2 B1 B2 :> Contractive (@cFunctor_map F A1 A2 B1 B2).
...@@ -412,6 +415,7 @@ Proof. ...@@ -412,6 +415,7 @@ Proof.
apply cofe_morC_map_ne; apply cFunctor_contractive=>i ?; split; by apply Hfg. apply cofe_morC_map_ne; apply cFunctor_contractive=>i ?; split; by apply Hfg.
Qed. Qed.
(** Discrete cofe *) (** Discrete cofe *)
Section discrete_cofe. Section discrete_cofe.
Context `{Equiv A, @Equivalence A ()}. Context `{Equiv A, @Equivalence A ()}.
...@@ -518,3 +522,10 @@ Proof. ...@@ -518,3 +522,10 @@ Proof.
intros A1 A2 B1 B2 n fg fg' Hfg. intros A1 A2 B1 B2 n fg fg' Hfg.
apply laterC_map_contractive => i ?. by apply cFunctor_ne, Hfg. apply laterC_map_contractive => i ?. by apply cFunctor_ne, Hfg.
Qed. Qed.
(** Notation for writing functors *)
Notation "∙" := idCF : cFunctor_scope.
Notation "F1 -n> F2" := (cofe_morCF F1%CF F2%CF) : cFunctor_scope.
Notation "( F1 , F2 , .. , Fn )" := (prodCF .. (prodCF F1%CF F2%CF) .. Fn%CF) : cFunctor_scope.
Notation "▶ F" := (laterCF F%CF) (at level 20, right associativity) : cFunctor_scope.
Coercion constCF : cofeT >-> cFunctor.
...@@ -20,7 +20,7 @@ Inductive binder := BAnon | BNamed : string → binder. ...@@ -20,7 +20,7 @@ Inductive binder := BAnon | BNamed : string → binder.
Definition cons_binder (mx : binder) (X : list string) : list string := Definition cons_binder (mx : binder) (X : list string) : list string :=
match mx with BAnon => X | BNamed x => x :: X end. match mx with BAnon => X | BNamed x => x :: X end.
Infix ":b:" := cons_binder (at level 60, right associativity). Infix ":b:" := cons_binder (at level 60, right associativity).
Delimit Scope binder_scope with binder. Delimit Scope binder_scope with bind.
Bind Scope binder_scope with binder. Bind Scope binder_scope with binder.
Instance binder_dec_eq (x1 x2 : binder) : Decision (x1 = x2). Instance binder_dec_eq (x1 x2 : binder) : Decision (x1 = x2).
Proof. solve_decision. Defined. Proof. solve_decision. Defined.
......
...@@ -34,7 +34,7 @@ End iProp_solution_sig. ...@@ -34,7 +34,7 @@ End iProp_solution_sig.
Module Export iProp_solution : iProp_solution_sig. Module Export iProp_solution : iProp_solution_sig.
Definition iProp_result (Λ : language) (Σ : iFunctor) : Definition iProp_result (Λ : language) (Σ : iFunctor) :
solution (uPredCF (resRF Λ (laterCF idCF) Σ)) := solver.result _. solution (uPredCF (resRF Λ ( ) Σ)) := solver.result _.
Definition iPreProp (Λ : language) (Σ : iFunctor) : cofeT := iProp_result Λ Σ. Definition iPreProp (Λ : language) (Σ : iFunctor) : cofeT := iProp_result Λ Σ.
Definition iGst (Λ : language) (Σ : iFunctor) : cmraT := Σ (iPreProp Λ Σ). Definition iGst (Λ : language) (Σ : iFunctor) : cmraT := Σ (iPreProp Λ Σ).
......
...@@ -5,7 +5,7 @@ Import uPred. ...@@ -5,7 +5,7 @@ Import uPred.
Class savedPropG (Λ : language) (Σ : gFunctors) (F : cFunctor) := Class savedPropG (Λ : language) (Σ : gFunctors) (F : cFunctor) :=
saved_prop_inG :> inG Λ Σ (agreeR (laterC (F (iPreProp Λ (globalF Σ))))). saved_prop_inG :> inG Λ Σ (agreeR (laterC (F (iPreProp Λ (globalF Σ))))).
Definition savedPropGF (F : cFunctor) : gFunctor := Definition savedPropGF (F : cFunctor) : gFunctor :=
GFunctor (agreeRF (laterCF F)). GFunctor (agreeRF ( F)).
Instance inGF_savedPropG `{inGF Λ Σ (savedPropGF F)} : savedPropG Λ Σ F. Instance inGF_savedPropG `{inGF Λ Σ (savedPropGF F)} : savedPropG Λ Σ F.
Proof. apply: inGF_inG. Qed. Proof. apply: inGF_inG. Qed.
......
...@@ -9,7 +9,7 @@ End ModelTest. ...@@ -9,7 +9,7 @@ End ModelTest.
Module SavedPropTest. Module SavedPropTest.
(* Test if we can really go "crazy higher order" *) (* Test if we can really go "crazy higher order" *)
Section sec. Section sec.
Definition F := (cofe_morCF idCF idCF). Definition F : cFunctor := ( -n> ).
Definition Σ : gFunctors := #[ savedPropGF F ]. Definition Σ : gFunctors := #[ savedPropGF F ].
Context {Λ : language}. Context {Λ : language}.
Notation iProp := (iPropG Λ Σ). Notation iProp := (iPropG Λ Σ).
......
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