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

Merge branch 'master' of gitlab.mpi-sws.org:FP/iris-coq

parents 6c6f5755 126d54c3
No related branches found
No related tags found
No related merge requests found
Showing with 69 additions and 26 deletions
......@@ -116,15 +116,11 @@ tests/list_reverse.v
tests/tree_sum.v
tests/counter.v
proofmode/coq_tactics.v
proofmode/pviewshifts.v
proofmode/environments.v
proofmode/intro_patterns.v
proofmode/spec_patterns.v
proofmode/sel_patterns.v
proofmode/tactics.v
proofmode/notation.v
proofmode/invariants.v
proofmode/weakestpre.v
proofmode/ghost_ownership.v
proofmode/classes.v
proofmode/class_instances.v
From iris.algebra Require Export excl local_updates.
From iris.algebra Require Import upred updates.
From iris.proofmode Require Import class_instances.
Local Arguments valid _ _ !_ /.
Local Arguments validN _ _ _ !_ /.
......@@ -224,6 +225,14 @@ End cmra.
Arguments authR : clear implicits.
Arguments authUR : clear implicits.
(* Proof mode class instances *)
Instance from_op_auth_frag {A : ucmraT} (a b1 b2 : A) :
FromOp a b1 b2 FromOp ( a) ( b1) ( b2).
Proof. done. Qed.
Instance into_op_auth_frag {A : ucmraT} (a b1 b2 : A) :
IntoOp a b1 b2 IntoOp ( a) ( b1) ( b2).
Proof. done. Qed.
(* Functor *)
Definition auth_map {A B} (f : A B) (x : auth A) : auth B :=
Auth (excl_map f <$> authoritative x) (f (auth_own x)).
......
......@@ -1208,8 +1208,8 @@ Section option.
Lemma Some_included x y : Some x Some y x y x y.
Proof. rewrite option_included; naive_solver. Qed.
Lemma Some_included' `{CMRATotal A} x y : Some x Some y x y x y.
Proof. rewrite Some_included; naive_solver. Qed.
Lemma Some_included' `{CMRATotal A} x y : Some x Some y x y.
Proof. rewrite Some_included. split. by intros [->|?]. eauto. Qed.
Lemma is_Some_included mx my : mx my is_Some mx is_Some my.
Proof. rewrite -!not_eq_None_Some option_included. naive_solver. Qed.
End option.
......
......@@ -450,6 +450,36 @@ Proof.
by apply prodC_map_ne; apply cFunctor_contractive.
Qed.
Instance compose_ne {A} {B B' : cofeT} (f : B -n> B') n :
Proper (dist n ==> dist n) (compose f : (A -c> B) A -c> B').
Proof. intros g g' Hf x; simpl. by rewrite (Hf x). Qed.
Definition cofe_funC_map {A B B'} (f : B -n> B') : (A -c> B) -n> (A -c> B') :=
@CofeMor (_ -c> _) (_ -c> _) (compose f) _.
Instance cofe_funC_map_ne {A B B'} n :
Proper (dist n ==> dist n) (@cofe_funC_map A B B').
Proof. intros f f' Hf g x. apply Hf. Qed.
Program Definition cofe_funCF (T : Type) (F : cFunctor) : cFunctor := {|
cFunctor_car A B := cofe_funC T (cFunctor_car F A B);
cFunctor_map A1 A2 B1 B2 fg := cofe_funC_map (cFunctor_map F fg)
|}.
Next Obligation.
intros ?? A1 A2 B1 B2 n ???; by apply cofe_funC_map_ne; apply cFunctor_ne.
Qed.
Next Obligation. intros F1 F2 A B ??. by rewrite /= /compose /= !cFunctor_id. Qed.
Next Obligation.
intros T F A1 A2 A3 B1 B2 B3 f g f' g' ??; simpl.
by rewrite !cFunctor_compose.
Qed.
Instance cofe_funCF_contractive (T : Type) (F : cFunctor) :
cFunctorContractive F cFunctorContractive (cofe_funCF T F).
Proof.
intros ?? A1 A2 B1 B2 n ???;
by apply cofe_funC_map_ne; apply cFunctor_contractive.
Qed.
Program Definition cofe_morCF (F1 F2 : cFunctor) : cFunctor := {|
cFunctor_car A B := cFunctor_car F1 B A -n> cFunctor_car F2 A B;
cFunctor_map A1 A2 B1 B2 fg :=
......@@ -759,6 +789,7 @@ Qed.
(** Notation for writing functors *)
Notation "∙" := idCF : cFunctor_scope.
Notation "T -c> F" := (cofe_funCF T%type F%CF) : cFunctor_scope.
Notation "F1 -n> F2" := (cofe_morCF F1%CF F2%CF) : cFunctor_scope.
Notation "F1 * F2" := (prodCF F1%CF F2%CF) : cFunctor_scope.
Notation "F1 + F2" := (sumCF F1%CF F2%CF) : cFunctor_scope.
......
......@@ -35,6 +35,3 @@ Global Instance frac_full_exclusive : Exclusive 1%Qp.
Proof.
move=> y /Qcle_not_lt [] /=. by rewrite -{1}(Qcplus_0_r 1) -Qcplus_lt_mono_l.
Qed.
Lemma invalid_plus_q: (q: Qp), ¬ (1 + q)%Qp.
Proof. intros q H. by apply (Qp_ge_1 q). Qed.
......@@ -1361,6 +1361,12 @@ Lemma option_validI {A : cmraT} (mx : option A) :
mx ⊣⊢ match mx with Some x => x | None => True end.
Proof. uPred.unseal. by destruct mx. Qed.
(* Functions *)
Lemma cofe_funC_equivI {A B} (f g : A -c> B) : f g ⊣⊢ x, f x g x.
Proof. by uPred.unseal. Qed.
Lemma cofe_moreC_equivI {A B : cofeT} (f g : A -n> B) : f g ⊣⊢ x, f x g x.
Proof. by uPred.unseal. Qed.
(* Timeless instances *)
Global Instance pure_timeless φ : TimelessP ( φ : uPred M)%I.
Proof.
......
......@@ -2,7 +2,7 @@ From iris.program_logic Require Export weakestpre adequacy.
From iris.heap_lang Require Export heap.
From iris.program_logic Require Import auth ownership.
From iris.heap_lang Require Import proofmode notation.
From iris.proofmode Require Import tactics weakestpre.
From iris.proofmode Require Import tactics.
Definition heapΣ : gFunctors := #[authΣ heapUR; irisΣ heap_lang].
......
......@@ -2,7 +2,7 @@ From iris.heap_lang Require Export lifting.
From iris.algebra Require Import upred_big_op gmap frac dec_agree.
From iris.program_logic Require Export invariants ghost_ownership.
From iris.program_logic Require Import ownership auth.
From iris.proofmode Require Import weakestpre.
From iris.proofmode Require Import tactics.
Import uPred.
(* TODO: The entire construction could be generalized to arbitrary languages that have
a finmap as their state. Or maybe even beyond "as their state", i.e. arbitrary
......
From iris.program_logic Require Export weakestpre.
From iris.heap_lang Require Export lang.
From iris.proofmode Require Import invariants tactics.
From iris.proofmode Require Import tactics.
From iris.program_logic Require Import auth.
From iris.heap_lang Require Import proofmode notation.
......
From iris.program_logic Require Export weakestpre.
From iris.heap_lang Require Export lang.
From iris.proofmode Require Import invariants tactics.
From iris.proofmode Require Import tactics.
From iris.heap_lang Require Import proofmode notation.
From iris.algebra Require Import excl.
......
From iris.program_logic Require Export weakestpre.
From iris.heap_lang Require Export lang.
From iris.proofmode Require Import invariants tactics.
From iris.proofmode Require Import tactics.
From iris.heap_lang Require Import proofmode notation.
From iris.algebra Require Import excl.
From iris.heap_lang.lib Require Import lock.
......
From iris.program_logic Require Export weakestpre.
From iris.heap_lang Require Export lang.
From iris.proofmode Require Import invariants.
From iris.proofmode Require Import tactics.
From iris.heap_lang Require Import proofmode notation.
From iris.algebra Require Import auth gset.
From iris.heap_lang.lib Require Export lock.
......
......@@ -2,7 +2,7 @@ From iris.program_logic Require Export weakestpre.
From iris.program_logic Require Import ownership ectx_lifting. (* for ownP *)
From iris.heap_lang Require Export lang.
From iris.heap_lang Require Import tactics.
From iris.proofmode Require Import weakestpre.
From iris.proofmode Require Import tactics.
From iris.prelude Require Import fin_maps.
Import uPred.
......
From iris.program_logic Require Export weakestpre.
From iris.proofmode Require Import coq_tactics.
From iris.proofmode Require Export weakestpre.
From iris.proofmode Require Export tactics.
From iris.heap_lang Require Export wp_tactics heap.
Import uPred.
......
......@@ -567,9 +567,12 @@ Proof.
apply Qp_eq; simpl. ring.
Qed.
Lemma Qp_ge_1 (q: Qp): ¬ ((1 + q)%Qp 1%Qp)%Qc.
Lemma Qp_not_plus_q_ge_1 (q: Qp): ¬ ((1 + q)%Qp 1%Qp)%Qc.
Proof.
intros Hle.
apply (Qcplus_le_mono_l q 0 1) in Hle.
apply Qcle_ngt in Hle. by destruct q.
apply Qcle_ngt in Hle. apply Hle, Qp_prf.
Qed.
Lemma Qp_ge_0 (q: Qp): (0 q)%Qc.
Proof. apply Qclt_le_weak, Qp_prf. Qed.
From iris.program_logic Require Export weakestpre.
From iris.algebra Require Import gmap auth agree gset coPset upred_big_op.
From iris.program_logic Require Import ownership.
From iris.proofmode Require Import tactics weakestpre.
From iris.proofmode Require Import tactics.
Import uPred.
Record adequate {Λ} (e1 : expr Λ) (σ1 : state Λ) (φ : val Λ Prop) := {
......
From iris.program_logic Require Export pviewshifts.
From iris.program_logic Require Export invariants.
From iris.algebra Require Export auth.
From iris.algebra Require Import gmap.
From iris.proofmode Require Import invariants.
From iris.proofmode Require Import tactics.
Import uPred.
(* The CMRA we need. *)
......
From iris.program_logic Require Export pviewshifts.
From iris.program_logic Require Export invariants.
From iris.algebra Require Import auth gmap agree upred_big_op.
From iris.proofmode Require Import tactics invariants.
From iris.proofmode Require Import tactics.
Import uPred.
(** The CMRAs we need. *)
......
From iris.program_logic Require Export invariants.
From iris.algebra Require Export frac.
From iris.proofmode Require Import invariants tactics.
From iris.proofmode Require Import tactics.
Import uPred.
Class cinvG Σ := cinv_inG :> inG Σ fracR.
......
(** Some derived lemmas for ectx-based languages *)
From iris.program_logic Require Export ectx_language weakestpre lifting.
From iris.program_logic Require Import ownership.
From iris.proofmode Require Import weakestpre.
From iris.proofmode Require Import tactics.
Section wp.
Context {expr val ectx state} {Λ : EctxLanguage expr val ectx state}.
......
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