Skip to content
Snippets Groups Projects
Commit ebcb867c authored by Robbert Krebbers's avatar Robbert Krebbers Committed by Ralf Jung
Browse files

Initial Iris commit

parent 02f213ce
No related branches found
No related tags found
No related merge requests found
......@@ -23,6 +23,7 @@ Arguments compose _ _ _ _ _ _ /.
Arguments flip _ _ _ _ _ _ /.
Arguments const _ _ _ _ /.
Typeclasses Transparent id compose flip const.
Instance: Params (@pair) 2.
(** Change [True] and [False] into notations in order to enable overloading.
We will use this in the file [assertions] to give [True] and [False] a
......@@ -792,6 +793,11 @@ Instance pointwise_transitive {A} `{R : relation B} :
Transitive R Transitive (pointwise_relation A R) | 9.
Proof. firstorder. Qed.
(** ** Unit *)
Instance unit_equiv : Equiv unit := λ _ _, True.
Instance unit_equivalence : Equivalence (@equiv unit _).
Proof. repeat split. Qed.
(** ** Products *)
Instance prod_map_injective {A A' B B'} (f : A A') (g : B B') :
Injective (=) (=) f Injective (=) (=) g
......@@ -825,6 +831,15 @@ Section prod_relation.
Proof. firstorder eauto. Qed.
End prod_relation.
Instance prod_equiv `{Equiv A,Equiv B} : Equiv (A * B) := prod_relation () ().
Instance pair_proper `{Equiv A, Equiv B} :
Proper (() ==> () ==> ()) (@pair A B) | 0 := _.
Instance fst_proper `{Equiv A, Equiv B} :
Proper (() ==> ()) (@fst A B) | 0 := _.
Instance snd_proper `{Equiv A, Equiv B} :
Proper (() ==> ()) (@snd A B) | 0 := _.
Typeclasses Opaque prod_equiv.
(** ** Other *)
Lemma or_l P Q : ¬Q P Q P.
Proof. tauto. Qed.
......
......@@ -249,7 +249,7 @@ Ltac unfold_elem_of :=
For goals that do not involve [≡], [⊆], [map], or quantifiers this tactic is
generally powerful enough. This tactic either fails or proves the goal. *)
Tactic Notation "solve_elem_of" tactic3(tac) :=
simpl in *;
setoid_subst;
decompose_empty;
unfold_elem_of;
solve [intuition (simplify_equality; tac)].
......@@ -261,7 +261,7 @@ fails or loops on very small goals generated by [solve_elem_of] already. We
use the [naive_solver] tactic as a substitute. This tactic either fails or
proves the goal. *)
Tactic Notation "esolve_elem_of" tactic3(tac) :=
simpl in *;
setoid_subst;
decompose_empty;
unfold_elem_of;
naive_solver tac.
......@@ -273,6 +273,11 @@ Section collection.
Global Instance: Lattice C.
Proof. split. apply _. firstorder auto. solve_elem_of. Qed.
Global Instance difference_proper : Proper (() ==> () ==> ()) ().
Proof.
intros X1 X2 HX Y1 Y2 HY; apply elem_of_equiv; intros x.
by rewrite !elem_of_difference, HX, HY.
Qed.
Lemma intersection_singletons x : {[x]} {[x]} {[x]}.
Proof. esolve_elem_of. Qed.
Lemma difference_twice X Y : (X Y) Y X Y.
......@@ -283,6 +288,8 @@ Section collection.
Proof. esolve_elem_of. Qed.
Lemma difference_union_distr_l X Y Z : (X Y) Z X Z Y Z.
Proof. esolve_elem_of. Qed.
Lemma difference_union_distr_r X Y Z : Z (X Y) (Z X) (Z Y).
Proof. esolve_elem_of. Qed.
Lemma difference_intersection_distr_l X Y Z : (X Y) Z X Z Y Z.
Proof. esolve_elem_of. Qed.
......@@ -298,6 +305,8 @@ Section collection.
Proof. unfold_leibniz. apply difference_diag. Qed.
Lemma difference_union_distr_l_L X Y Z : (X Y) Z = X Z Y Z.
Proof. unfold_leibniz. apply difference_union_distr_l. Qed.
Lemma difference_union_distr_r_L X Y Z : Z (X Y) = (Z X) (Z Y).
Proof. unfold_leibniz. apply difference_union_distr_r. Qed.
Lemma difference_intersection_distr_l_L X Y Z :
(X Y) Z = X Z Y Z.
Proof. unfold_leibniz. apply difference_intersection_distr_l. Qed.
......@@ -518,16 +527,13 @@ Section collection_monad.
Global Instance collection_fmap_proper {A B} (f : A B) :
Proper (() ==> ()) (fmap f).
Proof. intros X Y E. esolve_elem_of. Qed.
Global Instance collection_ret_proper {A} :
Proper ((=) ==> ()) (@mret M _ A).
Proof. intros X Y E. esolve_elem_of. Qed.
Proof. intros X Y [??]; split; esolve_elem_of. Qed.
Global Instance collection_bind_proper {A B} (f : A M B) :
Proper (() ==> ()) (mbind f).
Proof. intros X Y E. esolve_elem_of. Qed.
Proof. intros X Y [??]; split; esolve_elem_of. Qed.
Global Instance collection_join_proper {A} :
Proper (() ==> ()) (@mjoin M _ A).
Proof. intros X Y E. esolve_elem_of. Qed.
Proof. intros X Y [??]; split; esolve_elem_of. Qed.
Lemma collection_bind_singleton {A B} (f : A M B) x : {[ x ]} ≫= f f x.
Proof. esolve_elem_of. Qed.
......
......@@ -3,7 +3,7 @@
(** This file collects definitions and theorems on finite collections. Most
importantly, it implements a fold and size function and some useful induction
principles on finite collections . *)
Require Import Permutation ars listset.
Require Import Permutation relations listset.
Require Export numbers collections.
Instance collection_size `{Elements A C} : Size C := length elements.
......
......@@ -77,15 +77,15 @@ Proof. rewrite not_elem_of_dom. apply delete_partial_alter. Qed.
Lemma delete_insert_dom {A} (m : M A) i x :
i dom D m delete i (<[i:=x]>m) = m.
Proof. rewrite not_elem_of_dom. apply delete_insert. Qed.
Lemma map_disjoint_dom {A} (m1 m2 : M A) : m1 m2 dom D m1 dom D m2 ∅.
Lemma map_disjoint_dom {A} (m1 m2 : M A) : m1 m2 dom D m1 dom D m2 ∅.
Proof.
rewrite map_disjoint_spec, elem_of_equiv_empty.
setoid_rewrite elem_of_intersection.
setoid_rewrite elem_of_dom. unfold is_Some. naive_solver.
Qed.
Lemma map_disjoint_dom_1 {A} (m1 m2 : M A) : m1 m2 dom D m1 dom D m2 ∅.
Lemma map_disjoint_dom_1 {A} (m1 m2 : M A) : m1 m2 dom D m1 dom D m2 ∅.
Proof. apply map_disjoint_dom. Qed.
Lemma map_disjoint_dom_2 {A} (m1 m2 : M A) : dom D m1 dom D m2 m1 m2.
Lemma map_disjoint_dom_2 {A} (m1 m2 : M A) : dom D m1 dom D m2 m1 m2.
Proof. apply map_disjoint_dom. Qed.
Lemma dom_union {A} (m1 m2 : M A) : dom D (m1 m2) dom D m1 dom D m2.
Proof.
......
This diff is collapsed.
......@@ -74,6 +74,40 @@ Proof. destruct x; unfold is_Some; naive_solver. Qed.
Lemma not_eq_None_Some `(x : option A) : x None is_Some x.
Proof. rewrite eq_None_not_Some. split. apply dec_stable. tauto. Qed.
(** Lifting a relation point-wise to option *)
Inductive option_Forall2 {A B} (P: A B Prop) : option A option B Prop :=
| Some_Forall2 x y : P x y option_Forall2 P (Some x) (Some y)
| None_Forall2 : option_Forall2 P None None.
Definition option_relation {A B} (R: A B Prop) (P: A Prop) (Q: B Prop)
(mx : option A) (my : option B) : Prop :=
match mx, my with
| Some x, Some y => R x y
| Some x, None => P x
| None, Some y => Q y
| None, None => True
end.
(** Setoids *)
Section setoids.
Context `{Equiv A}.
Global Instance option_equiv : Equiv (option A) := option_Forall2 ().
Global Instance option_equivalence `{!Equivalence (() : relation A)} :
Equivalence (() : relation (option A)).
Proof.
split.
* by intros []; constructor.
* by destruct 1; constructor.
* destruct 1; inversion 1; constructor; etransitivity; eauto.
Qed.
Global Instance Some_proper : Proper (() ==> ()) (@Some A).
Proof. by constructor. Qed.
Global Instance option_leibniz `{!LeibnizEquiv A} : LeibnizEquiv (option A).
Proof.
intros x y; split; [destruct 1; fold_leibniz; congruence|].
by intros <-; destruct x; constructor; fold_leibniz.
Qed.
End setoids.
(** Equality on [option] is decidable. *)
Instance option_eq_None_dec {A} (x : option A) : Decision (x = None) :=
match x with Some _ => right (Some_ne_None _) | None => left eq_refl end.
......
......@@ -8,7 +8,7 @@ Require Export
option
vector
numbers
ars
relations
collections
fin_collections
listset
......
(* Copyright (c) 2012-2015, Robbert Krebbers. *)
(* This file is distributed under the terms of the BSD license. *)
Require Export numbers option.
Require Import Ascii String ars.
Require Import Ascii String relations.
Infix "+:+" := String.append (at level 60, right associativity) : C_scope.
Arguments String.append _ _ : simpl never.
......
......@@ -22,6 +22,11 @@ Section definitions.
| rtc_refl x : rtc x x
| rtc_l x y z : R x y rtc y z rtc x z.
(** The reflexive transitive closure for setoids. *)
Inductive rtcS `{Equiv A} : relation A :=
| rtcS_refl x y : x y rtcS x y
| rtcS_l x y z : R x y rtcS y z rtcS x z.
(** Reductions of exactly [n] steps. *)
Inductive nsteps : nat relation A :=
| nsteps_O x : nsteps 0 x x
......
(* Copyright (c) 2012-2015, Robbert Krebbers. *)
(* This file is distributed under the terms of the BSD license. *)
(** This file implements sets as functions into Prop. *)
Require Export prelude.
Record set (A : Type) : Type := mkSet { set_car : A Prop }.
Arguments mkSet {_} _.
Arguments set_car {_} _ _.
Definition set_all {A} : set A := mkSet (λ _, True).
Instance set_empty {A} : Empty (set A) := mkSet (λ _, False).
Instance set_singleton {A} : Singleton A (set A) := λ x, mkSet (x =).
Instance set_elem_of {A} : ElemOf A (set A) := λ x X, set_car X x.
Instance set_union {A} : Union (set A) := λ X1 X2, mkSet (λ x, x X1 x X2).
Instance set_intersection {A} : Intersection (set A) := λ X1 X2,
mkSet (λ x, x X1 x X2).
Instance set_difference {A} : Difference (set A) := λ X1 X2,
mkSet (λ x, x X1 x X2).
Instance set_collection : Collection A (set A).
Proof. by split; [split | |]; repeat intro. Qed.
Instance set_ret : MRet set := λ A (x : A), {[ x ]}.
Instance set_bind : MBind set := λ A B (f : A set B) (X : set A),
mkSet (λ b, a, b f a a X).
Instance set_fmap : FMap set := λ A B (f : A B) (X : set A),
mkSet (λ b, a, b = f a a X).
Instance set_join : MJoin set := λ A (XX : set (set A)),
mkSet (λ a, X, a X X XX).
Instance set_collection_monad : CollectionMonad set.
Proof. by split; try apply _. Qed.
Global Opaque set_union set_intersection.
\ No newline at end of file
......@@ -207,6 +207,26 @@ Ltac f_lia :=
end.
Ltac f_lia' := csimpl in *; f_lia.
Ltac setoid_subst_l x :=
match goal with
| H : x ?y |- _ =>
is_var x;
try match y with x _ => fail 2 end;
repeat match goal with
| |- context [ x ] => setoid_rewrite H
| H' : context [ x ] |- _ =>
try match H' with H => fail 2 end;
setoid_rewrite H in H'
end;
clear H
end.
Ltac setoid_subst :=
repeat match goal with
| _ => progress simplify_equality'
| H : ?x _ |- _ => setoid_subst_l x
| H : _ ?x |- _ => symmetry in H; setoid_subst_l x
end.
(** Given a tactic [tac2] generating a list of terms, [iter tac1 tac2]
runs [tac x] for each element [x] until [tac x] succeeds. If it does not
suceed for any element of the generated list, the whole tactic wil fail. *)
......
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