diff --git a/theories/base_logic/bi.v b/theories/base_logic/bi.v index 261641c77fcd502e479723dfb4caff195b0c749f..8f86b9c0210e0dbb9dcd159e44db8d77b0916960 100644 --- a/theories/base_logic/bi.v +++ b/theories/base_logic/bi.v @@ -197,7 +197,7 @@ Lemma bupd_ownM_updateP x (Φ : M → Prop) : x ~~>: Φ → uPred_ownM x ⊢ |==> ∃ y, ⌜Φ y⌠∧ uPred_ownM y. Proof. exact: uPred_primitive.bupd_ownM_updateP. Qed. -(* This is really just a special case of an entailment +(** This is really just a special case of an entailment between two [siProp], but we do not have the infrastructure to express the more general case. This temporary proof rule will be replaced by the proper one eventually. *) @@ -225,6 +225,14 @@ Proof. exact: uPred_primitive.discrete_valid. Qed. Lemma discrete_fun_validI {A} {B : A → ucmraT} (g : discrete_fun B) : ✓ g ⊣⊢ ∀ i, ✓ g i. Proof. exact: uPred_primitive.discrete_fun_validI. Qed. +(** This is really just a special case of an entailment +between two [siProp], but we do not have the infrastructure +to express the more general case. This temporary proof rule will +be replaced by the proper one eventually. *) +Lemma valid_entails {A B : cmraT} (a : A) (b : B) : + (∀ n, ✓{n} a → ✓{n} b) → ✓ a ⊢ ✓ b. +Proof. exact: uPred_primitive.valid_entails. Qed. + (** Consistency/soundness statement *) Lemma pure_soundness φ : (⊢@{uPredI M} ⌜ φ âŒ) → φ. Proof. apply pure_soundness. Qed. diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index 9c836e3b34f4242cf4e13f73272aa4b3e8b389e2..44a5aaea896733e684c0abb9aff30f9822163995 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -720,7 +720,7 @@ Proof. unseal=> ?. split=> n x ?. by apply (discrete_iff n). Qed. -(* This is really just a special case of an entailment +(** This is really just a special case of an entailment between two [siProp], but we do not have the infrastructure to express the more general case. This temporary proof rule will be replaced by the proper one eventually. *) @@ -818,6 +818,14 @@ Proof. unseal; split=> n x _. by rewrite /= -cmra_discrete_valid_iff. Qed. Lemma discrete_fun_validI {A} {B : A → ucmraT} (g : discrete_fun B) : ✓ g ⊣⊢ ∀ i, ✓ g i. Proof. by unseal. Qed. +(** This is really just a special case of an entailment +between two [siProp], but we do not have the infrastructure +to express the more general case. This temporary proof rule will +be replaced by the proper one eventually. *) +Lemma valid_entails {A B : cmraT} (a : A) (b : B) : + (∀ n, ✓{n} a → ✓{n} b) → ✓ a ⊢ ✓ b. +Proof. unseal=> Hval. split=>n x ?. apply Hval. Qed. + (** Consistency/soundness statement *) (** The lemmas [pure_soundness] and [internal_eq_soundness] should become an instance of [siProp] soundness in the future. *)