Newer
Older
(* Copyright (c) 2012-2019, Coq-std++ developers. *)
Robbert Krebbers
committed
(* This file is distributed under the terms of the BSD license. *)
(** 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
Robbert Krebbers
committed
structures. *)
From Coq Require Export Morphisms RelationClasses List Bool Utf8 Setoid.
From Coq Require Import Permutation.
Set Default Proof Using "Type".
Export ListNotations.
From Coq.Program Require Export Basics Syntax.
(** * 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 taht 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. *)
(** 3. Hide obligations from the results of the [Search] commands. *)
Add Search Blacklist "_obligation_".
(** * Sealing off definitions *)
Section seal.
Local Set Primitive Projections.
Record seal {A} (f : A) := { unseal : A; seal_eq : unseal = f }.
End seal.
Arguments unseal {_ _} _ : assert.
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 }.
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
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
| TCIf_true : P → Q → TCIf P Q R
| TCIf_false : R → TCIf P Q R.
Existing Class TCIf.
Hint Extern 0 (TCIf _ _ _) =>
first [apply TCIf_true; [apply _|]
|apply TCIf_false] : typeclass_instances.
(** * Typeclass opaque definitions *)
(** 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.
Arguments tc_opaque {_} _ /.
(** 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.
Existing Instance TCOr_l | 9.
Existing Instance TCOr_r | 10.
Inductive TCAnd (P1 P2 : Prop) : Prop := TCAnd_intro : P1 → P2 → TCAnd P1 P2.
Existing Class TCAnd.
Existing Instance TCAnd_intro.
Inductive TCTrue : Prop := TCTrue_intro : TCTrue.
Existing Class TCTrue.
Existing Instance TCTrue_intro.
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.
Existing Instance TCForall_nil.
Existing Instance TCForall_cons.
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.
Existing Instance TCForall2_nil.
Existing Instance TCForall2_cons.
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.
Existing Instance TCElemOf_here.
Existing Instance TCElemOf_further.
Inductive TCEq {A} (x : A) : A → Prop := TCEq_refl : TCEq x x.
Existing Class TCEq.
Existing Instance TCEq_refl.
Inductive TCDiag {A} (C : A → Prop) : A → A → Prop :=
| TCDiag_diag x : C x → TCDiag C x x.
Existing Class TCDiag.
Existing Instance TCDiag_diag.
Robbert Krebbers
committed
(** 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. *)
Delimit Scope stdpp_scope with stdpp.
Global Open Scope stdpp_scope.
Robbert Krebbers
committed
(** 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.
(** * Equality *)
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
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.
Robbert Krebbers
committed
Hint Extern 0 (_ = _) => reflexivity : core.
Hint Extern 100 (_ ≠ _) => discriminate : core.
Robbert Krebbers
committed
Instance: ∀ A, PreOrder (=@{A}).
Proof. split; repeat intro; congruence. Qed.
(** ** Setoid equality *)
(** 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.
(* No Hint Mode set because of Coq bug #5735
Hint Mode Equiv ! : typeclass_instances. *)
Infix "≡" := equiv (at level 70, no associativity) : stdpp_scope.
Robbert Krebbers
committed
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.
Robbert Krebbers
committed
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.
Robbert Krebbers
committed
(** 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.
Hint Mode LeibnizEquiv ! - : typeclass_instances.
Robbert Krebbers
committed
Lemma leibniz_equiv_iff `{LeibnizEquiv A, !Reflexive (≡@{A})} (x y : A) :
x ≡ y ↔ x = y.
Proof. split. apply leibniz_equiv. intros ->; reflexivity. Qed.
Ltac fold_leibniz := repeat
match goal with
Robbert Krebbers
committed
| H : context [ _ ≡@{?A} _ ] |- _ =>
setoid_rewrite (leibniz_equiv_iff (A:=A)) in H
Robbert Krebbers
committed
| |- context [ _ ≡@{?A} _ ] =>
setoid_rewrite (leibniz_equiv_iff (A:=A))
end.
Ltac unfold_leibniz := repeat
match goal with
Robbert Krebbers
committed
| H : context [ _ =@{?A} _ ] |- _ =>
setoid_rewrite <-(leibniz_equiv_iff (A:=A)) in H
Robbert Krebbers
committed
| |- context [ _ =@{?A} _ ] =>
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. *)
(** The following instance forces [setoid_replace] to use setoid equality
(for types that have an [Equiv] instance) rather than the standard Leibniz
equality. *)
Instance equiv_default_relation `{Equiv A} : DefaultRelation (≡) | 3 := {}.
Hint Extern 0 (_ ≡ _) => reflexivity : core.
Hint Extern 0 (_ ≡ _) => symmetry; assumption : core.
(** * Type classes *)
(** ** Decidable propositions *)
(** This type class by (Spitters/van der Weegen, 2011) collects decidable
Robbert Krebbers
committed
propositions. *)
Class Decision (P : Prop) := decide : {P} + {¬P}.
Hint Mode Decision ! : typeclass_instances.
Robbert Krebbers
committed
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).
Hint Mode RelDecision ! ! ! : typeclass_instances.
Arguments decide_rel {_ _} _ {_} _ _ : simpl never, assert.
Robbert Krebbers
committed
Notation EqDecision A := (RelDecision (=@{A})).
(** ** Inhabited types *)
(** This type class collects types that are inhabited. *)
Class Inhabited (A : Type) : Type := populate { inhabitant : A }.
Hint Mode Inhabited ! : typeclass_instances.
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.
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]. *)
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
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.
Arguments irreflexivity {_} _ {_} _ _ : assert.
Arguments inj {_ _ _ _} _ {_} _ _ _ : assert.
Arguments inj2 {_ _ _ _ _ _} _ {_} _ _ _ _ _: assert.
Arguments cancel {_ _ _} _ _ {_} _ : assert.
Arguments surj {_ _ _} _ {_} _ : assert.
Arguments idemp {_ _} _ {_} _ : assert.
Arguments comm {_ _ _} _ {_} _ _ : assert.
Arguments left_id {_ _} _ _ {_} _ : assert.
Arguments right_id {_ _} _ _ {_} _ : assert.
Arguments assoc {_ _} _ {_} _ _ _ : assert.
Arguments left_absorb {_ _} _ _ {_} _ : assert.
Arguments right_absorb {_ _} _ _ {_} _ : assert.
Arguments anti_symm {_ _} _ {_} _ _ _ _ : assert.
Arguments total {_} _ {_} _ _ : assert.
Arguments trichotomy {_} _ {_} _ _ : assert.
Arguments trichotomyT {_} _ {_} _ _ : assert.
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
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.
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.
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.
Class PartialOrder {A} (R : relation A) : Prop := {
partial_order_pre :> PreOrder R;
partial_order_anti_symm :> AntiSymm (=) R
}.
Class TotalOrder {A} (R : relation A) : Prop := {
total_order_partial :> PartialOrder R;
total_order_trichotomy :> Trichotomy (strict R)
}.
(** * Logic *)
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.
Hint Extern 0 (_ ↔ _) => reflexivity : core.
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.
Robbert Krebbers
committed
Instance: Comm (↔) (=@{A}).
Proof. red; intuition. Qed.
Robbert Krebbers
committed
Instance: Comm (↔) (λ x y, y =@{A} x).
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
Proof. red; intuition. Qed.
Instance: Comm (↔) (↔).
Proof. red; intuition. Qed.
Instance: Comm (↔) (∧).
Proof. red; intuition. Qed.
Instance: Assoc (↔) (∧).
Proof. red; intuition. Qed.
Instance: IdemP (↔) (∧).
Proof. red; intuition. Qed.
Instance: Comm (↔) (∨).
Proof. red; intuition. Qed.
Instance: Assoc (↔) (∨).
Proof. red; intuition. Qed.
Instance: IdemP (↔) (∨).
Proof. red; intuition. Qed.
Instance: LeftId (↔) True (∧).
Proof. red; intuition. Qed.
Instance: RightId (↔) True (∧).
Proof. red; intuition. Qed.
Instance: LeftAbsorb (↔) False (∧).
Proof. red; intuition. Qed.
Instance: RightAbsorb (↔) False (∧).
Proof. red; intuition. Qed.
Instance: LeftId (↔) False (∨).
Proof. red; intuition. Qed.
Instance: RightId (↔) False (∨).
Proof. red; intuition. Qed.
Instance: LeftAbsorb (↔) True (∨).
Proof. red; intuition. Qed.
Instance: RightAbsorb (↔) True (∨).
Proof. red; intuition. Qed.
Instance: LeftId (↔) True impl.
Proof. unfold impl. red; intuition. Qed.
Instance: 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.
(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.
Robbert Krebbers
committed
Instance impl_inhabited {A} `{Inhabited B} : Inhabited (A → B) :=
populate (λ _, inhabitant).
(** Ensure that [simpl] unfolds [id], [compose], and [flip] when fully
applied. *)
Arguments id _ _ / : assert.
Arguments compose _ _ _ _ _ _ / : assert.
Arguments flip _ _ _ _ _ _ / : assert.
Arguments const _ _ _ _ / : assert.
Robbert Krebbers
committed
Typeclasses Transparent id compose flip const.
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
Definition fun_map {A A' B B'} (f: A' → A) (g: B → B') (h : A → B) : A' → B' :=
g ∘ h ∘ f.
Instance const_proper `{R1 : relation A, R2 : relation B} (x : B) :
Reflexive R2 → Proper (R1 ==> R2) (λ _, x).
Proof. intros ? y1 y2; reflexivity. Qed.
Instance id_inj {A} : Inj (=) (=) (@id A).
Proof. intros ??; auto. Qed.
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.
Instance id_surj {A} : Surj (=) (@id A).
Proof. intros y; exists y; reflexivity. Qed.
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.
Instance id_comm {A B} (x : B) : Comm (=) (λ _ _ : A, x).
Proof. intros ?; reflexivity. Qed.
Instance id_assoc {A} (x : A) : Assoc (=) (λ _ _ : A, x).
Proof. intros ???; reflexivity. Qed.
Instance const1_assoc {A} : Assoc (=) (λ x _ : A, x).
Proof. intros ???; reflexivity. Qed.
Instance const2_assoc {A} : Assoc (=) (λ _ x : A, x).
Proof. intros ???; reflexivity. Qed.
Instance const1_idemp {A} : IdemP (=) (λ x _ : A, x).
Proof. intros ?; reflexivity. Qed.
Instance const2_idemp {A} : IdemP (=) (λ _ x : A, x).
Proof. intros ?; reflexivity. Qed.
(** ** Lists *)
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.
Hint Unfold Is_true : core.
Hint Immediate Is_true_eq_left : core.
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).
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).
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.
(** ** Unit *)
Instance unit_equiv : Equiv unit := λ _ _, True.
Robbert Krebbers
committed
Instance unit_equivalence : Equivalence (≡@{unit}).
Proof. repeat split. Qed.
Instance unit_leibniz : LeibnizEquiv unit.
Proof. intros [] []; reflexivity. Qed.
Instance unit_inhabited: Inhabited unit := populate ().
(** ** Empty *)
Instance Empty_set_equiv : Equiv Empty_set := λ _ _, True.
Instance Empty_set_equivalence : Equivalence (≡@{Empty_set}).
Proof. repeat split. Qed.
Instance Empty_set_leibniz : LeibnizEquiv Empty_set.
Proof. intros [] []; reflexivity. Qed.
(** ** Products *)
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").
Instance: Params (@pair) 2 := {}.
Instance: Params (@fst) 2 := {}.
Instance: Params (@snd) 2 := {}.
Notation curry := prod_curry.
Notation uncurry := prod_uncurry.
Definition curry3 {A B C D} (f : A → B → C → D) (p : A * B * C) : D :=
let '(a,b,c) := p in f a b c.
Definition curry4 {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.
Definition uncurry3 {A B C D} (f : A * B * C → D) (a : A) (b : B) (c : C) : D :=
f (a, b, c).
Definition uncurry4 {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).
Definition prod_map {A A' B B'} (f: A → A') (g: B → B') (p : A * B) : A' * B' :=
(f (p.1), g (p.2)).
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)).
Arguments prod_zip {_ _ _ _ _ _} _ _ !_ !_ / : assert.
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.
Robbert Krebbers
committed
Instance pair_inj : Inj2 (=) (=) (=) (@pair A B).
Proof. injection 1; auto. Qed.
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.
Robbert Krebbers
committed
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).
Section prod_relation.
Context `{R1 : relation A, R2 : relation B}.
Global Instance prod_relation_refl :
Reflexive R1 → Reflexive R2 → Reflexive (prod_relation R1 R2).
Proof. firstorder eauto. Qed.
Global Instance prod_relation_sym :
Symmetric R1 → Symmetric R2 → Symmetric (prod_relation R1 R2).
Proof. firstorder eauto. Qed.
Global Instance prod_relation_trans :
Transitive R1 → Transitive R2 → Transitive (prod_relation R1 R2).
Proof. firstorder eauto. Qed.
Global Instance prod_relation_equiv :
Equivalence R1 → Equivalence R2 → Equivalence (prod_relation R1 R2).
Proof. split; apply _. Qed.
Global Instance pair_proper' : Proper (R1 ==> R2 ==> prod_relation R1 R2) pair.
Proof. firstorder eauto. Qed.
Global Instance pair_inj' : Inj2 R1 R2 (prod_relation R1 R2) pair.
Proof. inversion_clear 1; eauto. Qed.
Global Instance fst_proper' : Proper (prod_relation R1 R2 ==> R1) fst.
Proof. firstorder eauto. Qed.
Global Instance snd_proper' : Proper (prod_relation R1 R2 ==> R2) snd.
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) := _.
Instance pair_equiv_inj `{Equiv A, Equiv B} : Inj2 (≡) (≡) (≡) (@pair A B) := _.
Instance fst_proper `{Equiv A, Equiv B} : Proper ((≡) ==> (≡)) (@fst A B) := _.
Instance snd_proper `{Equiv A, Equiv B} : Proper ((≡) ==> (≡)) (@snd A B) := _.
Typeclasses Opaque prod_equiv.
Instance prod_leibniz `{LeibnizEquiv A, LeibnizEquiv B} : LeibnizEquiv (A * B).
Proof. intros [??] [??] [??]; f_equal; apply leibniz_equiv; auto. Qed.
(** ** Sums *)
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.
Arguments sum_map {_ _ _ _} _ _ !_ / : assert.
Instance sum_inhabited_l {A B} (iA : Inhabited A) : Inhabited (A + B) :=
Robbert Krebbers
committed
match iA with populate x => populate (inl x) end.
Instance sum_inhabited_r {A B} (iB : Inhabited A) : Inhabited (A + B) :=
Robbert Krebbers
committed
match iB with populate y => populate (inl y) end.
Instance inl_inj : Inj (=) (=) (@inl A B).
Proof. injection 1; auto. Qed.
Instance inr_inj : Inj (=) (=) (@inr A B).
Proof. injection 1; auto. Qed.
700
701
702
703
704
705
706
707
708
709
710
711
712
713
714
715
716
717
718
719
720
721
722
723
724
725
726
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}
(R1 : relation A) (R2 : relation B) : relation (A + B) :=
| inl_related x1 x2 : R1 x1 x2 → sum_relation R1 R2 (inl x1) (inl x2)
| inr_related y1 y2 : R2 y1 y2 → sum_relation R1 R2 (inr y1) (inr y2).
Section sum_relation.
Context `{R1 : relation A, R2 : relation B}.
Global Instance sum_relation_refl :
Reflexive R1 → Reflexive R2 → Reflexive (sum_relation R1 R2).
Proof. intros ?? [?|?]; constructor; reflexivity. Qed.
Global Instance sum_relation_sym :
Symmetric R1 → Symmetric R2 → Symmetric (sum_relation R1 R2).
Proof. destruct 3; constructor; eauto. Qed.
Global Instance sum_relation_trans :
Transitive R1 → Transitive R2 → Transitive (sum_relation R1 R2).
Proof. destruct 3; inversion_clear 1; constructor; eauto. Qed.
Global Instance sum_relation_equiv :
Equivalence R1 → Equivalence R2 → Equivalence (sum_relation R1 R2).
Proof. split; apply _. Qed.
Global Instance inl_proper' : Proper (R1 ==> sum_relation R1 R2) inl.
Proof. constructor; auto. Qed.
Global Instance inr_proper' : Proper (R2 ==> sum_relation R1 R2) inr.
Proof. constructor; auto. Qed.
Global Instance inl_inj' : Inj R1 (sum_relation R1 R2) inl.
Proof. inversion_clear 1; auto. Qed.
Global Instance inr_inj' : Inj R2 (sum_relation R1 R2) inr.
Proof. inversion_clear 1; auto. Qed.
End sum_relation.
Instance sum_equiv `{Equiv A, Equiv B} : Equiv (A + B) := sum_relation (≡) (≡).
Instance inl_proper `{Equiv A, Equiv B} : Proper ((≡) ==> (≡)) (@inl A B) := _.
Instance inr_proper `{Equiv A, Equiv B} : Proper ((≡) ==> (≡)) (@inr A B) := _.
Instance inl_equiv_inj `{Equiv A, Equiv B} : Inj (≡) (≡) (@inl A B) := _.
Instance inr_equiv_inj `{Equiv A, Equiv B} : Inj (≡) (≡) (@inr A B) := _.
(** ** Option *)
Instance option_inhabited {A} : Inhabited (option A) := populate None.
(** ** Sigma types *)
Arguments existT {_ _} _ _ : assert.
Arguments projT1 {_ _} _ : assert.
Arguments projT2 {_ _} _ : assert.
Robbert Krebbers
committed
Arguments exist {_} _ _ _ : assert.
Arguments proj1_sig {_ _} _ : assert.
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) :
x↾Px = y↾Py → 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.
Arguments sig_map _ _ _ _ _ _ !_ / : assert.
Robbert Krebbers
committed
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.
(** 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
Jacques-Henri Jourdan
committed
[(⊆)] and element of [(∈)] relation, and disjointess [(##)]. *)
Hint Mode Empty ! : typeclass_instances.
Notation "∅" := empty (format "∅") : stdpp_scope.
Instance empty_inhabited `(Empty A) : Inhabited A := populate ∅.
Hint Mode Union ! : typeclass_instances.
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.
Infix "∪**" := (zip_with (zip_with (∪)))
(at level 50, left associativity) : stdpp_scope.
Infix "∪*∪**" := (zip_with (prod_zip (∪) (∪*)))
(at level 50, left associativity) : stdpp_scope.
Robbert Krebbers
committed
Definition union_list `{Empty A} `{Union A} : list A → A := fold_right (∪) ∅.
Arguments union_list _ _ _ !_ / : assert.
Notation "⋃ l" := (union_list l) (at level 20, format "⋃ l") : stdpp_scope.
Class DisjUnion A := disj_union: A → A → A.
Hint Mode DisjUnion ! : typeclass_instances.
Instance: Params (@disj_union) 2 := {}.
Infix "⊎" := disj_union (at level 50, left associativity) : stdpp_scope.
Notation "(⊎)" := disj_union (only parsing) : stdpp_scope.
Notation "( x ⊎.)" := (disj_union x) (only parsing) : stdpp_scope.
Notation "(.⊎ x )" := (λ y, disj_union y x) (only parsing) : stdpp_scope.
Hint Mode Intersection ! : typeclass_instances.
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.
Hint Mode Difference ! : typeclass_instances.
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.
Infix "∖**" := (zip_with (zip_with (∖)))
(at level 40, left associativity) : stdpp_scope.
Infix "∖*∖**" := (zip_with (prod_zip (∖) (∖*)))
(at level 50, left associativity) : stdpp_scope.
Robbert Krebbers
committed
Class Singleton A B := singleton: A → B.
Hint Mode Singleton - ! : typeclass_instances.
Notation "{[ x ]}" := (singleton x) (at level 1) : stdpp_scope.
Robbert Krebbers
committed
Notation "{[ x ; y ; .. ; z ]}" :=
Robbert Krebbers
committed
(union .. (union (singleton x) (singleton y)) .. (singleton z))
Robbert Krebbers
committed
Notation "{[ x , y ]}" := (singleton (x,y))
(at level 1, y at next level) : stdpp_scope.
Robbert Krebbers
committed
Notation "{[ x , y , z ]}" := (singleton (x,y,z))
(at level 1, y at next level, z at next level) : stdpp_scope.
Robbert Krebbers
committed
Class SubsetEq A := subseteq: relation A.
Hint Mode SubsetEq ! : typeclass_instances.
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.
Infix "⊆**" := (Forall2 (⊆*)) (at level 70) : stdpp_scope.
Infix "⊆1*" := (Forall2 (λ p q, p.1 ⊆ q.1)) (at level 70) : stdpp_scope.
Infix "⊆2*" := (Forall2 (λ p q, p.2 ⊆ q.2)) (at level 70) : stdpp_scope.
Infix "⊆1**" := (Forall2 (λ p q, p.1 ⊆* q.1)) (at level 70) : stdpp_scope.
Infix "⊆2**" := (Forall2 (λ p q, p.2 ⊆* q.2)) (at level 70) : stdpp_scope.
Hint Extern 0 (_ ⊆ _) => reflexivity : core.
Hint Extern 0 (_ ⊆* _) => reflexivity : core.
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.
Notation "( X ⊄.)" := (λ Y, X ⊄ Y) (only parsing) : stdpp_scope.
Notation "(.⊄ X )" := (λ Y, Y ⊄ X) (only parsing) : stdpp_scope.
Infix "⊂@{ A }" := (strict (⊆@{A})) (at level 70, only parsing) : stdpp_scope.
Notation "(⊂@{ A } )" := (strict (⊆@{A})) (only parsing) : stdpp_scope.
Notation "X ⊆ Y ⊆ Z" := (X ⊆ Y ∧ Y ⊆ Z) (at level 70, Y at next level) : stdpp_scope.
Notation "X ⊆ Y ⊂ Z" := (X ⊆ Y ∧ Y ⊂ Z) (at level 70, Y at next level) : stdpp_scope.
Notation "X ⊂ Y ⊆ Z" := (X ⊂ Y ∧ Y ⊆ Z) (at level 70, Y at next level) : stdpp_scope.
Notation "X ⊂ Y ⊂ Z" := (X ⊂ Y ∧ Y ⊂ Z) (at level 70, Y at next level) : stdpp_scope.
Definition option_to_set `{Singleton A C, Empty C} (mx : option A) : C :=
match mx with None => ∅ | Some x => {[ x ]} end.
Fixpoint list_to_set `{Singleton A C, Empty C, Union C} (l : list A) : C :=
match l with [] => ∅ | x :: l => {[ x ]} ∪ list_to_set l end.
Fixpoint list_to_set_disj `{Singleton A C, Empty C, DisjUnion C} (l : list A) : C :=
match l with [] => ∅ | x :: l => {[ x ]} ⊎ list_to_set_disj l end.
(** The class [Lexico A] is used for the lexicographic order on [A]. This order
is used to create finite maps, finite sets, etc, and is typically different from
the order [(⊆)]. *)
Class Lexico A := lexico: relation A.
Hint Mode Lexico ! : typeclass_instances.
Hint Mode ElemOf - ! : typeclass_instances.
Infix "∈" := elem_of (at level 70) : stdpp_scope.
Notation "(∈)" := elem_of (only parsing) : stdpp_scope.
Notation "( x ∈.)" := (elem_of x) (only parsing) : stdpp_scope.
Notation "(.∈ X )" := (λ x, elem_of x X) (only parsing) : stdpp_scope.
Notation "x ∉ X" := (¬x ∈ X) (at level 80) : stdpp_scope.
Notation "(∉)" := (λ x X, x ∉ X) (only parsing) : stdpp_scope.
Notation "( x ∉.)" := (λ X, x ∉ X) (only parsing) : stdpp_scope.
Notation "(.∉ X )" := (λ x, x ∉ X) (only parsing) : stdpp_scope.
Infix "∈@{ B }" := (@elem_of _ B _) (at level 70, only parsing) : stdpp_scope.
Notation "(∈@{ B } )" := (@elem_of _ B _) (only parsing) : stdpp_scope.
Notation "x ∉@{ B } X" := (¬x ∈@{B} X) (at level 80, only parsing) : stdpp_scope.
Notation "(∉@{ B } )" := (λ x X, x ∉@{B} X) (only parsing) : stdpp_scope.
Jacques-Henri Jourdan
committed
Hint Mode Disjoint ! : typeclass_instances.
Infix "##" := disjoint (at level 70) : stdpp_scope.
Notation "(##)" := disjoint (only parsing) : stdpp_scope.
Notation "( X ##.)" := (disjoint X) (only parsing) : stdpp_scope.
Notation "(.## X )" := (λ Y, Y ## X) (only parsing) : stdpp_scope.
Infix "##@{ A }" := (@disjoint A _) (at level 70, only parsing) : stdpp_scope.
Notation "(##@{ A } )" := (@disjoint A _) (only parsing) : stdpp_scope.
Infix "##*" := (Forall2 (##)) (at level 70) : stdpp_scope.
Notation "(##*)" := (Forall2 (##)) (only parsing) : stdpp_scope.
Infix "##**" := (Forall2 (##*)) (at level 70) : stdpp_scope.
Infix "##1*" := (Forall2 (λ p q, p.1 ## q.1)) (at level 70) : stdpp_scope.
Infix "##2*" := (Forall2 (λ p q, p.2 ## q.2)) (at level 70) : stdpp_scope.
Infix "##1**" := (Forall2 (λ p q, p.1 ##* q.1)) (at level 70) : stdpp_scope.
Infix "##2**" := (Forall2 (λ p q, p.2 ##* q.2)) (at level 70) : stdpp_scope.
Hint Extern 0 (_ ## _) => symmetry; eassumption : core.
Hint Extern 0 (_ ##* _) => symmetry; eassumption : core.
Class DisjointE E A := disjointE : E → A → A → Prop.
Hint Mode DisjointE - ! : typeclass_instances.
Jacques-Henri Jourdan
committed
Notation "X ##{ Γ } Y" := (disjointE Γ X Y)
(at level 70, format "X ##{ Γ } Y") : stdpp_scope.
Notation "(##{ Γ } )" := (disjointE Γ) (only parsing, Γ at level 1) : stdpp_scope.
Jacques-Henri Jourdan
committed
Notation "Xs ##{ Γ }* Ys" := (Forall2 (##{Γ}) Xs Ys)
(at level 70, format "Xs ##{ Γ }* Ys") : stdpp_scope.
Jacques-Henri Jourdan
committed
Notation "(##{ Γ }* )" := (Forall2 (##{Γ}))
(only parsing, Γ at level 1) : stdpp_scope.
Jacques-Henri Jourdan
committed
Notation "X ##{ Γ1 , Γ2 , .. , Γ3 } Y" := (disjoint (pair .. (Γ1, Γ2) .. Γ3) X Y)
(at level 70, format "X ##{ Γ1 , Γ2 , .. , Γ3 } Y") : stdpp_scope.
Jacques-Henri Jourdan
committed
Notation "Xs ##{ Γ1 , Γ2 , .. , Γ3 }* Ys" :=
(Forall2 (disjoint (pair .. (Γ1, Γ2) .. Γ3)) Xs Ys)
(at level 70, format "Xs ##{ Γ1 , Γ2 , .. , Γ3 }* Ys") : stdpp_scope.
Hint Extern 0 (_ ##{_} _) => symmetry; eassumption : core.
Robbert Krebbers
committed
Class DisjointList A := disjoint_list : list A → Prop.
Hint Mode DisjointList ! : typeclass_instances.
Instance: Params (@disjoint_list) 2 := {}.
Notation "## Xs" := (disjoint_list Xs) (at level 20, format "## Xs") : stdpp_scope.
Notation "##@{ A } Xs" :=
(@disjoint_list A _ Xs) (at level 20, only parsing) : stdpp_scope.
Robbert Krebbers
committed
Section disjoint_list.
Context `{Disjoint A, Union A, Empty A}.
Inductive disjoint_list_default : DisjointList A :=
Jacques-Henri Jourdan
committed
| disjoint_cons_2 (X : A) (Xs : list A) : X ## ⋃ Xs → ## Xs → ## (X :: Xs).
Global Existing Instance disjoint_list_default.
Robbert Krebbers
committed
Robbert Krebbers
committed
Proof. split; constructor. Qed.
Jacques-Henri Jourdan
committed
Lemma disjoint_list_cons X Xs : ## (X :: Xs) ↔ X ## ⋃ Xs ∧ ## Xs.
Robbert Krebbers
committed
Proof. split. inversion_clear 1; auto. intros [??]. constructor; auto. Qed.
End disjoint_list.
Robbert Krebbers
committed
Class Filter A B := filter: ∀ (P : A → Prop) `{∀ x, Decision (P x)}, B → B.
Hint Mode Filter - ! : typeclass_instances.
Class UpClose A B := up_close : A → B.
Hint Mode UpClose - ! : typeclass_instances.
Notation "↑ x" := (up_close x) (at level 20, format "↑ x").
(** * Monadic operations *)
(** We define operational type classes for the monadic operations bind, join
and fmap. We use these type classes merely for convenient overloading of
notations and do not formalize any theory on monads (we do not even define a
class with the monad laws). *)
Class MRet (M : Type → Type) := mret: ∀ {A}, A → M A.
Arguments mret {_ _ _} _ : assert.
Class MBind (M : Type → Type) := mbind : ∀ {A B}, (A → M B) → M A → M B.
Arguments mbind {_ _ _ _} _ !_ / : assert.
Class MJoin (M : Type → Type) := mjoin: ∀ {A}, M (M A) → M A.
Arguments mjoin {_ _ _} !_ / : assert.