Skip to content
Snippets Groups Projects
base.v 70.3 KiB
Newer Older
(** This file collects type class interfaces, notations, and general theorems
that are used throughout the whole development. Most importantly it contains
abstract interfaces for ordered structures, sets, and various other data
(* We want to ensure that [le] and [lt] refer to operations on [nat].
These two functions being defined both in [Coq.Bool] and in [Coq.Peano],
we must export [Coq.Peano] later than any export of [Coq.Bool]. *)
(* We also want to ensure that notations from [Coq.Utf8] take precedence
over the ones of [Coq.Peano] (see Coq PR#12950), so we import [Utf8] last. *)
From Coq Require Export Morphisms RelationClasses List Bool Setoid Peano Utf8.
From Coq Require Import Permutation.
Export ListNotations.
From Coq.Program Require Export Basics Syntax.
From stdpp Require Import options.
Michael Sammler's avatar
Michael Sammler committed
(** This notation is necessary to prevent [length] from being printed
as [strings.length] if strings.v is imported and later base.v. See
also strings.v and
https://gitlab.mpi-sws.org/iris/stdpp/-/merge_requests/144 and
https://gitlab.mpi-sws.org/iris/stdpp/-/merge_requests/129. *)
Notation length := Datatypes.length.

Ralf Jung's avatar
Ralf Jung committed
(** * Enable implicit generalization. *)
(** This option enables implicit generalization in arguments of the form
   [`{...}] (i.e., anonymous arguments).  Unfortunately, it also enables
   implicit generalization in [Instance].  We think that the fact that both
   behaviors are coupled together is a [bug in
   Coq](https://github.com/coq/coq/issues/6030). *)
Global Generalizable All Variables.

(** * Tweak program *)
(** 1. Since we only use Program to solve logical side-conditions, they should
always be made Opaque, otherwise we end up with performance problems due to
Coq blindly unfolding them.

Note that in most cases we use [Next Obligation. (* ... *) Qed.], for which
this option does not matter. However, sometimes we write things like
[Solve Obligations with naive_solver (* ... *)], and then the obligations
should surely be opaque. *)
Global Unset Transparent Obligations.

(** 2. Do not let Program automatically simplify obligations. The default
obligation tactic is [Tactics.program_simpl], which, among other things,
introduces all variables and gives them fresh names. As such, it becomes
impossible to refer to hypotheses in a robust way. *)
Obligation Tactic := idtac.

(** 3. Hide obligations from the results of the [Search] commands. *)
Add Search Blacklist "_obligation_".
Robbert Krebbers's avatar
Robbert Krebbers committed

(** * Sealing off definitions *)
Section seal.
  Local Set Primitive Projections.
  Record seal {A} (f : A) := { unseal : A; seal_eq : unseal = f }.
End seal.
Global Arguments unseal {_ _} _ : assert.
Global Arguments seal_eq {_ _} _ : assert.
(** * Non-backtracking type classes *)
(** The type class [TCNoBackTrack P] can be used to establish [P] without ever
backtracking on the instance of [P] that has been found. Backtracking may
normally happen when [P] contains evars that could be instanciated in different
ways depending on which instance is picked, and type class search somewhere else
depends on this evar.

The proper way of handling this would be by setting Coq's option
`Typeclasses Unique Instances`. However, this option seems to be broken, see Coq
issue #6714.

See https://gitlab.mpi-sws.org/FP/iris-coq/merge_requests/112 for a rationale
of this type class. *)
Class TCNoBackTrack (P : Prop) := { tc_no_backtrack : P }.
Global Hint Extern 0 (TCNoBackTrack _) => constructor; apply _ : typeclass_instances.
(* A conditional at the type class level. Note that [TCIf P Q R] is not the same
as [TCOr (TCAnd P Q) R]: the latter will backtrack to [R] if it fails to
Paolo G. Giarrusso's avatar
Paolo G. Giarrusso committed
establish [Q], i.e. does not have the behavior of a conditional. Furthermore,
note that [TCOr (TCAnd P Q) (TCAnd (TCNot P) R)] would not work; we generally
Robbert Krebbers's avatar
Robbert Krebbers committed
would not be able to prove the negation of [P]. *)
Robbert Krebbers's avatar
Robbert Krebbers committed
Inductive TCIf (P Q R : Prop) : Prop :=
  | TCIf_true : P  Q  TCIf P Q R
  | TCIf_false : R  TCIf P Q R.
Existing Class TCIf.

Global Hint Extern 0 (TCIf _ _ _) =>
  first [apply TCIf_true; [apply _|]
        |apply TCIf_false] : typeclass_instances.

(** * Typeclass opaque definitions *)
Ralf Jung's avatar
Ralf Jung committed
(** The constant [tc_opaque] is used to make definitions opaque for just type
class search. Note that [simpl] is set up to always unfold [tc_opaque]. *)
Definition tc_opaque {A} (x : A) : A := x.
Typeclasses Opaque tc_opaque.
Global Arguments tc_opaque {_} _ /.
Ralf Jung's avatar
Ralf Jung committed
(** Below we define type class versions of the common logical operators. It is
important to note that we duplicate the definitions, and do not declare the
existing logical operators as type classes. That is, we do not say:

  Existing Class or.
  Existing Class and.

If we could define the existing logical operators as classes, there is no way
of disambiguating whether a premise of a lemma should be solved by type class
resolution or not.

These classes are useful for two purposes: writing complicated type class
premises in a more concise way, and for efficiency. For example, using the [Or]
class, instead of defining two instances [P → Q1 → R] and [P → Q2 → R] we could
have one instance [P → Or Q1 Q2 → R]. When we declare the instance that way, we
avoid the need to derive [P] twice. *)
Inductive TCOr (P1 P2 : Prop) : Prop :=
  | TCOr_l : P1  TCOr P1 P2
  | TCOr_r : P2  TCOr P1 P2.
Existing Class TCOr.
Global Existing Instance TCOr_l | 9.
Global Existing Instance TCOr_r | 10.
Global Hint Mode TCOr ! ! : typeclass_instances.
Inductive TCAnd (P1 P2 : Prop) : Prop := TCAnd_intro : P1  P2  TCAnd P1 P2.
Global Existing Instance TCAnd_intro.
Global Hint Mode TCAnd ! ! : typeclass_instances.
Inductive TCTrue : Prop := TCTrue_intro : TCTrue.
Existing Class TCTrue.
Global Existing Instance TCTrue_intro.
Michael Sammler's avatar
Michael Sammler committed
(** The class [TCFalse] is not stricly necessary as one could also use
[False]. However, users might expect that TCFalse exists and if it
does not, it can cause hard to diagnose bugs due to automatic
generalization. *)
Inductive TCFalse : Prop :=.
Existing Class TCFalse.

(** The class [TCUnless] can be used to check that search for [P]
fails. This is useful as a guard for certain instances together with
classes like [TCFastDone] (see [tactics.v]) to prevent infinite loops
(e.g. when saturating the context). *)
Notation TCUnless P := (TCIf P TCFalse TCTrue).

Inductive TCForall {A} (P : A  Prop) : list A  Prop :=
  | TCForall_nil : TCForall P []
  | TCForall_cons x xs : P x  TCForall P xs  TCForall P (x :: xs).
Existing Class TCForall.
Global Existing Instance TCForall_nil.
Global Existing Instance TCForall_cons.
Global Hint Mode TCForall ! ! ! : typeclass_instances.
(** The class [TCForall2 P l k] is commonly used to transform an input list [l]
into an output list [k], or the converse. Therefore there are two modes, either
[l] input and [k] output, or [k] input and [l] input. *)
Robbert Krebbers's avatar
Robbert Krebbers committed
Inductive TCForall2 {A B} (P : A  B  Prop) : list A  list B  Prop :=
  | TCForall2_nil : TCForall2 P [] []
  | TCForall2_cons x y xs ys :
     P x y  TCForall2 P xs ys  TCForall2 P (x :: xs) (y :: ys).
Existing Class TCForall2.
Global Existing Instance TCForall2_nil.
Global Existing Instance TCForall2_cons.
Global Hint Mode TCForall2 ! ! ! ! - : typeclass_instances.
Global Hint Mode TCForall2 ! ! ! - ! : typeclass_instances.
Michael Sammler's avatar
Michael Sammler committed
Inductive TCExists {A} (P : A  Prop) : list A  Prop :=
  | TCExists_cons_hd x l : P x  TCExists P (x :: l)
  | TCExists_cons_tl x l: TCExists P l  TCExists P (x :: l).
Existing Class TCExists.
Global Existing Instance TCExists_cons_hd | 10.
Global Existing Instance TCExists_cons_tl | 20.
Michael Sammler's avatar
Michael Sammler committed
Global Hint Mode TCExists ! ! ! : typeclass_instances.

Inductive TCElemOf {A} (x : A) : list A  Prop :=
  | TCElemOf_here xs : TCElemOf x (x :: xs)
  | TCElemOf_further y xs : TCElemOf x xs  TCElemOf x (y :: xs).
Existing Class TCElemOf.
Global Existing Instance TCElemOf_here.
Global Existing Instance TCElemOf_further.
Global Hint Mode TCElemOf ! ! ! : typeclass_instances.
(** We declare both arguments [x] and [y] of [TCEq x y] as outputs, which means
[TCEq] can also be used to unify evars. This is harmless: since the only
instance of [TCEq] is [TCEq_refl] below, it can never cause loops. See
https://gitlab.mpi-sws.org/iris/iris/merge_requests/391 for a use case. *)
Inductive TCEq {A} (x : A) : A  Prop := TCEq_refl : TCEq x x.
Existing Class TCEq.
Global Existing Instance TCEq_refl.
Global Hint Mode TCEq ! - - : typeclass_instances.
Lemma TCEq_eq {A} (x1 x2 : A) : TCEq x1 x2  x1 = x2.
Proof. split; destruct 1; reflexivity. Qed.

Inductive TCDiag {A} (C : A  Prop) : A  A  Prop :=
  | TCDiag_diag x : C x  TCDiag C x x.
Existing Class TCDiag.
Global Existing Instance TCDiag_diag.
Global Hint Mode TCDiag ! ! ! - : typeclass_instances.
Global Hint Mode TCDiag ! ! - ! : typeclass_instances.
(** Given a proposition [P] that is a type class, [tc_to_bool P] will return
[true] iff there is an instance of [P]. It is often useful in Ltac programming,
where one can do [lazymatch tc_to_bool P with true => .. | false => .. end]. *)
Definition tc_to_bool (P : Prop)
  {p : bool} `{TCIf P (TCEq p true) (TCEq p false)} : bool := p.

(** Throughout this development we use [stdpp_scope] for all general purpose
notations that do not belong to a more specific scope. *)
Declare Scope stdpp_scope.
Delimit Scope stdpp_scope with stdpp.
Global Open Scope stdpp_scope.
(** Change [True] and [False] into notations in order to enable overloading.
We will use this to give [True] and [False] a different interpretation for
embedded logics. *)
Notation "'True'" := True (format "True") : type_scope.
Notation "'False'" := False (format "False") : type_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Gregory Malecha's avatar
Gregory Malecha committed
(** Change [forall] into a notation in order to enable overloading. *)
Notation "'forall' x .. y , P" := (forall x, .. (forall y, P) ..)
Gregory Malecha's avatar
Gregory Malecha committed
  (at level 200, x binder, y binder, right associativity,
   only parsing) : type_scope.

Robbert Krebbers's avatar
Robbert Krebbers committed

(** Introduce some Haskell style like notations. *)
Notation "(=)" := eq (only parsing) : stdpp_scope.
Notation "( x =.)" := (eq x) (only parsing) : stdpp_scope.
Notation "(.= x )" := (λ y, eq y x) (only parsing) : stdpp_scope.
Notation "(≠)" := (λ x y, x  y) (only parsing) : stdpp_scope.
Notation "( x ≠.)" := (λ y, x  y) (only parsing) : stdpp_scope.
Notation "(.≠ x )" := (λ y, y  x) (only parsing) : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Infix "=@{ A }" := (@eq A)
  (at level 70, only parsing, no associativity) : stdpp_scope.
Notation "(=@{ A } )" := (@eq A) (only parsing) : stdpp_scope.
Notation "(≠@{ A } )" := (λ X Y, ¬X =@{A} Y) (only parsing) : stdpp_scope.
Notation "X ≠@{ A } Y":= (¬X =@{ A } Y)
  (at level 70, only parsing, no associativity) : stdpp_scope.
Global Hint Extern 0 (_ = _) => reflexivity : core.
Global Hint Extern 100 (_  _) => discriminate : core.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Instance:  A, PreOrder (=@{A}).
Proof. split; repeat intro; congruence. Qed.

(** ** Setoid equality *)
Ralf Jung's avatar
Ralf Jung committed
(** We define an operational type class for setoid equality, i.e., the
"canonical" equivalence for a type. The typeclass is tied to the \equiv
symbol. This is based on (Spitters/van der Weegen, 2011). *)
Class Equiv A := equiv: relation A.
Ralf Jung's avatar
Ralf Jung committed
(* No Hint Mode set because of Coq bug #14441.
Global Hint Mode Equiv ! : typeclass_instances. *)
Infix "≡" := equiv (at level 70, no associativity) : stdpp_scope.
Infix "≡@{ A }" := (@equiv A _)
  (at level 70, only parsing, no associativity) : stdpp_scope.

Notation "(≡)" := equiv (only parsing) : stdpp_scope.
Notation "( X ≡.)" := (equiv X) (only parsing) : stdpp_scope.
Notation "(.≡ X )" := (λ Y, Y  X) (only parsing) : stdpp_scope.
Notation "(≢)" := (λ X Y, ¬X  Y) (only parsing) : stdpp_scope.
Notation "X ≢ Y":= (¬X  Y) (at level 70, no associativity) : stdpp_scope.
Notation "( X ≢.)" := (λ Y, X  Y) (only parsing) : stdpp_scope.
Notation "(.≢ X )" := (λ Y, Y  X) (only parsing) : stdpp_scope.
Notation "(≡@{ A } )" := (@equiv A _) (only parsing) : stdpp_scope.
Notation "(≢@{ A } )" := (λ X Y, ¬X ≡@{A} Y) (only parsing) : stdpp_scope.
Notation "X ≢@{ A } Y":= (¬X ≡@{ A } Y)
  (at level 70, only parsing, no associativity) : stdpp_scope.
(** The type class [LeibnizEquiv] collects setoid equalities that coincide
with Leibniz equality. We provide the tactic [fold_leibniz] to transform such
setoid equalities into Leibniz equalities, and [unfold_leibniz] for the
reverse. *)
Class LeibnizEquiv A `{Equiv A} := leibniz_equiv x y : x  y  x = y.
Global Hint Mode LeibnizEquiv ! - : typeclass_instances.
Lemma leibniz_equiv_iff `{LeibnizEquiv A, !Reflexive (≡@{A})} (x y : A) :
Proof. split; [apply leibniz_equiv|]. intros ->; reflexivity. Qed.
Ltac fold_leibniz := repeat
  match goal with
    setoid_rewrite (leibniz_equiv_iff (A:=A)) in H
    setoid_rewrite (leibniz_equiv_iff (A:=A))
  end.
Ltac unfold_leibniz := repeat
  match goal with
    setoid_rewrite <-(leibniz_equiv_iff (A:=A)) in H
    setoid_rewrite <-(leibniz_equiv_iff (A:=A))
  end.

Definition equivL {A} : Equiv A := (=).

(** A [Params f n] instance forces the setoid rewriting mechanism not to
rewrite in the first [n] arguments of the function [f]. We will declare such
instances for all operational type classes in this development. *)
Global Instance: Params (@equiv) 2 := {}.

(** The following instance forces [setoid_replace] to use setoid equality
(for types that have an [Equiv] instance) rather than the standard Leibniz
equality. *)
Global Instance equiv_default_relation `{Equiv A} : DefaultRelation () | 3 := {}.
Global Hint Extern 0 (_  _) => reflexivity : core.
Global Hint Extern 0 (_  _) => symmetry; assumption : core.


(** * Type classes *)
(** ** Decidable propositions *)
(** This type class by (Spitters/van der Weegen, 2011) collects decidable
Class Decision (P : Prop) := decide : {P} + {¬P}.
Global Hint Mode Decision ! : typeclass_instances.
Global Arguments decide _ {_} : simpl never, assert.

(** Although [RelDecision R] is just [∀ x y, Decision (R x y)], we make this
an explicit class instead of a notation for two reasons:

- It allows us to control [Hint Mode] more precisely. In particular, if it were
  defined as a notation, the above [Hint Mode] for [Decision] would not prevent
  diverging instance search when looking for [RelDecision (@eq ?A)], which would
  result in it looking for [Decision (@eq ?A x y)], i.e. an instance where the
  head position of [Decision] is not en evar.
- We use it to avoid inefficient computation due to eager evaluation of
  propositions by [vm_compute]. This inefficiency arises for example if
  [(x = y) := (f x = f y)]. Since [decide (x = y)] evaluates to
  [decide (f x = f y)], this would then lead to evaluation of [f x] and [f y].
  Using the [RelDecision], the [f] is hidden under a lambda, which prevents
  unnecessary evaluation. *)
Class RelDecision {A B} (R : A  B  Prop) :=
  decide_rel x y :> Decision (R x y).
Global Hint Mode RelDecision ! ! ! : typeclass_instances.
Global Arguments decide_rel {_ _} _ {_} _ _ : simpl never, assert.

(** ** Inhabited types *)
(** This type class collects types that are inhabited. *)
Class Inhabited (A : Type) : Type := populate { inhabitant : A }.
Global Hint Mode Inhabited ! : typeclass_instances.
Global Arguments populate {_} _ : assert.

(** ** Proof irrelevant types *)
(** This type class collects types that are proof irrelevant. That means, all
elements of the type are equal. We use this notion only used for propositions,
but by universe polymorphism we can generalize it. *)
Class ProofIrrel (A : Type) : Prop := proof_irrel (x y : A) : x = y.
Global Hint Mode ProofIrrel ! : typeclass_instances.

(** ** Common properties *)
(** These operational type classes allow us to refer to common mathematical
properties in a generic way. For example, for injectivity of [(k ++.)] it
allows us to write [inj (k ++.)] instead of [app_inv_head k]. *)
Class Inj {A B} (R : relation A) (S : relation B) (f : A  B) : Prop :=
  inj x y : S (f x) (f y)  R x y.
Class Inj2 {A B C} (R1 : relation A) (R2 : relation B)
    (S : relation C) (f : A  B  C) : Prop :=
  inj2 x1 x2 y1 y2 : S (f x1 x2) (f y1 y2)  R1 x1 y1  R2 x2 y2.
Class Cancel {A B} (S : relation B) (f : A  B) (g : B  A) : Prop :=
  cancel :  x, S (f (g x)) x.
Class Surj {A B} (R : relation B) (f : A  B) :=
  surj y :  x, R (f x) y.
Class IdemP {A} (R : relation A) (f : A  A  A) : Prop :=
  idemp x : R (f x x) x.
Class Comm {A B} (R : relation A) (f : B  B  A) : Prop :=
  comm x y : R (f x y) (f y x).
Class LeftId {A} (R : relation A) (i : A) (f : A  A  A) : Prop :=
  left_id x : R (f i x) x.
Class RightId {A} (R : relation A) (i : A) (f : A  A  A) : Prop :=
  right_id x : R (f x i) x.
Class Assoc {A} (R : relation A) (f : A  A  A) : Prop :=
  assoc x y z : R (f x (f y z)) (f (f x y) z).
Class LeftAbsorb {A} (R : relation A) (i : A) (f : A  A  A) : Prop :=
  left_absorb x : R (f i x) i.
Class RightAbsorb {A} (R : relation A) (i : A) (f : A  A  A) : Prop :=
  right_absorb x : R (f x i) i.
Class AntiSymm {A} (R S : relation A) : Prop :=
  anti_symm x y : S x y  S y x  R x y.
Class Total {A} (R : relation A) := total x y : R x y  R y x.
Class Trichotomy {A} (R : relation A) :=
  trichotomy x y : R x y  x = y  R y x.
Class TrichotomyT {A} (R : relation A) :=
  trichotomyT x y : {R x y} + {x = y} + {R y x}.

Notation Involutive R f := (Cancel R f f).
Lemma involutive {A} {R : relation A} (f : A  A) `{Involutive R f} x :
  R (f (f x)) x.
Proof. auto. Qed.
Global Arguments irreflexivity {_} _ {_} _ _ : assert.
Global Arguments inj {_ _ _ _} _ {_} _ _ _ : assert.
Global Arguments inj2 {_ _ _ _ _ _} _ {_} _ _ _ _ _: assert.
Global Arguments cancel {_ _ _} _ _ {_} _ : assert.
Global Arguments surj {_ _ _} _ {_} _ : assert.
Global Arguments idemp {_ _} _ {_} _ : assert.
Global Arguments comm {_ _ _} _ {_} _ _ : assert.
Global Arguments left_id {_ _} _ _ {_} _ : assert.
Global Arguments right_id {_ _} _ _ {_} _ : assert.
Global Arguments assoc {_ _} _ {_} _ _ _ : assert.
Global Arguments left_absorb {_ _} _ _ {_} _ : assert.
Global Arguments right_absorb {_ _} _ _ {_} _ : assert.
Global Arguments anti_symm {_ _} _ {_} _ _ _ _ : assert.
Global Arguments total {_} _ {_} _ _ : assert.
Global Arguments trichotomy {_} _ {_} _ _ : assert.
Global Arguments trichotomyT {_} _ {_} _ _ : assert.

Lemma not_symmetry `{R : relation A, !Symmetric R} x y : ¬R x y  ¬R y x.
Proof. intuition. Qed.
Lemma symmetry_iff `(R : relation A) `{!Symmetric R} x y : R x y  R y x.
Proof. intuition. Qed.

Lemma not_inj `{Inj A B R R' f} x y : ¬R x y  ¬R' (f x) (f y).
Proof. intuition. Qed.
Lemma not_inj2_1 `{Inj2 A B C R R' R'' f} x1 x2 y1 y2 :
  ¬R x1 x2  ¬R'' (f x1 y1) (f x2 y2).
Proof. intros HR HR''. destruct (inj2 f x1 y1 x2 y2); auto. Qed.
Lemma not_inj2_2 `{Inj2 A B C R R' R'' f} x1 x2 y1 y2 :
  ¬R' y1 y2  ¬R'' (f x1 y1) (f x2 y2).
Proof. intros HR' HR''. destruct (inj2 f x1 y1 x2 y2); auto. Qed.

Lemma inj_iff {A B} {R : relation A} {S : relation B} (f : A  B)
  `{!Inj R S f} `{!Proper (R ==> S) f} x y : S (f x) (f y)  R x y.
Proof. firstorder. Qed.
Global Instance inj2_inj_1 `{Inj2 A B C R1 R2 R3 f} y : Inj R1 R3 (λ x, f x y).
Proof. repeat intro; edestruct (inj2 f); eauto. Qed.
Global Instance inj2_inj_2 `{Inj2 A B C R1 R2 R3 f} x : Inj R2 R3 (f x).
Proof. repeat intro; edestruct (inj2 f); eauto. Qed.

Lemma cancel_inj `{Cancel A B R1 f g, !Equivalence R1, !Proper (R2 ==> R1) f} :
  Inj R1 R2 g.
Proof.
  intros x y E. rewrite <-(cancel f g x), <-(cancel f g y), E. reflexivity.
Qed.
Lemma cancel_surj `{Cancel A B R1 f g} : Surj R1 f.
Proof. intros y. exists (g y). auto. Qed.

(** The following lemmas are specific versions of the projections of the above
type classes for Leibniz equality. These lemmas allow us to enforce Coq not to
use the setoid rewriting mechanism. *)
Lemma idemp_L {A} f `{!@IdemP A (=) f} x : f x x = x.
Proof. auto. Qed.
Lemma comm_L {A B} f `{!@Comm A B (=) f} x y : f x y = f y x.
Proof. auto. Qed.
Lemma left_id_L {A} i f `{!@LeftId A (=) i f} x : f i x = x.
Proof. auto. Qed.
Lemma right_id_L {A} i f `{!@RightId A (=) i f} x : f x i = x.
Proof. auto. Qed.
Lemma assoc_L {A} f `{!@Assoc A (=) f} x y z : f x (f y z) = f (f x y) z.
Proof. auto. Qed.
Lemma left_absorb_L {A} i f `{!@LeftAbsorb A (=) i f} x : f i x = i.
Proof. auto. Qed.
Lemma right_absorb_L {A} i f `{!@RightAbsorb A (=) i f} x : f x i = i.
Proof. auto. Qed.

(** ** Generic orders *)
(** The classes [PreOrder], [PartialOrder], and [TotalOrder] use an arbitrary
relation [R] instead of [⊆] to support multiple orders on the same type. *)
Definition strict {A} (R : relation A) : relation A := λ X Y, R X Y  ¬R Y X.
Global Instance: Params (@strict) 2 := {}.
Class PartialOrder {A} (R : relation A) : Prop := {
  partial_order_pre :> PreOrder R;
  partial_order_anti_symm :> AntiSymm (=) R
}.
Global Hint Mode PartialOrder ! ! : typeclass_instances.
Class TotalOrder {A} (R : relation A) : Prop := {
  total_order_partial :> PartialOrder R;
  total_order_trichotomy :> Trichotomy (strict R)
}.
Global Hint Mode TotalOrder ! ! : typeclass_instances.
Global Instance prop_inhabited : Inhabited Prop := populate True.
Notation "(∧)" := and (only parsing) : stdpp_scope.
Notation "( A ∧.)" := (and A) (only parsing) : stdpp_scope.
Notation "(.∧ B )" := (λ A, A  B) (only parsing) : stdpp_scope.
Notation "(∨)" := or (only parsing) : stdpp_scope.
Notation "( A ∨.)" := (or A) (only parsing) : stdpp_scope.
Notation "(.∨ B )" := (λ A, A  B) (only parsing) : stdpp_scope.
Notation "(↔)" := iff (only parsing) : stdpp_scope.
Notation "( A ↔.)" := (iff A) (only parsing) : stdpp_scope.
Notation "(.↔ B )" := (λ A, A  B) (only parsing) : stdpp_scope.
Global Hint Extern 0 (_  _) => reflexivity : core.
Global Hint Extern 0 (_  _) => symmetry; assumption : core.

Lemma or_l P Q : ¬Q  P  Q  P.
Proof. tauto. Qed.
Lemma or_r P Q : ¬P  P  Q  Q.
Proof. tauto. Qed.
Lemma and_wlog_l (P Q : Prop) : (Q  P)  Q  (P  Q).
Proof. tauto. Qed.
Lemma and_wlog_r (P Q : Prop) : P  (P  Q)  (P  Q).
Proof. tauto. Qed.
Lemma impl_transitive (P Q R : Prop) : (P  Q)  (Q  R)  (P  R).
Proof. tauto. Qed.
Lemma forall_proper {A} (P Q : A  Prop) :
  ( x, P x  Q x)  ( x, P x)  ( x, Q x).
Proof. firstorder. Qed.
Lemma exist_proper {A} (P Q : A  Prop) :
  ( x, P x  Q x)  ( x, P x)  ( x, Q x).
Proof. firstorder. Qed.
Global Instance eq_comm {A} : Comm () (=@{A}).
Global Instance flip_eq_comm {A} : Comm () (λ x y, y =@{A} x).
Global Instance iff_comm : Comm () ().
Global Instance and_comm : Comm () ().
Global Instance and_assoc : Assoc () ().
Global Instance and_idemp : IdemP () ().
Global Instance or_comm : Comm () ().
Global Instance or_assoc : Assoc () ().
Global Instance or_idemp : IdemP () ().
Global Instance True_and : LeftId () True ().
Global Instance and_True : RightId () True ().
Global Instance False_and : LeftAbsorb () False ().
Global Instance and_False : RightAbsorb () False ().
Global Instance False_or : LeftId () False ().
Global Instance or_False : RightId () False ().
Global Instance True_or : LeftAbsorb () True ().
Global Instance or_True : RightAbsorb () True ().
Global Instance True_impl : LeftId () True impl.
Proof. unfold impl. red; intuition. Qed.
Global Instance impl_True : RightAbsorb () True impl.
Proof. unfold impl. red; intuition. Qed.


(** * Common data types *)
(** ** Functions *)
Notation "(→)" := (λ A B, A  B) (only parsing) : stdpp_scope.
Notation "( A →.)" := (λ B, A  B) (only parsing) : stdpp_scope.
Notation "(.→ B )" := (λ A, A  B) (only parsing) : stdpp_scope.
Notation "t $ r" := (t r)
  (at level 65, right associativity, only parsing) : stdpp_scope.
Notation "($)" := (λ f x, f x) (only parsing) : stdpp_scope.
Notation "(.$ x )" := (λ f, f x) (only parsing) : stdpp_scope.
Infix "∘" := compose : stdpp_scope.
Notation "(∘)" := compose (only parsing) : stdpp_scope.
Notation "( f ∘.)" := (compose f) (only parsing) : stdpp_scope.
Notation "(.∘ f )" := (λ g, compose g f) (only parsing) : stdpp_scope.
Global Instance impl_inhabited {A} `{Inhabited B} : Inhabited (A  B) :=
  populate (λ _, inhabitant).

(** Ensure that [simpl] unfolds [id], [compose], and [flip] when fully
applied. *)
Global Arguments id _ _ / : assert.
Global Arguments compose _ _ _ _ _ _ / : assert.
Global Arguments flip _ _ _ _ _ _ / : assert.
Global Arguments const _ _ _ _ / : assert.
Typeclasses Transparent id compose flip const.

Definition fun_map {A A' B B'} (f: A'  A) (g: B  B') (h : A  B) : A'  B' :=
  g  h  f.

Global Instance const_proper `{R1 : relation A, R2 : relation B} (x : B) :
  Reflexive R2  Proper (R1 ==> R2) (λ _, x).
Proof. intros ? y1 y2; reflexivity. Qed.

Global Instance id_inj {A} : Inj (=) (=) (@id A).
Global Instance compose_inj {A B C} R1 R2 R3 (f : A  B) (g : B  C) :
  Inj R1 R2 f  Inj R2 R3 g  Inj R1 R3 (g  f).
Proof. red; intuition. Qed.

Global Instance id_surj {A} : Surj (=) (@id A).
Proof. intros y; exists y; reflexivity. Qed.
Global Instance compose_surj {A B C} R (f : A  B) (g : B  C) :
  Surj (=) f  Surj R g  Surj R (g  f).
Proof.
  intros ?? x. unfold compose. destruct (surj g x) as [y ?].
  destruct (surj f y) as [z ?]. exists z. congruence.
Qed.

Global Instance id_comm {A B} (x : B) : Comm (=) (λ _ _ : A, x).
Proof. intros ?; reflexivity. Qed.
Global Instance id_assoc {A} (x : A) : Assoc (=) (λ _ _ : A, x).
Proof. intros ???; reflexivity. Qed.
Global Instance const1_assoc {A} : Assoc (=) (λ x _ : A, x).
Proof. intros ???; reflexivity. Qed.
Global Instance const2_assoc {A} : Assoc (=) (λ _ x : A, x).
Proof. intros ???; reflexivity. Qed.
Global Instance const1_idemp {A} : IdemP (=) (λ x _ : A, x).
Proof. intros ?; reflexivity. Qed.
Global Instance const2_idemp {A} : IdemP (=) (λ _ x : A, x).
Proof. intros ?; reflexivity. Qed.

(** ** Lists *)
Global Instance list_inhabited {A} : Inhabited (list A) := populate [].

Definition zip_with {A B C} (f : A  B  C) : list A  list B  list C :=
  fix go l1 l2 :=
  match l1, l2 with x1 :: l1, x2 :: l2 => f x1 x2 :: go l1 l2 | _ , _ => [] end.
Notation zip := (zip_with pair).

(** ** Booleans *)
(** The following coercion allows us to use Booleans as propositions. *)
Coercion Is_true : bool >-> Sortclass.
Global Hint Unfold Is_true : core.
Global Hint Immediate Is_true_eq_left : core.
Global Hint Resolve orb_prop_intro andb_prop_intro : core.
Notation "(&&)" := andb (only parsing).
Notation "(||)" := orb (only parsing).
Infix "&&*" := (zip_with (&&)) (at level 40).
Infix "||*" := (zip_with (||)) (at level 50).

Global Instance bool_inhabated : Inhabited bool := populate true.
Definition bool_le (β1 β2 : bool) : Prop := negb β1 || β2.
Infix "=.>" := bool_le (at level 70).
Infix "=.>*" := (Forall2 bool_le) (at level 70).
Global Instance: PartialOrder bool_le.
Proof. repeat split; repeat intros [|]; compute; tauto. Qed.
Lemma andb_True b1 b2 : b1 && b2  b1  b2.
Proof. destruct b1, b2; simpl; tauto. Qed.
Lemma orb_True b1 b2 : b1 || b2  b1  b2.
Proof. destruct b1, b2; simpl; tauto. Qed.
Lemma negb_True b : negb b  ¬b.
Proof. destruct b; simpl; tauto. Qed.
Lemma Is_true_false (b : bool) : b = false  ¬b.
Proof. now intros -> ?. Qed.
Global Instance unit_equiv : Equiv unit := λ _ _, True.
Global Instance unit_equivalence : Equivalence (≡@{unit}).
Global Instance unit_leibniz : LeibnizEquiv unit.
Proof. intros [] []; reflexivity. Qed.
Global Instance unit_inhabited: Inhabited unit := populate ().
(** ** Empty *)
Global Instance Empty_set_equiv : Equiv Empty_set := λ _ _, True.
Global Instance Empty_set_equivalence : Equivalence (≡@{Empty_set}).
Proof. repeat split. Qed.
Global Instance Empty_set_leibniz : LeibnizEquiv Empty_set.
Proof. intros [] []; reflexivity. Qed.

Notation "( x ,.)" := (pair x) (only parsing) : stdpp_scope.
Notation "(., y )" := (λ x, (x,y)) (only parsing) : stdpp_scope.
Notation "p .1" := (fst p) (at level 2, left associativity, format "p .1").
Notation "p .2" := (snd p) (at level 2, left associativity, format "p .2").
Global Instance: Params (@pair) 2 := {}.
Global Instance: Params (@fst) 2 := {}.
Global Instance: Params (@snd) 2 := {}.
(** The Coq standard library swapped the names of curry/uncurry, see
https://github.com/coq/coq/pull/12716/ 
FIXME: Remove this workaround once the lowest Coq version we support is 8.13. *)
Notation curry := prod_uncurry.
Global Instance: Params (@curry) 3 := {}.
Global Instance: Params (@uncurry) 3 := {}.

Definition uncurry3 {A B C D} (f : A  B  C  D) (p : A * B * C) : D :=
  let '(a,b,c) := p in f a b c.
Global Instance: Params (@uncurry3) 4 := {}.
Definition uncurry4 {A B C D E} (f : A  B  C  D  E) (p : A * B * C * D) : E :=
  let '(a,b,c,d) := p in f a b c d.
Global Instance: Params (@uncurry4) 5 := {}.
Definition curry3 {A B C D} (f : A * B * C  D) (a : A) (b : B) (c : C) : D :=
  f (a, b, c).
Global Instance: Params (@curry3) 4 := {}.
Definition curry4 {A B C D E} (f : A * B * C * D  E)
  (a : A) (b : B) (c : C) (d : D) : E := f (a, b, c, d).
Global Instance: Params (@curry4) 5 := {}.
Definition prod_map {A A' B B'} (f: A  A') (g: B  B') (p : A * B) : A' * B' :=
  (f (p.1), g (p.2)).
Global Arguments prod_map {_ _ _ _} _ _ !_ / : assert.
Definition prod_zip {A A' A'' B B' B''} (f : A  A'  A'') (g : B  B'  B'')
    (p : A * B) (q : A' * B') : A'' * B'' := (f (p.1) (q.1), g (p.2) (q.2)).
Global Arguments prod_zip {_ _ _ _ _ _} _ _ !_ !_ / : assert.
Global Instance prod_inhabited {A B} (iA : Inhabited A)
    (iB : Inhabited B) : Inhabited (A * B) :=
  match iA, iB with populate x, populate y => populate (x,y) end.
(** Note that we need eta for products for the [uncurry_curry] lemmas to hold
in non-applied form ([uncurry (curry f) = f]). *)
Lemma curry_uncurry {A B C} (f : A  B  C) : curry (uncurry f) = f.
Proof. reflexivity. Qed.
Lemma uncurry_curry {A B C} (f : A * B  C) p : uncurry (curry f) p = f p.
Proof. destruct p; reflexivity. Qed.
Lemma curry3_uncurry3 {A B C D} (f : A  B  C  D) : curry3 (uncurry3 f) = f.
Proof. reflexivity. Qed.
Lemma uncurry3_curry3 {A B C D} (f : A * B * C  D) p :
  uncurry3 (curry3 f) p = f p.
Proof. destruct p as [[??] ?]; reflexivity. Qed.
Lemma curry4_uncurry4 {A B C D E} (f : A  B  C  D  E) :
  curry4 (uncurry4 f) = f.
Proof. reflexivity. Qed.
Lemma uncurry4_curry4 {A B C D E} (f : A * B * C * D  E) p :
  uncurry4 (curry4 f) p = f p.
Proof. destruct p as [[[??] ?] ?]; reflexivity. Qed.

Global Instance pair_inj {A B} : Inj2 (=) (=) (=) (@pair A B).
Proof. injection 1; auto. Qed.
Global Instance prod_map_inj {A A' B B'} (f : A  A') (g : B  B') :
  Inj (=) (=) f  Inj (=) (=) g  Inj (=) (=) (prod_map f g).
Proof.
  intros ?? [??] [??] ?; simpl in *; f_equal;
    [apply (inj f)|apply (inj g)]; congruence.
Qed.
Definition prod_relation {A B} (R1 : relation A) (R2 : relation B) :
  relation (A * B) := λ x y, R1 (x.1) (y.1)  R2 (x.2) (y.2).
  Context `{RA : relation A, RB : relation B}.
  Global Instance prod_relation_refl :
    Reflexive RA  Reflexive RB  Reflexive (prod_relation RA RB).
  Proof. firstorder eauto. Qed.
  Global Instance prod_relation_sym :
    Symmetric RA  Symmetric RB  Symmetric (prod_relation RA RB).
  Proof. firstorder eauto. Qed.
  Global Instance prod_relation_trans :
    Transitive RA  Transitive RB  Transitive (prod_relation RA RB).
  Proof. firstorder eauto. Qed.
  Global Instance prod_relation_equiv :
    Equivalence RA  Equivalence RB  Equivalence (prod_relation RA RB).
  Global Instance pair_proper' : Proper (RA ==> RB ==> prod_relation RA RB) pair.
  Proof. firstorder eauto. Qed.
  Global Instance pair_inj' : Inj2 RA RB (prod_relation RA RB) pair.
  Proof. inversion_clear 1; eauto. Qed.
  Global Instance fst_proper' : Proper (prod_relation RA RB ==> RA) fst.
  Proof. firstorder eauto. Qed.
  Global Instance snd_proper' : Proper (prod_relation RA RB ==> RB) snd.
  Proof. firstorder eauto. Qed.

  Global Instance curry_proper' `{RC : relation C} :
    Proper ((prod_relation RA RB ==> RC) ==> RA ==> RB ==> RC) curry.
  Proof. firstorder eauto. Qed.
  Global Instance uncurry_proper' `{RC : relation C} :
    Proper ((RA ==> RB ==> RC) ==> prod_relation RA RB ==> RC) uncurry.
  Proof. intros f1 f2 Hf [x1 y1] [x2 y2] []; apply Hf; assumption. Qed.

  Global Instance curry3_proper' `{RC : relation C, RD : relation D} :
    Proper ((prod_relation (prod_relation RA RB) RC ==> RD) ==>
            RA ==> RB ==> RC ==> RD) curry3.
  Proof. firstorder eauto. Qed.
  Global Instance uncurry3_proper' `{RC : relation C, RD : relation D} :
    Proper ((RA ==> RB ==> RC ==> RD) ==>
            prod_relation (prod_relation RA RB) RC ==> RD) uncurry3.
  Proof. intros f1 f2 Hf [[??] ?] [[??] ?] [[??] ?]; apply Hf; assumption. Qed.

  Global Instance curry4_proper' `{RC : relation C, RD : relation D, RE : relation E} :
    Proper ((prod_relation (prod_relation (prod_relation RA RB) RC) RD ==> RE) ==>
            RA ==> RB ==> RC ==> RD ==> RE) curry4.
  Proof. firstorder eauto. Qed.
  Global Instance uncurry5_proper' `{RC : relation C, RD : relation D, RE : relation E} :
    Proper ((RA ==> RB ==> RC ==> RD ==> RE) ==>
            prod_relation (prod_relation (prod_relation RA RB) RC) RD ==> RE) uncurry4.
  Proof.
    intros f1 f2 Hf [[[??] ?] ?] [[[??] ?] ?] [[[??] ?] ?]; apply Hf; assumption.
  Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Instance prod_equiv `{Equiv A,Equiv B} : Equiv (A * B) :=
  prod_relation () ().

Robbert Krebbers's avatar
Robbert Krebbers committed
(** Below we make [prod_equiv] type class opaque, so we first lift all
instances *)
Section prod_setoid.
  Context `{Equiv A, Equiv B}.

  Global Instance prod_equivalence :
    Equivalence (≡@{A})  Equivalence (≡@{B})  Equivalence (≡@{A * B}) := _.

  Global Instance pair_proper : Proper (() ==> () ==> (≡@{A*B})) pair := _.
  Global Instance pair_equiv_inj : Inj2 () () (≡@{A*B}) pair := _.
  Global Instance fst_proper : Proper ((≡@{A*B}) ==> ()) fst := _.
  Global Instance snd_proper : Proper ((≡@{A*B}) ==> ()) snd := _.

  Global Instance curry_proper `{Equiv C} :
    Proper (((≡@{A*B}) ==> (≡@{C})) ==> () ==> () ==> ()) curry := _.
  Global Instance uncurry_proper `{Equiv C} :
    Proper ((() ==> () ==> ()) ==> (≡@{A*B}) ==> (≡@{C})) uncurry := _.

  Global Instance curry3_proper `{Equiv C, Equiv D} :
    Proper (((≡@{A*B*C}) ==> (≡@{D})) ==>
            () ==> () ==> () ==> ()) curry3 := _.
  Global Instance uncurry3_proper `{Equiv C, Equiv D} :
    Proper ((() ==> () ==> () ==> ()) ==>
            (≡@{A*B*C}) ==> (≡@{D})) uncurry3 := _.

  Global Instance curry4_proper `{Equiv C, Equiv D, Equiv E} :
    Proper (((≡@{A*B*C*D}) ==> (≡@{E})) ==>
            () ==> () ==> () ==> () ==> ()) curry4 := _.
  Global Instance uncurry4_proper `{Equiv C, Equiv D, Equiv E} :
    Proper ((() ==> () ==> () ==> () ==> ()) ==>
            (≡@{A*B*C*D}) ==> (≡@{E})) uncurry4 := _.
End prod_setoid.

Typeclasses Opaque prod_equiv.
Global Instance prod_leibniz `{LeibnizEquiv A, LeibnizEquiv B} :
  LeibnizEquiv (A * B).
Proof. intros [??] [??] [??]; f_equal; apply leibniz_equiv; auto. Qed.
Definition sum_map {A A' B B'} (f: A  A') (g: B  B') (xy : A + B) : A' + B' :=
  match xy with inl x => inl (f x) | inr y => inr (g y) end.
Global Arguments sum_map {_ _ _ _} _ _ !_ / : assert.
Global Instance sum_inhabited_l {A B} (iA : Inhabited A) : Inhabited (A + B) :=
  match iA with populate x => populate (inl x) end.
Global Instance sum_inhabited_r {A B} (iB : Inhabited B) : Inhabited (A + B) :=
  match iB with populate y => populate (inr y) end.
Global Instance inl_inj {A B} : Inj (=) (=) (@inl A B).
Proof. injection 1; auto. Qed.
Global Instance inr_inj {A B} : Inj (=) (=) (@inr A B).
Proof. injection 1; auto. Qed.
Global Instance sum_map_inj {A A' B B'} (f : A  A') (g : B  B') :
  Inj (=) (=) f  Inj (=) (=) g  Inj (=) (=) (sum_map f g).
Proof. intros ?? [?|?] [?|?] [=]; f_equal; apply (inj _); auto. Qed.

Inductive sum_relation {A B}
     (RA : relation A) (RB : relation B) : relation (A + B) :=
  | inl_related x1 x2 : RA x1 x2  sum_relation RA RB (inl x1) (inl x2)
  | inr_related y1 y2 : RB y1 y2  sum_relation RA RB (inr y1) (inr y2).

Section sum_relation.
  Context `{RA : relation A, RB : relation B}.
  Global Instance sum_relation_refl :
    Reflexive RA  Reflexive RB  Reflexive (sum_relation RA RB).
  Proof. intros ?? [?|?]; constructor; reflexivity. Qed.
  Global Instance sum_relation_sym :
    Symmetric RA  Symmetric RB  Symmetric (sum_relation RA RB).
  Proof. destruct 3; constructor; eauto. Qed.
  Global Instance sum_relation_trans :
    Transitive RA  Transitive RB  Transitive (sum_relation RA RB).
  Proof. destruct 3; inversion_clear 1; constructor; eauto. Qed.
  Global Instance sum_relation_equiv :
    Equivalence RA  Equivalence RB  Equivalence (sum_relation RA RB).
  Proof. split; apply _. Qed.
  Global Instance inl_proper' : Proper (RA ==> sum_relation RA RB) inl.
  Proof. constructor; auto. Qed.
  Global Instance inr_proper' : Proper (RB ==> sum_relation RA RB) inr.
  Proof. constructor; auto. Qed.
  Global Instance inl_inj' : Inj RA (sum_relation RA RB) inl.
  Proof. inversion_clear 1; auto. Qed.
  Global Instance inr_inj' : Inj RB (sum_relation RA RB) inr.
  Proof. inversion_clear 1; auto. Qed.
End sum_relation.

Global Instance sum_equiv `{Equiv A, Equiv B} : Equiv (A + B) := sum_relation () ().
Global Instance inl_proper `{Equiv A, Equiv B} : Proper (() ==> ()) (@inl A B) := _.
Global Instance inr_proper `{Equiv A, Equiv B} : Proper (() ==> ()) (@inr A B) := _.
Global Instance inl_equiv_inj `{Equiv A, Equiv B} : Inj () () (@inl A B) := _.
Global Instance inr_equiv_inj `{Equiv A, Equiv B} : Inj () () (@inr A B) := _.
Typeclasses Opaque sum_equiv.

Global Instance option_inhabited {A} : Inhabited (option A) := populate None.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Arguments existT {_ _} _ _ : assert.
Global Arguments projT1 {_ _} _ : assert.
Global Arguments projT2 {_ _} _ : assert.
Global Arguments exist {_} _ _ _ : assert.
Global Arguments proj1_sig {_ _} _ : assert.
Global Arguments proj2_sig {_ _} _ : assert.
Notation "x ↾ p" := (exist _ x p) (at level 20) : stdpp_scope.
Notation "` x" := (proj1_sig x) (at level 10, format "` x") : stdpp_scope.
Lemma proj1_sig_inj {A} (P : A  Prop) x (Px : P x) y (Py : P y) :
  xPx = yPy  x = y.
Proof. injection 1; trivial. Qed.
Section sig_map.
  Context `{P : A  Prop} `{Q : B  Prop} (f : A  B) (Hf :  x, P x  Q (f x)).
  Definition sig_map (x : sig P) : sig Q := f (`x)  Hf _ (proj2_sig x).
  Global Instance sig_map_inj:
    ( x, ProofIrrel (P x))  Inj (=) (=) f  Inj (=) (=) sig_map.
  Proof.
    intros ?? [x Hx] [y Hy]. injection 1. intros Hxy.
    apply (inj f) in Hxy; subst. rewrite (proof_irrel _ Hy). auto.
  Qed.
End sig_map.
Global Arguments sig_map _ _ _ _ _ _ !_ / : assert.
Definition proj1_ex {P : Prop} {Q : P  Prop} (p :  x, Q x) : P :=
  let '(ex_intro _ x _) := p in x.
Definition proj2_ex {P : Prop} {Q : P  Prop} (p :  x, Q x) : Q (proj1_ex p) :=
  let '(ex_intro _ x H) := p in H.
Robbert Krebbers's avatar
Robbert Krebbers committed

(** * Operations on sets *)
(** We define operational type classes for the traditional operations and
relations on sets: the empty set [∅], the union [(∪)],
intersection [(∩)], and difference [(∖)], the singleton [{[_]}], the subset
[(⊆)] and element of [(∈)] relation, and disjointess [(##)]. *)
Robbert Krebbers's avatar
Robbert Krebbers committed
Class Empty A := empty: A.
Global Hint Mode Empty ! : typeclass_instances.
Notation "∅" := empty (format "∅") : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Instance empty_inhabited `(Empty A) : Inhabited A := populate ∅.
Robbert Krebbers's avatar
Robbert Krebbers committed
Class Union A := union: A  A  A.
Global Hint Mode Union ! : typeclass_instances.
Global Instance: Params (@union) 2 := {}.
Infix "∪" := union (at level 50, left associativity) : stdpp_scope.
Notation "(∪)" := union (only parsing) : stdpp_scope.
Notation "( x ∪.)" := (union x) (only parsing) : stdpp_scope.
Notation "(.∪ x )" := (λ y, union y x) (only parsing) : stdpp_scope.
Infix "∪*" := (zip_with ()) (at level 50, left associativity) : stdpp_scope.
Notation "(∪*)" := (zip_with ()) (only parsing) : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Definition union_list `{Empty A} `{Union A} : list A  A := fold_right () ∅.
Global Arguments union_list _ _ _ !_ / : assert.
Notation "⋃ l" := (union_list l) (at level 20, format "⋃  l") : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed
Class Intersection A := intersection: A  A  A.
Global Hint Mode Intersection ! : typeclass_instances.
Global Instance: Params (@intersection) 2 := {}.
Infix "∩" := intersection (at level 40) : stdpp_scope.
Notation "(∩)" := intersection (only parsing) : stdpp_scope.
Notation "( x ∩.)" := (intersection x) (only parsing) : stdpp_scope.
Notation "(.∩ x )" := (λ y, intersection y x) (only parsing) : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Class Difference A := difference: A  A  A.
Global Hint Mode Difference ! : typeclass_instances.
Global Instance: Params (@difference) 2 := {}.
Infix "∖" := difference (at level 40, left associativity) : stdpp_scope.
Notation "(∖)" := difference (only parsing) : stdpp_scope.
Notation "( x ∖.)" := (difference x) (only parsing) : stdpp_scope.
Notation "(.∖ x )" := (λ y, difference y x) (only parsing) : stdpp_scope.
Infix "∖*" := (zip_with ()) (at level 40, left associativity) : stdpp_scope.
Notation "(∖*)" := (zip_with ()) (only parsing) : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Class Singleton A B := singleton: A  B.
Global Hint Mode Singleton - ! : typeclass_instances.
Global Instance: Params (@singleton) 3 := {}.
Notation "{[ x ]}" := (singleton x) (at level 1) : stdpp_scope.
  (union .. (union (singleton x) (singleton y)) .. (singleton z))
  (at level 1) : stdpp_scope.
Class SubsetEq A := subseteq: relation A.
Global Hint Mode SubsetEq ! : typeclass_instances.
Global Instance: Params (@subseteq) 2 := {}.
Infix "⊆" := subseteq (at level 70) : stdpp_scope.
Notation "(⊆)" := subseteq (only parsing) : stdpp_scope.
Notation "( X ⊆.)" := (subseteq X) (only parsing) : stdpp_scope.
Notation "(.⊆ X )" := (λ Y, Y  X) (only parsing) : stdpp_scope.
Notation "X ⊈ Y" := (¬X  Y) (at level 70) : stdpp_scope.
Notation "(⊈)" := (λ X Y, X  Y) (only parsing) : stdpp_scope.
Notation "( X ⊈.)" := (λ Y, X  Y) (only parsing) : stdpp_scope.
Notation "(.⊈ X )" := (λ Y, Y  X) (only parsing) : stdpp_scope.

Infix "⊆@{ A }" := (@subseteq A _) (at level 70, only parsing) : stdpp_scope.
Notation "(⊆@{ A } )" := (@subseteq A _) (only parsing) : stdpp_scope.

Infix "⊆*" := (Forall2 ()) (at level 70) : stdpp_scope.
Notation "(⊆*)" := (Forall2 ()) (only parsing) : stdpp_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Hint Extern 0 (_  _) => reflexivity : core.
Global Hint Extern 0 (_ ⊆* _) => reflexivity : core.
Infix "⊂" := (strict ()) (at level 70) : stdpp_scope.
Notation "(⊂)" := (strict ()) (only parsing) : stdpp_scope.
Notation "( X ⊂.)" := (strict () X) (only parsing) : stdpp_scope.
Notation "(.⊂ X )" := (λ Y, Y  X) (only parsing) : stdpp_scope.
Notation "X ⊄ Y" := (¬X  Y) (at level 70) : stdpp_scope.
Notation "(⊄)" := (λ X Y, X  Y) (only parsing) : stdpp_scope.