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

Turn iProp into a notation.

This avoids Coq distinguishing iProp and uPred (iResUR _) when it
should not.
parent 3ba59838
No related branches found
No related tags found
No related merge requests found
...@@ -116,7 +116,7 @@ Module Type iProp_solution_sig. ...@@ -116,7 +116,7 @@ Module Type iProp_solution_sig.
Parameter iPreProp : gFunctors cofeT. Parameter iPreProp : gFunctors cofeT.
Definition iResUR (Σ : gFunctors) : ucmraT := Definition iResUR (Σ : gFunctors) : ucmraT :=
iprodUR (λ i, gmapUR gname (Σ i (iPreProp Σ))). iprodUR (λ i, gmapUR gname (Σ i (iPreProp Σ))).
Definition iProp (Σ : gFunctors) : cofeT := uPredC (iResUR Σ). Notation iProp Σ := (uPredC (iResUR Σ)).
Parameter iProp_unfold: {Σ}, iProp Σ -n> iPreProp Σ. Parameter iProp_unfold: {Σ}, iProp Σ -n> iPreProp Σ.
Parameter iProp_fold: {Σ}, iPreProp Σ -n> iProp Σ. Parameter iProp_fold: {Σ}, iPreProp Σ -n> iProp Σ.
...@@ -134,7 +134,7 @@ Module Export iProp_solution : iProp_solution_sig. ...@@ -134,7 +134,7 @@ Module Export iProp_solution : iProp_solution_sig.
Definition iPreProp (Σ : gFunctors) : cofeT := iProp_result Σ. Definition iPreProp (Σ : gFunctors) : cofeT := iProp_result Σ.
Definition iResUR (Σ : gFunctors) : ucmraT := Definition iResUR (Σ : gFunctors) : ucmraT :=
iprodUR (λ i, gmapUR gname (Σ i (iPreProp Σ))). iprodUR (λ i, gmapUR gname (Σ i (iPreProp Σ))).
Definition iProp (Σ : gFunctors) : cofeT := uPredC (iResUR Σ). Notation iProp Σ := (uPredC (iResUR Σ)).
Definition iProp_unfold {Σ} : iProp Σ -n> iPreProp Σ := Definition iProp_unfold {Σ} : iProp Σ -n> iPreProp Σ :=
solution_fold (iProp_result Σ). solution_fold (iProp_result Σ).
...@@ -145,8 +145,6 @@ Module Export iProp_solution : iProp_solution_sig. ...@@ -145,8 +145,6 @@ Module Export iProp_solution : iProp_solution_sig.
Proof. apply solution_fold_unfold. Qed. Proof. apply solution_fold_unfold. Qed.
End iProp_solution. End iProp_solution.
Bind Scope uPred_scope with iProp.
(** * Properties of the solution to the recursive domain equation *) (** * Properties of the solution to the recursive domain equation *)
Lemma iProp_unfold_equivI {Σ} (P Q : iProp Σ) : Lemma iProp_unfold_equivI {Σ} (P Q : iProp Σ) :
......
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