diff --git a/theories/ars.v b/theories/ars.v
index 9a4fb2241c948fe322d0c2a3aff8a6c4141eded3..57004f7ebb1669f179674ce8c42748d3fbe7ab4c 100644
--- a/theories/ars.v
+++ b/theories/ars.v
@@ -4,6 +4,7 @@
 These are particularly useful as we define the operational semantics as a
 small step semantics. This file defines a hint database [ars] containing
 some theorems on abstract rewriting systems. *)
+Require Import Wf_nat.
 Require Export tactics base.
 
 (** * Definitions *)
@@ -47,13 +48,13 @@ Hint Constructors rtc nsteps bsteps tc : ars.
 Section rtc.
   Context `{R : relation A}.
 
-  Global Instance: Reflexive (rtc R).
-  Proof rtc_refl R.
-  Global Instance rtc_trans: Transitive (rtc R).
-  Proof. red; induction 1; eauto with ars. Qed.
+  Global Instance rtc_reflexive: Reflexive (rtc R).
+  Proof. red. apply rtc_refl. Qed.
+  Global Instance rtc_transitive: Transitive (rtc R).
+  Proof. red. induction 1; eauto with ars. Qed.
   Lemma rtc_once x y : R x y → rtc R x y.
   Proof. eauto with ars. Qed.
-  Global Instance: subrelation R (rtc R).
+  Instance rtc_once_subrel: subrelation R (rtc R).
   Proof. exact @rtc_once. Qed.
   Lemma rtc_r x y z : rtc R x y → R y z → rtc R x z.
   Proof. intros. etransitivity; eauto with ars. Qed.
@@ -62,8 +63,9 @@ Section rtc.
   Proof. inversion_clear 1; eauto. Qed.
 
   Lemma rtc_ind_r (P : A → A → Prop)
-    (Prefl : ∀ x, P x x) (Pstep : ∀ x y z, rtc R x y → R y z → P x y → P x z) :
-    ∀ y z, rtc R y z → P y z.
+    (Prefl : ∀ x, P x x)
+    (Pstep : ∀ x y z, rtc R x y → R y z → P x y → P x z) :
+    ∀ x z, rtc R x z → P x z.
   Proof.
     cut (∀ y z, rtc R y z → ∀ x, rtc R x y → P x y → P x z).
     { eauto using rtc_refl. }
@@ -99,7 +101,7 @@ Section rtc.
     bsteps R n x y → bsteps R (m + n) x y.
   Proof. apply bsteps_weaken. auto with arith. Qed.
   Lemma bsteps_S n x y :  bsteps R n x y → bsteps R (S n) x y.
-  Proof. apply bsteps_weaken. auto with arith. Qed.
+  Proof. apply bsteps_weaken. lia. Qed.
   Lemma bsteps_trans n m x y z :
     bsteps R n x y → bsteps R m y z → bsteps R (n + m) x z.
   Proof. induction 1; simpl; eauto using bsteps_plus_l with ars. Qed.
@@ -108,7 +110,31 @@ Section rtc.
   Lemma bsteps_rtc n x y : bsteps R n x y → rtc R x y.
   Proof. induction 1; eauto with ars. Qed.
   Lemma rtc_bsteps x y : rtc R x y → ∃ n, bsteps R n x y.
-  Proof. induction 1. exists 0. auto with ars. firstorder eauto with ars. Qed.
+  Proof.
+    induction 1.
+    * exists 0. constructor.
+    * naive_solver eauto with ars.
+  Qed.
+
+  Lemma bsteps_ind_r (P : nat → A → Prop) (x : A)
+    (Prefl : ∀ n, P n x)
+    (Pstep : ∀ n y z, bsteps R n x y → R y z → P n y → P (S n) z) :
+    ∀ n z, bsteps R n x z → P n z.
+  Proof.
+    cut (∀ m y z, bsteps R m y z → ∀ n,
+      bsteps R n x y →
+      (∀ m', n ≤ m' ∧ m' ≤ n + m → P m' y) →
+      P (n + m) z).
+    { intros help ?. change n with (0 + n). eauto with ars. }
+    induction 1 as [|m x' y z p2 p3 IH]; [by eauto with arith|].
+    intros n p1 H. rewrite <-plus_n_Sm.
+    apply (IH (S n)); [by eauto using bsteps_r |].
+    intros [|m'] [??]; [lia |].
+    apply Pstep with x'.
+    * apply bsteps_weaken with n; intuition lia.
+    * done.
+    * apply H; intuition lia.
+  Qed.
 
   Global Instance tc_trans: Transitive (tc R).
   Proof. red; induction 1; eauto with ars. Qed.
@@ -116,7 +142,7 @@ Section rtc.
   Proof. intros. etransitivity; eauto with ars. Qed.
   Lemma tc_rtc x y : tc R x y → rtc R x y.
   Proof. induction 1; eauto with ars. Qed.
-  Global Instance: subrelation (tc R) (rtc R).
+  Instance tc_once_subrel: subrelation (tc R) (rtc R).
   Proof. exact @tc_rtc. Qed.
 
   Lemma looping_red x : looping R x → red R x.
@@ -137,23 +163,73 @@ Section rtc.
   Qed.
 End rtc.
 
-Hint Resolve rtc_once rtc_r tc_r : ars.
+(* Avoid too eager type class resolution *)
+Hint Extern 5 (subrelation _ (rtc _)) =>
+  eapply @rtc_once_subrel : typeclass_instances.
+Hint Extern 5 (subrelation _ (tc _)) =>
+  eapply @tc_once_subrel : typeclass_instances.
+
+Hint Resolve
+  rtc_once rtc_r
+  tc_r
+  bsteps_once bsteps_r bsteps_refl bsteps_trans : ars.
 
 (** * Theorems on sub relations *)
 Section subrel.
   Context {A} (R1 R2 : relation A) (Hsub : subrelation R1 R2).
 
   Lemma red_subrel x : red R1 x → red R2 x.
-  Proof. intros [y ?]. exists y. now apply Hsub. Qed.
+  Proof. intros [y ?]. exists y. by apply Hsub. Qed.
   Lemma nf_subrel x : nf R2 x → nf R1 x.
-  Proof. intros H1 H2. destruct H1. now apply red_subrel. Qed.
-
-  Global Instance rtc_subrel: subrelation (rtc R1) (rtc R2).
-  Proof. induction 1; [left|eright]; eauto; now apply Hsub. Qed.
-  Global Instance nsteps_subrel: subrelation (nsteps R1 n) (nsteps R2 n).
-  Proof. induction 1; [left|eright]; eauto; now apply Hsub. Qed.
-  Global Instance bsteps_subrel: subrelation (bsteps R1 n) (bsteps R2 n).
-  Proof. induction 1; [left|eright]; eauto; now apply Hsub. Qed.
-  Global Instance tc_subrel: subrelation (tc R1) (tc R2).
-  Proof. induction 1; [left|eright]; eauto; now apply Hsub. Qed.
+  Proof. intros H1 H2. destruct H1. by apply red_subrel. Qed.
+
+  Instance rtc_subrel: subrelation (rtc R1) (rtc R2).
+  Proof. induction 1; [left|eright]; eauto; by apply Hsub. Qed.
+  Instance nsteps_subrel: subrelation (nsteps R1 n) (nsteps R2 n).
+  Proof. induction 1; [left|eright]; eauto; by apply Hsub. Qed.
+  Instance bsteps_subrel: subrelation (bsteps R1 n) (bsteps R2 n).
+  Proof. induction 1; [left|eright]; eauto; by apply Hsub. Qed.
+  Instance tc_subrel: subrelation (tc R1) (tc R2).
+  Proof. induction 1; [left|eright]; eauto; by apply Hsub. Qed.
 End subrel.
+
+Hint Extern 5 (subrelation (rtc _) (rtc _)) =>
+  eapply @rtc_subrel : typeclass_instances.
+Hint Extern 5 (subrelation (nsteps _) (nsteps _)) =>
+  eapply @nsteps_subrel : typeclass_instances.
+Hint Extern 5 (subrelation (bsteps _) (bsteps _)) =>
+  eapply @bsteps_subrel : typeclass_instances.
+Hint Extern 5 (subrelation (tc _) (tc _)) =>
+  eapply @tc_subrel : typeclass_instances.
+
+Notation wf := well_founded.
+
+Section wf.
+  Context `{R : relation A}.
+
+  (** A trick by Thomas Braibant to compute with well-founded recursions:
+  it lazily adds [2^n] [Acc_intro] constructors in front of a well foundedness
+  proof, so that the actual proof is never reached in practise. *)
+  Fixpoint wf_guard (n : nat) (wfR : wf R) : wf R :=
+    match n with
+    | 0 => wfR
+    | S n => λ x, Acc_intro x (λ y _, wf_guard n (wf_guard n wfR) y)
+    end.
+
+  Lemma wf_projected `(R2 : relation B) (f : A → B) :
+    (∀ x y, R x y → R2 (f x) (f y)) →
+    wf R2 → wf R.
+  Proof.
+    intros Hf Hwf.
+    cut (∀ y, Acc R2 y → ∀ x, y = f x → Acc R x).
+    { intros aux x. apply (aux (f x)); auto. }
+    induction 1 as [y _ IH]. intros x ?. subst.
+    constructor. intros. apply (IH (f y)); auto.
+  Qed.
+End wf.
+
+(* Generally we do not want [wf_guard] to be expanded (neither by tactics,
+nor by conversion tests in the kernel), but in some cases we do need it for
+computation (that is, we cannot make it opaque). We use the [Strategy]
+command to make its expanding behavior less eager. *)
+Strategy 100 [wf_guard].
diff --git a/theories/base.v b/theories/base.v
index 38fdebfd2d0681688fbd2adcc0c54d80c04007ef..6ef39bcb3924243ac9c4c3b0613eee532d07bb4e 100644
--- a/theories/base.v
+++ b/theories/base.v
@@ -12,9 +12,11 @@ Require Export Morphisms RelationClasses List Bool Utf8 Program Setoid NArith.
 (** The following coercion allows us to use Booleans as propositions. *)
 Coercion Is_true : bool >-> Sortclass.
 
-(** Ensure that [simpl] unfolds [id] and [compose] when fully applied. *)
+(** Ensure that [simpl] unfolds [id], [compose], and [flip] when fully
+applied. *)
 Arguments id _ _/.
 Arguments compose _ _ _ _ _ _ /.
+Arguments flip _ _ _ _ _ _/.
 
 (** 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
@@ -23,6 +25,9 @@ semantics. *)
 Notation "'True'" := True : type_scope.
 Notation "'False'" := False : type_scope.
 
+Notation curry := prod_curry.
+Notation uncurry := prod_uncurry.
+
 (** Throughout this development we use [C_scope] for all general purpose
 notations that do not belong to a more specific scope. *)
 Delimit Scope C_scope with C.
@@ -38,16 +43,32 @@ Notation "(≠ x )" := (λ y, y ≠ x) (only parsing) : C_scope.
 
 Hint Extern 0 (?x = ?x) => reflexivity.
 
-Notation "(→)" := (λ x y, x → y) (only parsing) : C_scope.
-Notation "( T →)" := (λ y, T → y) (only parsing) : C_scope.
-Notation "(→ T )" := (λ y, y → T) (only parsing) : C_scope.
+Notation "(→)" := (λ A B, A → B) (only parsing) : C_scope.
+Notation "( A →)" := (λ B, A → B) (only parsing) : C_scope.
+Notation "(→ B )" := (λ A, A → B) (only parsing) : C_scope.
+
 Notation "t $ r" := (t r)
   (at level 65, right associativity, only parsing) : C_scope.
+Notation "($)" := (λ f x, f x) (only parsing) : C_scope.
+Notation "($ x )" := (λ f, f x) (only parsing) : C_scope.
+
 Infix "∘" := compose : C_scope.
 Notation "(∘)" := compose (only parsing) : C_scope.
 Notation "( f ∘)" := (compose f) (only parsing) : C_scope.
 Notation "(∘ f )" := (λ g, compose g f) (only parsing) : C_scope.
 
+Notation "(∧)" := and (only parsing) : C_scope.
+Notation "( A ∧)" := (and A) (only parsing) : C_scope.
+Notation "(∧ B )" := (λ A, A ∧ B) (only parsing) : C_scope.
+
+Notation "(∨)" := or (only parsing) : C_scope.
+Notation "( A ∨)" := (or A) (only parsing) : C_scope.
+Notation "(∨ B )" := (λ A, A ∨ B) (only parsing) : C_scope.
+
+Notation "(↔)" := iff (only parsing) : C_scope.
+Notation "( A ↔)" := (iff A) (only parsing) : C_scope.
+Notation "(↔ B )" := (λ A, A ↔ B) (only parsing) : C_scope.
+
 (** Set convenient implicit arguments for [existT] and introduce notations. *)
 Arguments existT {_ _} _ _.
 Notation "x ↾ p" := (exist _ x p) (at level 20) : C_scope.
@@ -61,12 +82,12 @@ Class PropHolds (P : Prop) := prop_holds: P.
 
 Hint Extern 0 (PropHolds _) => assumption : typeclass_instances.
 Instance: Proper (iff ==> iff) PropHolds.
-Proof. now repeat intro. Qed.
+Proof. repeat intro; trivial. Qed.
 
 Ltac solve_propholds :=
   match goal with
-  | [ |- PropHolds (?P) ] => apply _
-  | [ |- ?P ] => change (PropHolds P); apply _
+  | |- PropHolds (?P) => apply _
+  | |- ?P => change (PropHolds P); apply _
   end.
 
 (** ** Decidable propositions *)
@@ -77,6 +98,28 @@ on a type [A] we write [`{∀ x y : A, Decision (x = y)}] and use it by writing
 Class Decision (P : Prop) := decide : {P} + {¬P}.
 Arguments decide _ {_}.
 
+(** ** Inhabited types *)
+(** This type class collects types that are inhabited. *)
+Class Inhabited (A : Type) : Prop := populate { _ : A }.
+Arguments populate {_} _.
+
+Instance unit_inhabited: Inhabited unit := populate ().
+Instance list_inhabited {A} : Inhabited (list A) := populate [].
+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.
+Instance sum_inhabited_l {A B} (iA : Inhabited A) : Inhabited (A + B) :=
+  match iA with
+  | populate x => populate (inl x)
+  end.
+Instance sum_inhabited_r {A B} (iB : Inhabited A) : Inhabited (A + B) :=
+  match iB with
+  | populate y => populate (inl y)
+  end.
+Instance option_inhabited {A} : Inhabited (option A) := populate None.
+
 (** ** Setoid equality *)
 (** We define an operational type class for setoid equality. This is based on
 (Spitters/van der Weegen, 2011). *)
@@ -99,13 +142,14 @@ Instance: Params (@equiv) 2.
 (for types that have an [Equiv] instance) rather than the standard Leibniz
 equality. *)
 Instance equiv_default_relation `{Equiv A} : DefaultRelation (≡) | 3.
-Hint Extern 0 (?x ≡ ?x) => reflexivity.
+Hint Extern 0 (_ ≡ _) => reflexivity.
+Hint Extern 0 (_ ≡ _) => symmetry; assumption.
 
 (** ** Operations on collections *)
-(** We define operational type classes for the standard operations and
+(** We define operational type classes for the traditional operations and
 relations on collections: the empty collection [∅], the union [(∪)],
-intersection [(∩)], difference [(∖)], and the singleton [{[_]}]
-operation, and the subset [(⊆)] and element of [(∈)] relation. *)
+intersection [(∩)], and difference [(∖)], the singleton [{[_]}], the subset
+[(⊆)] and element of [(∈)] relation, and disjointess [(⊥)]. *)
 Class Empty A := empty: A.
 Notation "∅" := empty : C_scope.
 
@@ -116,6 +160,11 @@ Notation "(∪)" := union (only parsing) : C_scope.
 Notation "( x ∪)" := (union x) (only parsing) : C_scope.
 Notation "(∪ x )" := (λ y, union y x) (only parsing) : C_scope.
 
+Definition union_list `{Empty A}
+  `{Union A} : list A → A := fold_right (∪) ∅.
+Arguments union_list _ _ _ !_ /.
+Notation "⋃ l" := (union_list l) (at level 20, format "⋃  l") : C_scope.
+
 Class Intersection A := intersection: A → A → A.
 Instance: Params (@intersection) 2.
 Infix "∩" := intersection (at level 40) : C_scope.
@@ -147,7 +196,18 @@ Notation "(⊈)" := (λ X Y, X ⊈ Y) (only parsing) : C_scope.
 Notation "( X ⊈ )" := (λ Y, X ⊈ Y) (only parsing) : C_scope.
 Notation "( ⊈ X )" := (λ Y, Y ⊈ X) (only parsing) : C_scope.
 
-Hint Extern 0 (?x ⊆ ?x) => reflexivity.
+Hint Extern 0 (_ ⊆ _) => reflexivity.
+
+Class Subset A := subset: A → A → Prop.
+Instance: Params (@subset) 2.
+Infix "⊂" := subset (at level 70) : C_scope.
+Notation "(⊂)" := subset (only parsing) : C_scope.
+Notation "( X ⊂ )" := (subset X) (only parsing) : C_scope.
+Notation "( ⊂ X )" := (λ Y, subset Y X) (only parsing) : C_scope.
+Notation "X ⊄  Y" := (¬X ⊂ Y) (at level 70) : C_scope.
+Notation "(⊄)" := (λ X Y, X ⊄ Y) (only parsing) : C_scope.
+Notation "( X ⊄ )" := (λ Y, X ⊄ Y) (only parsing) : C_scope.
+Notation "( ⊄ X )" := (λ Y, Y ⊄ X) (only parsing) : C_scope.
 
 Class ElemOf A B := elem_of: A → B → Prop.
 Instance: Params (@elem_of) 3.
@@ -167,84 +227,171 @@ Notation "(⊥)" := disjoint (only parsing) : C_scope.
 Notation "( X ⊥)" := (disjoint X) (only parsing) : C_scope.
 Notation "(⊥ X )" := (λ Y, disjoint Y X) (only parsing) : C_scope.
 
+Inductive list_disjoint `{Disjoint A} : list A → Prop :=
+  | disjoint_nil :
+     list_disjoint []
+  | disjoint_cons X Xs :
+     Forall (⊥ X) Xs →
+     list_disjoint Xs →
+     list_disjoint (X :: Xs).
+Lemma list_disjoint_cons_inv `{Disjoint A} X Xs :
+  list_disjoint (X :: Xs) →
+  Forall (⊥ X) Xs ∧ list_disjoint Xs.
+Proof. inversion_clear 1; auto. Qed.
+
+Instance generic_disjoint `{ElemOf A B} : Disjoint B | 100 :=
+  λ X Y, ∀ x, x ∉ X ∨ x ∉ Y.
+
+Class Filter A B :=
+  filter: ∀ (P : A → Prop) `{∀ x, Decision (P x)}, B → B.
+(* Arguments filter {_ _ _} _ {_} !_ / : simpl nomatch. *)
+
+(** ** Monadic operations *)
+(** We define operational type classes for the monadic operations bind, join 
+and fmap. These type classes are defined in a non-standard way by taking the
+function as a parameter of the class. For example, we define
+<<
+  Class FMapD := fmap: ∀ {A B}, (A → B) → M A → M B.
+>>
+instead of
+<<
+  Class FMap {A B} (f : A → B) := fmap: M A → M B.
+>>
+This approach allows us to define [fmap] on lists such that [simpl] unfolds it
+in the appropriate way, and so that it can be used for mutual recursion
+(the mapped function [f] is not part of the fixpoint) as well. This is a hack,
+and should be replaced by something more appropriate in future versions. *)
+
+(* 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.
+Instance: Params (@mret) 3.
+Arguments mret {_ _ _} _.
+
+Class MBindD (M : Type → Type) {A B} (f : A → M B) := mbind: M A → M B.
+Notation MBind M := (∀ {A B} (f : A → M B), MBindD M f)%type.
+Instance: Params (@mbind) 5.
+Arguments mbind {_ _ _} _ {_} !_ / : simpl nomatch.
+
+Class MJoin (M : Type → Type) := mjoin: ∀ {A}, M (M A) → M A.
+Instance: Params (@mjoin) 3.
+Arguments mjoin {_ _ _} !_ / : simpl nomatch.
+
+Class FMapD (M : Type → Type) {A B} (f : A → B) := fmap: M A → M B.
+Notation FMap M := (∀ {A B} (f : A → B), FMapD M f)%type.
+Instance: Params (@fmap) 6.
+Arguments fmap {_ _ _} _ {_} !_ / : simpl nomatch.
+
+Notation "m ≫= f" := (mbind f m) (at level 60, right associativity) : C_scope.
+Notation "( m ≫=)" := (λ f, mbind f m) (only parsing) : C_scope.
+Notation "(≫= f )" := (mbind f) (only parsing) : C_scope.
+Notation "(≫=)" := (λ m f, mbind f m) (only parsing) : C_scope.
+
+Notation "x ← y ; z" := (y ≫= (λ x : _, z))
+  (at level 65, only parsing, next at level 35, right associativity) : C_scope.
+Infix "<$>" := fmap (at level 65, right associativity) : C_scope.
+
+Class MGuard (M : Type → Type) :=
+  mguard: ∀ P {dec : Decision P} {A}, M A → M A.
+Notation "'guard' P ; o" := (mguard P o)
+  (at level 65, only parsing, next at level 35, right associativity) : C_scope.
+
 (** ** Operations on maps *)
 (** In this section we define operational type classes for the operations
 on maps. In the file [fin_maps] we will axiomatize finite maps.
 The function lookup [m !! k] should yield the element at key [k] in [m]. *)
-Class Lookup (K : Type) (M : Type → Type) :=
-  lookup: ∀ {A}, K → M A → option A.
+Class Lookup (K A M : Type) :=
+  lookup: K → M → option A.
 Instance: Params (@lookup) 4.
 
 Notation "m !! i" := (lookup i m) (at level 20) : C_scope.
 Notation "(!!)" := lookup (only parsing) : C_scope.
 Notation "( m !!)" := (λ i, lookup i m) (only parsing) : C_scope.
 Notation "(!! i )" := (lookup i) (only parsing) : C_scope.
+Arguments lookup _ _ _ _ !_ !_ / : simpl nomatch.
 
 (** The function insert [<[k:=a]>m] should update the element at key [k] with
 value [a] in [m]. *)
-Class Insert (K : Type) (M : Type → Type) :=
-  insert: ∀ {A}, K → A → M A → M A.
+Class Insert (K A M : Type) :=
+  insert: K → A → M → M.
 Instance: Params (@insert) 4.
 Notation "<[ k := a ]>" := (insert k a)
   (at level 5, right associativity, format "<[ k := a ]>") : C_scope.
+Arguments insert _ _ _ _ !_ _ !_ / : simpl nomatch.
 
 (** The function delete [delete k m] should delete the value at key [k] in
 [m]. If the key [k] is not a member of [m], the original map should be
 returned. *)
-Class Delete (K : Type) (M : Type → Type) :=
-  delete: ∀ {A}, K → M A → M A.
-Instance: Params (@delete) 4.
+Class Delete (K M : Type) :=
+  delete: K → M → M.
+Instance: Params (@delete) 3.
+Arguments delete _ _ _ !_ !_ / : simpl nomatch.
 
 (** The function [alter f k m] should update the value at key [k] using the
 function [f], which is called with the original value. *)
-Class Alter (K : Type) (M : Type → Type) :=
-  alter: ∀ {A}, (A → A) → K → M A → M A.
-Instance: Params (@alter) 4.
+Class AlterD (K A M : Type) (f : A → A) :=
+  alter: K → M → M.
+Notation Alter K A M := (∀ (f : A → A), AlterD K A M f)%type.
+Instance: Params (@alter) 5.
+Arguments alter {_ _ _} _ {_} !_ !_ / : simpl nomatch.
 
 (** The function [alter f k m] should update the value at key [k] using the
 function [f], which is called with the original value at key [k] or [None]
 if [k] is not a member of [m]. The value at [k] should be deleted if [f] 
 yields [None]. *)
-Class PartialAlter (K : Type) (M : Type → Type) :=
-  partial_alter: ∀ {A}, (option A → option A) → K → M A → M A.
+Class PartialAlter (K A M : Type) :=
+  partial_alter: (option A → option A) → K → M → M.
 Instance: Params (@partial_alter) 4.
+Arguments partial_alter _ _ _ _ _ !_ !_ / : simpl nomatch.
 
 (** The function [dom C m] should yield the domain of [m]. That is a finite
 collection of type [C] that contains the keys that are a member of [m]. *)
-Class Dom (K : Type) (M : Type → Type) :=
-  dom: ∀ {A} C `{Empty C} `{Union C} `{Singleton K C}, M A → C.
-Instance: Params (@dom) 8.
+Class Dom (K M : Type) :=
+  dom: ∀ C `{Empty C} `{Union C} `{Singleton K C}, M → C.
+Instance: Params (@dom) 7.
+Arguments dom _ _ _ _ _ _ _ !_ / : simpl nomatch.
 
 (** The function [merge f m1 m2] should merge the maps [m1] and [m2] by
 constructing a new map whose value at key [k] is [f (m1 !! k) (m2 !! k)]
 provided that [k] is a member of either [m1] or [m2].*)
-Class Merge (M : Type → Type) :=
-  merge: ∀ {A}, (option A → option A → option A) → M A → M A → M A.
+Class Merge (A M : Type) :=
+  merge: (option A → option A → option A) → M → M → M.
 Instance: Params (@merge) 3.
+Arguments merge _ _ _ _ !_ !_ / : simpl nomatch.
 
 (** We lift the insert and delete operation to lists of elements. *)
-Definition insert_list `{Insert K M} {A} (l : list (K * A)) (m : M A) : M A :=
+Definition insert_list `{Insert K A M} (l : list (K * A)) (m : M) : M :=
   fold_right (λ p, <[ fst p := snd p ]>) m l.
 Instance: Params (@insert_list) 4.
-Definition delete_list `{Delete K M} {A} (l : list K) (m : M A) : M A :=
+Definition delete_list `{Delete K M} (l : list K) (m : M) : M :=
   fold_right delete m l.
-Instance: Params (@delete_list) 4.
-
-(** The function [union_with f m1 m2] should yield the union of [m1] and [m2]
-using the function [f] to combine values of members that are in both [m1] and
-[m2]. *)
-Class UnionWith (M : Type → Type) :=
-  union_with: ∀ {A}, (A → A → A) → M A → M A → M A.
+Instance: Params (@delete_list) 3.
+
+Definition insert_consecutive `{Insert nat A M}
+    (i : nat) (l : list A) (m : M) : M :=
+  fold_right (λ x f i, <[i:=x]>(f (S i))) (λ _, m) l i.
+Instance: Params (@insert_consecutive) 3.
+
+(** The function [union_with f m1 m2] is supposed to yield the union of [m1]
+and [m2] using the function [f] to combine values of members that are in
+both [m1] and [m2]. *)
+Class UnionWith (A M : Type) :=
+  union_with: (A → A → option A) → M → M → M.
 Instance: Params (@union_with) 3.
 
-(** Similarly for the intersection and difference. *)
-Class IntersectionWith (M : Type → Type) :=
-  intersection_with: ∀ {A}, (A → A → A) → M A → M A → M A.
+(** Similarly for intersection and difference. *)
+Class IntersectionWith (A M : Type) :=
+  intersection_with: (A → A → option A) → M → M → M.
 Instance: Params (@intersection_with) 3.
-Class DifferenceWith (M : Type → Type) :=
-  difference_with: ∀ {A}, (A → A → option A) → M A → M A → M A.
+Class DifferenceWith (A M : Type) :=
+  difference_with: (A → A → option A) → M → M → M.
 Instance: Params (@difference_with) 3.
 
+Definition intersection_with_list `{IntersectionWith A M}
+  (f : A → A → option A) : M → list M → M := fold_right (intersection_with f).
+Arguments intersection_with_list _ _ _ _ _ !_ /.
+
 (** ** 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
@@ -261,6 +408,12 @@ Class RightId {A} R (i : A) (f : A → A → A) :=
   right_id: ∀ x, R (f x i) x.
 Class Associative {A} R (f : A → A → A) :=
   associative: ∀ x y z, R (f x (f y z)) (f (f x y) z).
+Class LeftAbsorb {A} R (i : A) (f : A → A → A) :=
+  left_absorb: ∀ x, R (f i x) i.
+Class RightAbsorb {A} R (i : A) (f : A → A → A) :=
+  right_absorb: ∀ x, R (f x i) i.
+Class AntiSymmetric {A} (R : A → A → Prop) :=
+  anti_symmetric: ∀ x y, R x y → R y x → x = y.
 
 Arguments injective {_ _ _ _} _ {_} _ _ _.
 Arguments idempotent {_ _} _ {_} _.
@@ -268,6 +421,44 @@ Arguments commutative {_ _ _} _ {_} _ _.
 Arguments left_id {_ _} _ _ {_} _.
 Arguments right_id {_ _} _ _ {_} _.
 Arguments associative {_ _} _ {_} _ _ _.
+Arguments left_absorb {_ _} _ _ {_} _.
+Arguments right_absorb {_ _} _ _ {_} _.
+Arguments anti_symmetric {_} _ {_} _ _ _ _.
+
+Instance: Commutative (↔) (↔).
+Proof. red. intuition. Qed.
+Instance: Commutative (↔) (∧).
+Proof. red. intuition. Qed.
+Instance: Associative (↔) (∧).
+Proof. red. intuition. Qed.
+Instance: Idempotent (↔) (∧).
+Proof. red. intuition. Qed.
+Instance: Commutative (↔) (∨).
+Proof. red. intuition. Qed.
+Instance: Associative (↔) (∨).
+Proof. red. intuition. Qed.
+Instance: Idempotent (↔) (∨).
+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.
 
 (** The following lemmas are more specific versions of the projections of the
 above type classes. These lemmas allow us to enforce Coq not to use the setoid
@@ -287,33 +478,12 @@ Proof. auto. Qed.
 Lemma associative_eq {A} (f : A → A → A) `{!Associative (=) f} x y z :
   f x (f y z) = f (f x y) z.
 Proof. auto. Qed.
-
-(** ** Monadic operations *)
-(** We do use the operation type classes for monads merely for convenient
-overloading of notations and do not formalize any theory on monads (we do not
-define a class with the monad laws). *)
-Section monad_ops.
-  Context (M : Type → Type).
-
-  Class MRet := mret: ∀ {A}, A → M A.
-  Class MBind := mbind: ∀ {A B}, (A → M B) → M A → M B.
-  Class MJoin := mjoin: ∀ {A}, M (M A) → M A.
-  Class FMap := fmap: ∀ {A B}, (A → B) → M A → M B.
-End monad_ops.
-
-Instance: Params (@mret) 3.
-Arguments mret {M MRet A} _.
-Instance: Params (@mbind) 4.
-Arguments mbind {M MBind A B} _ _.
-Instance: Params (@mjoin) 3.
-Arguments mjoin {M MJoin A} _.
-Instance: Params (@fmap) 4.
-Arguments fmap {M FMap A B} _ _.
-
-Notation "m ≫= f" := (mbind f m) (at level 60, right associativity) : C_scope.
-Notation "x ← y ; z" := (y ≫= (λ x : _, z))
-  (at level 65, next at level 35, right associativity) : C_scope.
-Infix "<$>" := fmap (at level 65, right associativity, only parsing) : C_scope.
+Lemma left_absorb_eq {A} (i : A) (f : A → A → A) `{!LeftAbsorb (=) i f} x :
+  f i x = i.
+Proof. auto. Qed.
+Lemma right_absorb_eq {A} (i : A) (f : A → A → A) `{!RightAbsorb (=) i f} x :
+  f x i = i.
+Proof. auto. Qed.
 
 (** ** Axiomatization of ordered structures *)
 (** A pre-order equiped with a smallest element. *)
@@ -321,13 +491,17 @@ Class BoundedPreOrder A `{Empty A} `{SubsetEq A} := {
   bounded_preorder :>> PreOrder (⊆);
   subseteq_empty x : ∅ ⊆ x
 }.
+Class PartialOrder A `{SubsetEq A} := {
+  po_preorder :>> PreOrder (⊆);
+  po_antisym :> AntiSymmetric (⊆)
+}.
 
 (** We do not include equality in the following interfaces so as to avoid the
 need for proofs that the  relations and operations respect setoid equality.
 Instead, we will define setoid equality in a generic way as
 [λ X Y, X ⊆ Y ∧ Y ⊆ X]. *)
 Class BoundedJoinSemiLattice A `{Empty A} `{SubsetEq A} `{Union A} := {
-  jsl_preorder :>> BoundedPreOrder A;
+  bjsl_preorder :>> BoundedPreOrder A;
   subseteq_union_l x y : x ⊆ x ∪ y;
   subseteq_union_r x y : y ⊆ x ∪ y;
   union_least x y z : x ⊆ z → y ⊆ z → x ∪ y ⊆ z
@@ -338,21 +512,28 @@ Class MeetSemiLattice A `{Empty A} `{SubsetEq A} `{Intersection A} := {
   subseteq_intersection_r x y : x ∩ y ⊆ y;
   intersection_greatest x y z : z ⊆ x → z ⊆ y → z ⊆ x ∩ y
 }.
-
+Class LowerBoundedLattice A `{Empty A} `{SubsetEq A}
+    `{Union A} `{Intersection A} := {
+  lbl_bjsl :>> BoundedJoinSemiLattice A;
+  lbl_msl :>> MeetSemiLattice A
+}.
 (** ** Axiomatization of collections *)
-(** The class [Collection A C] axiomatizes a collection of type [C] with
-elements of type [A]. Since [C] is not dependent on [A], we use the monomorphic
-[Map] type class instead of the polymorphic [FMap]. *)
-Class Map A C := map: (A → A) → (C → C).
+(** The class [SimpleCollection A C] axiomatizes a collection of type [C] with
+elements of type [A]. *)
 Instance: Params (@map) 3.
-Class Collection A C `{ElemOf A C} `{Empty C} `{Union C}
-    `{Intersection C} `{Difference C} `{Singleton A C} `{Map A C} := {
+Class SimpleCollection A C `{ElemOf A C}
+  `{Empty C} `{Singleton A C} `{Union C} := {
   not_elem_of_empty (x : A) : x ∉ ∅;
   elem_of_singleton (x y : A) : x ∈ {[ y ]} ↔ x = y;
-  elem_of_union X Y (x : A) : x ∈ X ∪ Y ↔ x ∈ X ∨ x ∈ Y;
+  elem_of_union X Y (x : A) : x ∈ X ∪ Y ↔ x ∈ X ∨ x ∈ Y
+}.
+Class Collection A C `{ElemOf A C} `{Empty C} `{Singleton A C}
+    `{Union C} `{Intersection C} `{Difference C} `{IntersectionWith A C} := {
+  collection_simple :>> SimpleCollection A C;
   elem_of_intersection X Y (x : A) : x ∈ X ∩ Y ↔ x ∈ X ∧ x ∈ Y;
   elem_of_difference X Y (x : A) : x ∈ X ∖ Y ↔ x ∈ X ∧ x ∉ Y;
-  elem_of_map f X (x : A) : x ∈ map f X ↔ ∃ y, x = f y ∧ y ∈ X
+  elem_of_intersection_with (f : A → A → option A) X Y (x : A) :
+    x ∈ intersection_with f X Y ↔ ∃ x1 x2, x1 ∈ X ∧ x2 ∈ Y ∧ f x1 x2 = Some x
 }.
 
 (** We axiomative a finite collection as a collection whose elements can be
@@ -360,21 +541,62 @@ enumerated as a list. These elements, given by the [elements] function, may be
 in any order and should not contain duplicates. *)
 Class Elements A C := elements: C → list A.
 Instance: Params (@elements) 3.
-Class FinCollection A C `{Empty C} `{Union C} `{Intersection C} `{Difference C}
-    `{Singleton A C} `{ElemOf A C} `{Map A C} `{Elements A C} := {
+
+(** We redefine the standard library's [In] and [NoDup] using type classes. *)
+Inductive elem_of_list {A} : ElemOf A (list A) :=
+  | elem_of_list_here (x : A) l : x ∈ x :: l
+  | elem_of_list_further (x y : A) l : x ∈ l → x ∈ y :: l.
+Existing Instance elem_of_list.
+
+Inductive NoDup {A} : list A → Prop :=
+  | NoDup_nil_2 : NoDup []
+  | NoDup_cons_2 x l : x ∉ l → NoDup l → NoDup (x :: l).
+
+(** Decidability of equality of the carrier set is admissible, but we add it
+anyway so as to avoid cycles in type class search. *)
+Class FinCollection A C `{ElemOf A C} `{Empty C} `{Singleton A C}
+    `{Union C} `{Intersection C} `{Difference C} `{IntersectionWith A C}
+    `{Filter A C} `{Elements A C} `{∀ x y : A, Decision (x = y)} := {
   fin_collection :>> Collection A C;
-  elements_spec X x : x ∈ X ↔ In x (elements X);
+  elem_of_filter X P `{∀ x, Decision (P x)} x :
+    x ∈ filter P X ↔ P x ∧ x ∈ X;
+  elements_spec X x : x ∈ X ↔ x ∈ elements X;
   elements_nodup X : NoDup (elements X)
 }.
 Class Size C := size: C → nat.
+Arguments size {_ _} !_ / : simpl nomatch.
 Instance: Params (@size) 2.
 
+(** The class [Collection M] axiomatizes a type constructor [M] that can be
+used to construct a collection [M A] with elements of type [A]. The advantage
+of this class, compared to [Collection], is that it also axiomatizes the
+the monadic operations. The disadvantage, is that not many inhabits are
+possible (we will only provide an inhabitant using unordered lists without
+duplicates removed). More interesting implementations typically need
+decidability of equality, or a total order on the elements, which do not fit
+in a type constructor of type [Type → Type]. *)
+Class CollectionMonad M `{∀ A, ElemOf A (M A)}
+    `{∀ A, Empty (M A)} `{∀ A, Singleton A (M A)} `{∀ A, Union (M A)}
+    `{!MBind M} `{!MRet M} `{!FMap M} `{!MJoin M} := {
+  collection_monad_simple A :> SimpleCollection A (M A);
+  elem_of_bind {A B} (f : A → M B) (X : M A) (x : B) :
+    x ∈ X ≫= f ↔ ∃ y, x ∈ f y ∧ y ∈ X;
+  elem_of_ret {A} (x y : A) :
+    x ∈ mret y ↔ x = y;
+  elem_of_fmap {A B} (f : A → B) (X : M A) (x : B) :
+    x ∈ f <$> X ↔ ∃ y, x = f y ∧ y ∈ X;
+  elem_of_join {A} (X : M (M A)) (x : A) :
+    x ∈ mjoin X ↔ ∃ Y, x ∈ Y ∧ Y ∈ X
+}.
+
 (** The function [fresh X] yields an element that is not contained in [X]. We
 will later prove that [fresh] is [Proper] with respect to the induced setoid
 equality on collections. *)
 Class Fresh A C := fresh: C → A.
 Instance: Params (@fresh) 3.
-Class FreshSpec A C `{!Fresh A C} `{!ElemOf A C} := {
+Class FreshSpec A C `{ElemOf A C}
+    `{Empty C} `{Singleton A C} `{Union C} `{Fresh A C} := {
+  fresh_collection_simple :>> SimpleCollection A C;
   fresh_proper_alt X Y : (∀ x, x ∈ X ↔ x ∈ Y) → fresh X = fresh Y;
   is_fresh (X : C) : fresh X ∉ X
 }.
@@ -382,7 +604,7 @@ Class FreshSpec A C `{!Fresh A C} `{!ElemOf A C} := {
 (** * Miscellaneous *)
 Lemma proj1_sig_inj {A} (P : A → Prop) x (Px : P x) y (Py : P y) :
   x↾Px = y↾Py → x = y.
-Proof. now injection 1. Qed.
+Proof. injection 1; trivial. Qed.
 
 Lemma symmetry_iff `(R : relation A) `{!Symmetric R} (x y : A) :
   R x y ↔ R y x.
@@ -406,6 +628,24 @@ Definition fst_map {A A' B} (f : A → A') (p : A * B) : A' * B :=
   (f (fst p), snd p).
 Definition snd_map {A B B'} (f : B → B') (p : A * B) : A * B' :=
   (fst p, f (snd p)).
+Arguments fst_map {_ _ _} _ !_ /.
+Arguments snd_map {_ _ _} _ !_ /.
+
+Instance: ∀ {A A' B} (f : A → A'),
+  Injective (=) (=) f → Injective (=) (=) (@fst_map A A' B f).
+Proof.
+  intros ????? [??] [??]; simpl; intro; f_equal.
+  * apply (injective f). congruence.
+  * congruence.
+Qed.
+Instance: ∀ {A B B'} (f : B → B'),
+  Injective (=) (=) f → Injective (=) (=) (@snd_map A B B' f).
+Proof.
+  intros ????? [??] [??]; simpl; intro; f_equal.
+  * congruence.
+  * apply (injective f). congruence.
+Qed.
+
 Definition prod_relation {A B} (R1 : relation A) (R2 : relation B) :
   relation (A * B) := λ x y, R1 (fst x) (fst y) ∧ R2 (snd x) (snd y).
 
@@ -436,29 +676,29 @@ Definition lift_relation {A B} (R : relation A)
   (f : B → A) : relation B := λ x y, R (f x) (f y).
 Definition lift_relation_equivalence {A B} (R : relation A) (f : B → A) :
   Equivalence R → Equivalence (lift_relation R f).
-Proof. unfold lift_relation. firstorder. Qed.
+Proof. unfold lift_relation. firstorder auto. Qed.
 Hint Extern 0 (Equivalence (lift_relation _ _)) =>
   eapply @lift_relation_equivalence : typeclass_instances.
 
 Instance: ∀ A B (x : B), Commutative (=) (λ _ _ : A, x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance: ∀ A (x : A), Associative (=) (λ _ _ : A, x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance: ∀ A, Associative (=) (λ x _ : A, x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance: ∀ A, Associative (=) (λ _ x : A, x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance: ∀ A, Idempotent (=) (λ x _ : A, x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance: ∀ A, Idempotent (=) (λ _ x : A, x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 
 Instance left_id_propholds {A} (R : relation A) i f :
   LeftId R i f → ∀ x, PropHolds (R (f i x) x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance right_id_propholds {A} (R : relation A) i f :
   RightId R i f → ∀ x, PropHolds (R (f x i) x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
 Instance idem_propholds {A} (R : relation A) f :
   Idempotent R f → ∀ x, PropHolds (R (f x x) x).
-Proof. easy. Qed.
+Proof. red. trivial. Qed.
diff --git a/theories/collections.v b/theories/collections.v
index 8cef0c32b6f07cb99ac6b82487ca4a1db7d9917d..b96b6836cd89a3c643b5cd559c8e2f92c0b068ae 100644
--- a/theories/collections.v
+++ b/theories/collections.v
@@ -6,91 +6,116 @@ collections. *)
 Require Export base tactics orders.
 
 (** * Theorems *)
-Section collection.
-  Context `{Collection A B}.
+Section simple_collection.
+  Context `{SimpleCollection A C}.
 
   Lemma elem_of_empty x : x ∈ ∅ ↔ False.
-  Proof. split. apply not_elem_of_empty. easy. Qed.
+  Proof. split. apply not_elem_of_empty. done. Qed.
   Lemma elem_of_union_l x X Y : x ∈ X → x ∈ X ∪ Y.
   Proof. intros. apply elem_of_union. auto. Qed.
   Lemma elem_of_union_r x X Y : x ∈ Y → x ∈ X ∪ Y.
   Proof. intros. apply elem_of_union. auto. Qed.
-  Lemma not_elem_of_singleton x y : x ∉ {[ y ]} ↔ x ≠ y.
-  Proof. now rewrite elem_of_singleton. Qed.
-  Lemma not_elem_of_union x X Y : x ∉ X ∪ Y ↔ x ∉ X ∧ x ∉ Y.
-  Proof. rewrite elem_of_union. tauto. Qed.
 
-  Global Instance collection_subseteq: SubsetEq B := λ X Y,
+  Global Instance collection_subseteq: SubsetEq C := λ X Y,
     ∀ x, x ∈ X → x ∈ Y.
-  Global Instance: BoundedJoinSemiLattice B.
-  Proof. firstorder. Qed.
-  Global Instance: MeetSemiLattice B.
-  Proof. firstorder. Qed.
+  Global Instance: BoundedJoinSemiLattice C.
+  Proof. firstorder auto. Qed.
 
   Lemma elem_of_subseteq X Y : X ⊆ Y ↔ ∀ x, x ∈ X → x ∈ Y.
-  Proof. easy. Qed.
+  Proof. done. Qed.
   Lemma elem_of_equiv X Y : X ≡ Y ↔ ∀ x, x ∈ X ↔ x ∈ Y.
   Proof. firstorder. Qed.
   Lemma elem_of_equiv_alt X Y :
     X ≡ Y ↔ (∀ x, x ∈ X → x ∈ Y) ∧ (∀ x, x ∈ Y → x ∈ X).
   Proof. firstorder. Qed.
-
+  Lemma elem_of_subseteq_singleton x X : x ∈ X ↔ {[ x ]} ⊆ X.
+  Proof.
+    split.
+    * intros ??. rewrite elem_of_singleton. intro. by subst.
+    * intros Ex. by apply (Ex x), elem_of_singleton.
+  Qed.
   Global Instance singleton_proper : Proper ((=) ==> (≡)) singleton.
-  Proof. repeat intro. now subst. Qed.
-  Global Instance elem_of_proper: Proper ((=) ==> (≡) ==> iff) (∈).
+  Proof. repeat intro. by subst. Qed.
+  Global Instance elem_of_proper: Proper ((=) ==> (≡) ==> iff) (∈) | 5.
   Proof. intros ???. subst. firstorder. Qed.
 
-  Lemma empty_ne_singleton x : ∅ ≢ {[ x ]}.
+  Lemma elem_of_union_list (Xs : list C) (x : A) :
+    x ∈ ⋃ Xs ↔ ∃ X, X ∈ Xs ∧ x ∈ X.
   Proof.
-    intros [_ E]. apply (elem_of_empty x).
-    apply E. now apply elem_of_singleton.
+    split.
+    * induction Xs; simpl; intros HXs.
+      + by apply elem_of_empty in HXs.
+      + setoid_rewrite elem_of_cons.
+        apply elem_of_union in HXs. naive_solver.
+    * intros [X []]. induction 1; simpl.
+      + by apply elem_of_union_l.
+      + intros. apply elem_of_union_r; auto.
   Qed.
-End collection.
 
-(** * Theorems about map *)
-Section map.
-  Context `{Collection A C}.
+  Lemma non_empty_singleton x : {[ x ]} ≢ ∅.
+  Proof. intros [E _]. by apply (elem_of_empty x), E, elem_of_singleton. Qed.
 
-  Lemma elem_of_map_1 (f : A → A) (X : C) (x : A) :
-    x ∈ X → f x ∈ map f X.
-  Proof. intros. apply (elem_of_map _). eauto. Qed.
-  Lemma elem_of_map_1_alt (f : A → A) (X : C) (x : A) y :
-    x ∈ X → y = f x → y ∈ map f X.
-  Proof. intros. apply (elem_of_map _). eauto. Qed.
-  Lemma elem_of_map_2 (f : A → A) (X : C) (x : A) :
-    x ∈ map f X → ∃ y, x = f y ∧ y ∈ X.
-  Proof. intros. now apply (elem_of_map _). Qed.
-End map.
+  Lemma not_elem_of_singleton x y : x ∉ {[ y ]} ↔ x ≠ y.
+  Proof. by rewrite elem_of_singleton. Qed.
+  Lemma not_elem_of_union x X Y : x ∉ X ∪ Y ↔ x ∉ X ∧ x ∉ Y.
+  Proof. rewrite elem_of_union. tauto. Qed.
+
+  Context `{∀ X Y : C, Decision (X ⊆ Y)}.
+
+  Global Instance elem_of_dec_slow (x : A) (X : C) : Decision (x ∈ X) | 100.
+  Proof.
+    refine (cast_if (decide_rel (⊆) {[ x ]} X));
+      by rewrite elem_of_subseteq_singleton.
+  Defined.
+End simple_collection.
+
+Ltac decompose_empty := repeat
+  match goal with
+  | H : _ ∪ _ ≡ ∅ |- _ => apply empty_union in H; destruct H
+  | H : _ ∪ _ ≢ ∅ |- _ => apply non_empty_union in H; destruct H
+  | H : {[ _ ]} ≡ ∅ |- _ => destruct (non_empty_singleton _ H)
+  end.
 
 (** * Tactics *)
 (** The first pass consists of eliminating all occurrences of [(∪)], [(∩)],
 [(∖)], [map], [∅], [{[_]}], [(≡)], and [(⊆)], by rewriting these into
 logically equivalent propositions. For example we rewrite [A → x ∈ X ∪ ∅] into
 [A → x ∈ X ∨ False]. *)
-Ltac unfold_elem_of := repeat
-  match goal with
-  | H : context [ _ ⊆ _ ] |- _ => setoid_rewrite elem_of_subseteq in H
-  | H : context [ _ ≡ _ ] |- _ => setoid_rewrite elem_of_equiv_alt in H
-  | H : context [ _ ∈ ∅ ] |- _ => setoid_rewrite elem_of_empty in H
-  | H : context [ _ ∈ {[ _ ]} ] |- _ => setoid_rewrite elem_of_singleton in H
-  | H : context [ _ ∈ _ ∪ _ ] |- _ => setoid_rewrite elem_of_union in H
-  | H : context [ _ ∈ _ ∩ _ ] |- _ => setoid_rewrite elem_of_intersection in H
-  | H : context [ _ ∈ _ ∖ _ ] |- _ => setoid_rewrite elem_of_difference in H
-  | H : context [ _ ∈ map _ _ ] |- _ => setoid_rewrite elem_of_map in H
+Ltac unfold_elem_of :=
+  repeat_on_hyps (fun H =>
+    repeat match type of H with
+    | context [ _ ⊆ _ ] => setoid_rewrite elem_of_subseteq in H
+    | context [ _ ⊂ _ ] => setoid_rewrite subset_spec in H
+    | context [ _ ≡ _ ] => setoid_rewrite elem_of_equiv_alt in H
+    | context [ _ ∈ ∅ ] => setoid_rewrite elem_of_empty in H
+    | context [ _ ∈ {[ _ ]} ] => setoid_rewrite elem_of_singleton in H
+    | context [ _ ∈ _ ∪ _ ] => setoid_rewrite elem_of_union in H
+    | context [ _ ∈ _ ∩ _ ] => setoid_rewrite elem_of_intersection in H
+    | context [ _ ∈ _ ∖ _ ] => setoid_rewrite elem_of_difference in H
+    | context [ _ ∈ _ <$> _ ] => setoid_rewrite elem_of_fmap in H
+    | context [ _ ∈ mret _ ] => setoid_rewrite elem_of_ret in H
+    | context [ _ ∈ _ ≫= _ ] => setoid_rewrite elem_of_bind in H
+    | context [ _ ∈ mjoin _ ] => setoid_rewrite elem_of_join in H
+    end);
+  repeat match goal with
   | |- context [ _ ⊆ _ ] => setoid_rewrite elem_of_subseteq
+  | |- context [ _ ⊂ _ ] => setoid_rewrite subset_spec
   | |- context [ _ ≡ _ ] => setoid_rewrite elem_of_equiv_alt
   | |- context [ _ ∈ ∅ ] => setoid_rewrite elem_of_empty
   | |- context [ _ ∈ {[ _ ]} ] => setoid_rewrite elem_of_singleton
   | |- context [ _ ∈ _ ∪ _ ] => setoid_rewrite elem_of_union
   | |- context [ _ ∈ _ ∩ _ ] => setoid_rewrite elem_of_intersection
   | |- context [ _ ∈ _ ∖ _ ] => setoid_rewrite elem_of_difference
-  | |- context [ _ ∈ map _ _ ] => setoid_rewrite elem_of_map
+  | |- context [ _ ∈ _ <$> _ ] => setoid_rewrite elem_of_fmap
+  | |- context [ _ ∈ mret _ ] => setoid_rewrite elem_of_ret
+  | |- context [ _ ∈ _ ≫= _ ] => setoid_rewrite elem_of_bind
+  | |- context [ _ ∈ mjoin _ ] => setoid_rewrite elem_of_join
   end.
 
 (** The tactic [solve_elem_of tac] composes the above tactic with [intuition].
 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" tactic(tac) :=
+Tactic Notation "solve_elem_of" tactic3(tac) :=
   simpl in *;
   unfold_elem_of;
   solve [intuition (simplify_equality; tac)].
@@ -101,19 +126,20 @@ Tactic Notation "solve_elem_of" := solve_elem_of auto.
 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" tactic(tac) :=
+Tactic Notation "esolve_elem_of" tactic3(tac) :=
   simpl in *;
   unfold_elem_of;
   naive_solver tac.
 Tactic Notation "esolve_elem_of" := esolve_elem_of eauto.
 
-(** Given an assumption [H : _ ∈ _], the tactic [destruct_elem_of H] will
+(** Given a hypothesis [H : _ ∈ _], the tactic [destruct_elem_of H] will
 recursively split [H] for [(∪)], [(∩)], [(∖)], [map], [∅], [{[_]}]. *)
-Tactic Notation "destruct_elem_of" hyp(H) :=
+Tactic Notation "decompose_elem_of" hyp(H) :=
   let rec go H :=
   lazymatch type of H with
   | _ ∈ ∅ => apply elem_of_empty in H; destruct H
-  | _ ∈ {[ ?l' ]} => apply elem_of_singleton in H; subst l'
+  | ?x ∈ {[ ?y ]} =>
+    apply elem_of_singleton in H; try first [subst y | subst x]
   | _ ∈ _ ∪ _ =>
     let H1 := fresh in let H2 := fresh in apply elem_of_union in H;
     destruct H as [H1|H2]; [go H1 | go H2]
@@ -123,21 +149,104 @@ Tactic Notation "destruct_elem_of" hyp(H) :=
   | _ ∈ _ ∖ _ =>
     let H1 := fresh in let H2 := fresh in apply elem_of_difference in H;
     destruct H as [H1 H2]; go H1; go H2
-  | _ ∈ map _ _ =>
-    let H1 := fresh in apply elem_of_map in H;
-    destruct H as [?[? H1]]; go H1
+  | ?x ∈ _ <$> _ =>
+    let H1 := fresh in apply elem_of_fmap in H;
+    destruct H as [? [? H1]]; try (subst x); go H1
+  | _ ∈ _ ≫= _ =>
+    let H1 := fresh in let H2 := fresh in apply elem_of_bind in H;
+    destruct H as [? [H1 H2]]; go H1; go H2
+  | ?x ∈ mret ?y =>
+    apply elem_of_ret in H; try first [subst y | subst x]
+  | _ ∈ mjoin _ ≫= _ =>
+    let H1 := fresh in let H2 := fresh in apply elem_of_join in H;
+    destruct H as [? [H1 H2]]; go H1; go H2
   | _ => idtac
   end in go H.
+Tactic Notation "decompose_elem_of" :=
+  repeat_on_hyps (fun H => decompose_elem_of H).
+
+Section collection.
+  Context `{Collection A C}.
+
+  Global Instance: LowerBoundedLattice C.
+  Proof. split. apply _. firstorder auto. Qed.
+
+  Lemma intersection_singletons x : {[x]} ∩ {[x]} ≡ {[x]}.
+  Proof. esolve_elem_of. Qed.
+  Lemma difference_twice X Y : (X ∖ Y) ∖ Y ≡ X ∖ Y.
+  Proof. esolve_elem_of. Qed.
+
+  Lemma empty_difference X Y : X ⊆ Y → X ∖ Y ≡ ∅.
+  Proof. esolve_elem_of. Qed.
+  Lemma difference_diag X : X ∖ X ≡ ∅.
+  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_intersection_distr_l X Y Z : (X ∩ Y) ∖ Z ≡ X ∖ Z ∩ Y ∖ Z.
+  Proof. esolve_elem_of. Qed.
+
+  Lemma elem_of_intersection_with_list (f : A → A → option A) Xs Y x :
+    x ∈ intersection_with_list f Y Xs ↔ ∃ xs y,
+      Forall2 (∈) xs Xs ∧ y ∈ Y ∧ foldr (λ x, (≫= f x)) (Some y) xs = Some x.
+  Proof.
+    split.
+    * revert x. induction Xs; simpl; intros x HXs.
+      + eexists [], x. intuition.
+      + rewrite elem_of_intersection_with in HXs.
+        destruct HXs as (x1 & x2 & Hx1 & Hx2 & ?).
+        destruct (IHXs x2) as (xs & y & hy & ? & ?); trivial.
+        eexists (x1 :: xs), y. intuition (simplify_option_equality; auto).
+    * intros (xs & y & Hxs & ? & Hx). revert x Hx.
+      induction Hxs; intros; simplify_option_equality; [done |].
+      rewrite elem_of_intersection_with. naive_solver.
+  Qed.
+
+  Lemma intersection_with_list_ind (P Q : A → Prop) f Xs Y :
+    (∀ y, y ∈ Y → P y) →
+    Forall (λ X, ∀ x, x ∈ X → Q x) Xs →
+    (∀ x y z, Q x → P y → f x y = Some z → P z) →
+    ∀ x, x ∈ intersection_with_list f Y Xs → P x.
+  Proof.
+    intros HY HXs Hf.
+    induction Xs; simplify_option_equality; [done |].
+    intros x Hx. rewrite elem_of_intersection_with in Hx.
+    decompose_Forall. destruct Hx as (? & ? & ? & ? & ?). eauto.
+  Qed.
+
+  Context `{∀ X Y : C, Decision (X ⊆ Y)}.
+
+  Lemma not_elem_of_intersection x X Y : x ∉ X ∩ Y ↔ x ∉ X ∨ x ∉ Y.
+  Proof.
+    rewrite elem_of_intersection.
+    destruct (decide (x ∈ X)); tauto.
+  Qed.
+  Lemma not_elem_of_difference x X Y : x ∉ X ∖ Y ↔ x ∉ X ∨ x ∈ Y.
+  Proof.
+    rewrite elem_of_difference.
+    destruct (decide (x ∈ Y)); tauto.
+  Qed.
+  Lemma union_difference X Y : X ⊆ Y → Y ≡ X ∪ Y ∖ X.
+  Proof.
+    split; intros x; rewrite !elem_of_union, elem_of_difference.
+    * destruct (decide (x ∈ X)); intuition.
+    * intuition.
+  Qed.
+  Lemma non_empty_difference X Y : X ⊂ Y → Y ∖ X ≢ ∅.
+  Proof.
+    intros [HXY1 HXY2] Hdiff. destruct HXY2. intros x.
+    destruct (decide (x ∈ X)); esolve_elem_of.
+  Qed.
+End collection.
 
 (** * Sets without duplicates up to an equivalence *)
 Section no_dup.
-  Context `{Collection A B} (R : relation A) `{!Equivalence R}.
+  Context `{SimpleCollection A B} (R : relation A) `{!Equivalence R}.
 
   Definition elem_of_upto (x : A) (X : B) := ∃ y, y ∈ X ∧ R x y.
   Definition no_dup (X : B) := ∀ x y, x ∈ X → y ∈ X → R x y → x = y.
 
   Global Instance: Proper ((≡) ==> iff) (elem_of_upto x).
-  Proof. firstorder. Qed.
+  Proof. intros ??? E. unfold elem_of_upto. by setoid_rewrite E. Qed.
   Global Instance: Proper (R ==> (≡) ==> iff) elem_of_upto.
   Proof.
     intros ?? E1 ?? E2. split; intros [z [??]]; exists z.
@@ -177,7 +286,7 @@ End no_dup.
 
 (** * Quantifiers *)
 Section quantifiers.
-  Context `{Collection A B} (P : A → Prop).
+  Context `{SimpleCollection A B} (P : A → Prop).
 
   Definition cforall X := ∀ x, x ∈ X → P x.
   Definition cexists X := ∃ x, x ∈ X ∧ P x.
@@ -208,10 +317,10 @@ End quantifiers.
 Section more_quantifiers.
   Context `{Collection A B}.
 
-  Lemma cforall_weak (P Q : A → Prop) (Hweak : ∀ x, P x → Q x) X :
+  Lemma cforall_weaken (P Q : A → Prop) (Hweaken : ∀ x, P x → Q x) X :
     cforall P X → cforall Q X.
   Proof. unfold cforall. naive_solver. Qed.
-  Lemma cexists_weak (P Q : A → Prop) (Hweak : ∀ x, P x → Q x) X :
+  Lemma cexists_weaken (P Q : A → Prop) (Hweaken : ∀ x, P x → Q x) X :
     cexists P X → cexists Q X.
   Proof. unfold cexists. naive_solver. Qed.
 End more_quantifiers.
@@ -220,13 +329,13 @@ End more_quantifiers.
 (** We collect some properties on the [fresh] operation. In particular we
 generalize [fresh] to generate lists of fresh elements. *)
 Section fresh.
-  Context `{Collection A C} `{Fresh A C} `{!FreshSpec A C} .
+  Context `{FreshSpec A C} .
 
   Definition fresh_sig (X : C) : { x : A | x ∉ X } :=
     exist (∉ X) (fresh X) (is_fresh X).
 
   Global Instance fresh_proper: Proper ((≡) ==> (=)) fresh.
-  Proof. intros ???. now apply fresh_proper_alt, elem_of_equiv. Qed.
+  Proof. intros ???. by apply fresh_proper_alt, elem_of_equiv. Qed.
 
   Fixpoint fresh_list (n : nat) (X : C) : list A :=
     match n with
@@ -238,18 +347,18 @@ Section fresh.
   Proof.
     intros ? n ?. subst.
     induction n; simpl; intros ?? E; f_equal.
-    * now rewrite E.
-    * apply IHn. now rewrite E.
+    * by rewrite E.
+    * apply IHn. by rewrite E.
   Qed.
 
   Lemma fresh_list_length n X : length (fresh_list n X) = n.
   Proof. revert X. induction n; simpl; auto. Qed.
 
-  Lemma fresh_list_is_fresh n X x : In x (fresh_list n X) → x ∉ X.
+  Lemma fresh_list_is_fresh n X x : x ∈ fresh_list n X → x ∉ X.
   Proof.
-    revert X. induction n; simpl.
-    * easy.
-    * intros X [?| Hin]. subst.
+    revert X. induction n; intros X; simpl.
+    * by rewrite elem_of_nil.
+    * rewrite elem_of_cons. intros [?| Hin]; subst.
       + apply is_fresh.
       + apply IHn in Hin. solve_elem_of.
   Qed.
@@ -262,3 +371,76 @@ Section fresh.
     solve_elem_of.
   Qed.
 End fresh.
+
+Definition option_collection `{Singleton A C} `{Empty C} (x : option A) : C :=
+  match x with
+  | None => ∅
+  | Some a => {[ a ]}
+  end.
+
+Section collection_monad.
+  Context `{CollectionMonad M}.
+
+  Global Instance collection_guard: MGuard M := λ P dec A x,
+    if dec then x else ∅.
+
+  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.
+  Global Instance collection_bind_proper {A B} (f : A → M B) :
+    Proper ((≡) ==> (≡)) (mbind f).
+  Proof. intros X Y E. esolve_elem_of. Qed.
+  Global Instance collection_join_proper {A} :
+    Proper ((≡) ==> (≡)) (@mjoin M _ A).
+  Proof. intros X Y E. esolve_elem_of. Qed.
+
+  Lemma collection_fmap_compose {A B C} (f : A → B) (g : B → C) X :
+    g ∘ f <$> X ≡ g <$> (f <$> X).
+  Proof. esolve_elem_of. Qed.
+  Lemma elem_of_fmap_1 {A B} (f : A → B) (X : M A) (y : B) :
+    y ∈ f <$> X → ∃ x, y = f x ∧ x ∈ X.
+  Proof. esolve_elem_of. Qed.
+  Lemma elem_of_fmap_2 {A B} (f : A → B) (X : M A) (x : A) :
+    x ∈ X → f x ∈ f <$> X.
+  Proof. esolve_elem_of. Qed.
+  Lemma elem_of_fmap_2_alt {A B} (f : A → B) (X : M A) (x : A) (y : B) :
+    x ∈ X → y = f x → y ∈ f <$> X.
+  Proof. esolve_elem_of. Qed.
+
+  Lemma elem_of_mapM {A B} (f : A → M B) l k :
+    l ∈ mapM f k ↔ Forall2 (λ x y, x ∈ f y) l k.
+  Proof.
+    split.
+    * revert l. induction k; esolve_elem_of.
+    * induction 1; esolve_elem_of.
+  Qed.
+  Lemma mapM_length {A B} (f : A → M B) l k :
+    l ∈ mapM f k → length l = length k.
+  Proof. revert l; induction k; esolve_elem_of. Qed.
+
+  Lemma elem_of_mapM_fmap {A B} (f : A → B) (g : B → M A) l k :
+    Forall (λ x, ∀ y, y ∈ g x → f y = x) l →
+    k ∈ mapM g l → fmap f k = l.
+  Proof.
+    intros Hl. revert k.
+    induction Hl; simpl; intros;
+      decompose_elem_of; simpl; f_equal; auto.
+  Qed.
+
+  Lemma elem_of_mapM_Forall {A B} (f : A → M B) (P : B → Prop) l k :
+    l ∈ mapM f k →
+    Forall (λ x, ∀ y, y ∈ f x → P y) k →
+    Forall P l.
+  Proof. rewrite elem_of_mapM. apply Forall2_Forall_l. Qed.
+  Lemma elem_of_mapM_Forall2_l {A B C} (f : A → M B) (P : B → C → Prop) l1 l2 k :
+    l1 ∈ mapM f k →
+    Forall2 (λ x y, ∀ z, z ∈ f x → P z y) k l2 →
+    Forall2 P l1 l2.
+  Proof.
+    rewrite elem_of_mapM. intros Hl1. revert l2.
+    induction Hl1; inversion_clear 1; constructor; auto.
+  Qed.
+End collection_monad.
diff --git a/theories/decidable.v b/theories/decidable.v
index 3bdfbf967dd79a1e01bc9e5b9a08974153fec38c..9ea1ccbd5cbf37456ee90720aa775cb41d89c666 100644
--- a/theories/decidable.v
+++ b/theories/decidable.v
@@ -3,7 +3,12 @@
 (** This file collects theorems, definitions, tactics, related to propositions
 with a decidable equality. Such propositions are collected by the [Decision]
 type class. *)
-Require Export base.
+Require Export base tactics.
+
+Hint Extern 200 (Decision _) => progress (lazy beta) : typeclass_instances.
+
+Lemma dec_stable `{Decision P} : ¬¬P → P.
+Proof. firstorder. Qed.
 
 (** We introduce [decide_rel] to avoid inefficienct computation due to eager
 evaluation of propositions by [vm_compute]. This inefficiency occurs if
@@ -14,10 +19,10 @@ Definition decide_rel {A B} (R : A → B → Prop) {dec : ∀ x y, Decision (R x
   (x : A) (y : B) : Decision (R x y) := dec x y.
 Lemma decide_rel_correct {A B} (R : A → B → Prop) `{∀ x y, Decision (R x y)}
   (x : A) (y : B) : decide_rel R x y = decide (R x y).
-Proof. easy. Qed.
+Proof. done. Qed.
 
 (** The tactic [case_decide] performs case analysis on an arbitrary occurrence
-of [decide] or [decide_rel] in the conclusion or assumptions. *)
+of [decide] or [decide_rel] in the conclusion or hypotheses. *)
 Ltac case_decide :=
   match goal with
   | H : context [@decide ?P ?dec] |- _ =>
@@ -34,21 +39,21 @@ Ltac case_decide :=
 with instance resolution to automatically generate decision procedures. *)
 Ltac solve_trivial_decision :=
   match goal with
-  | [ |- Decision (?P) ] => apply _
-  | [ |- sumbool ?P (¬?P) ] => change (Decision P); apply _
+  | |- Decision (?P) => apply _
+  | |- sumbool ?P (¬?P) => change (Decision P); apply _
   end.
-Ltac solve_decision :=
-  intros; first [ solve_trivial_decision
-                | unfold Decision; decide equality; solve_trivial_decision ].
+Ltac solve_decision := intros; first
+  [ solve_trivial_decision
+  | unfold Decision; decide equality; solve_trivial_decision ].
 
 (** We can convert decidable propositions to booleans. *)
 Definition bool_decide (P : Prop) {dec : Decision P} : bool :=
   if dec then true else false.
 
 Lemma bool_decide_unpack (P : Prop) {dec : Decision P} : bool_decide P → P.
-Proof. unfold bool_decide. now destruct dec. Qed.
+Proof. unfold bool_decide. by destruct dec. Qed.
 Lemma bool_decide_pack (P : Prop) {dec : Decision P} : P → bool_decide P.
-Proof. unfold bool_decide. now destruct dec. Qed.
+Proof. unfold bool_decide. by destruct dec. Qed.
 
 (** * Decidable Sigma types *)
 (** Leibniz equality on Sigma types requires the equipped proofs to be
@@ -70,38 +75,57 @@ Proof.
   * intro. destruct x as [x Hx], y as [y Hy].
     simpl in *. subst. f_equal.
     revert Hx Hy. case (bool_decide (P y)).
-    + now intros [] [].
-    + easy.
+    + by intros [] [].
+    + done.
 Qed.
 
+(** The following combinators are useful to create Decision proofs in
+combination with the [refine] tactic. *)
+Notation cast_if S := (if S then left _ else right _).
+Notation cast_if_and S1 S2 := (if S1 then cast_if S2 else right _).
+Notation cast_if_and3 S1 S2 S3 := (if S1 then cast_if_and S2 S3 else right _).
+Notation cast_if_and4 S1 S2 S3 S4 :=
+  (if S1 then cast_if_and3 S2 S3 S4 else right _).
+Notation cast_if_or S1 S2 := (if S1 then left _ else cast_if S2).
+Notation cast_if_not S := (if S then right _ else left _).
+
 (** * Instances of Decision *)
 (** Instances of [Decision] for operators of propositional logic. *)
-Program Instance True_dec: Decision True := left _.
-Program Instance False_dec: Decision False := right _.
-Program Instance and_dec `(P_dec : Decision P) `(Q_dec : Decision Q) :
-    Decision (P ∧ Q) :=
-  match P_dec with
-  | left _ => match Q_dec with left _ => left _ | right _ => right _ end
-  | right _ => right _
-  end.
-Solve Obligations using intuition.
-Program Instance or_dec `(P_dec : Decision P) `(Q_dec : Decision Q) :
-    Decision (P ∨ Q) :=
-  match P_dec with
-  | left _ => left _
-  | right _ => match Q_dec with left _ => left _ | right _ => right _ end
-  end.
-Solve Obligations using intuition.
+Instance True_dec: Decision True := left I.
+Instance False_dec: Decision False := right (False_rect False).
+
+Section prop_dec.
+  Context `(P_dec : Decision P) `(Q_dec : Decision Q).
+
+  Global Instance not_dec: Decision (¬P).
+  Proof. refine (cast_if_not P_dec); intuition. Defined.
+  Global Instance and_dec: Decision (P ∧ Q).
+  Proof. refine (cast_if_and P_dec Q_dec); intuition. Defined.
+  Global Instance or_dec: Decision (P ∨ Q).
+  Proof. refine (cast_if_or P_dec Q_dec); intuition. Defined.
+  Global Instance impl_dec: Decision (P → Q).
+  Proof. refine (if P_dec then cast_if Q_dec else left _); intuition. Defined.
+End prop_dec.
 
 (** Instances of [Decision] for common data types. *)
-Program Instance prod_eq_dec `(A_dec : ∀ x y : A, Decision (x = y))
-    `(B_dec : ∀ x y : B, Decision (x = y)) (x y : A * B) : Decision (x = y) :=
-  match A_dec (fst x) (fst y) with
-  | left _ =>
-    match B_dec (snd x) (snd y) with
-    | left _ => left _
-    | right _ => right _
-    end
-  | right _ => right _
+Instance bool_eq_dec (x y : bool) : Decision (x = y).
+Proof. solve_decision. Defined.
+Instance unit_eq_dec (x y : unit) : Decision (x = y).
+Proof. refine (left _); by destruct x, y. Defined.
+Instance prod_eq_dec `(A_dec : ∀ x y : A, Decision (x = y))
+  `(B_dec : ∀ x y : B, Decision (x = y)) (x y : A * B) : Decision (x = y).
+Proof.
+  refine (cast_if_and (A_dec (fst x) (fst y)) (B_dec (snd x) (snd y)));
+    abstract (destruct x, y; simpl in *; congruence).
+Defined.
+Instance sum_eq_dec `(A_dec : ∀ x y : A, Decision (x = y))
+  `(B_dec : ∀ x y : B, Decision (x = y)) (x y : A + B) : Decision (x = y).
+Proof. solve_decision. Defined.
+
+Instance curry_dec `(P_dec : ∀ (x : A) (y : B), Decision (P x y)) p :
+    Decision (curry P p) :=
+  match p as p return Decision (curry P p) with
+  | (x,y) => P_dec x y
   end.
-Solve Obligations using intuition (simpl;congruence).
+Instance uncurry_dec `(P_dec : ∀ (p : A * B), Decision (P p)) x y :
+  Decision (uncurry P x y) := P_dec (x,y).
diff --git a/theories/fin_collections.v b/theories/fin_collections.v
index 06caeb1784d064725ee799ae937d5928b400152b..35808597bc1ab411f30f29422ec9b88371ca8ad4 100644
--- a/theories/fin_collections.v
+++ b/theories/fin_collections.v
@@ -3,12 +3,12 @@
 (** 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.
-Require Export collections listset.
+Require Import Permutation ars.
+Require Export collections numbers listset.
 
-Instance collection_size `{Elements A C} : Size C := λ X, length (elements X).
-Definition collection_fold `{Elements A C} {B} (f : A → B → B)
-  (b : B) (X : C) : B := fold_right f b (elements X).
+Instance collection_size `{Elements A C} : Size C := length ∘ elements.
+Definition collection_fold `{Elements A C} {B}
+  (f : A → B → B) (b : B) : C → B := foldr f b ∘ elements.
 
 Section fin_collection.
 Context `{FinCollection A C}.
@@ -18,145 +18,145 @@ Proof.
   intros ?? E. apply NoDup_Permutation.
   * apply elements_nodup.
   * apply elements_nodup.
-  * intros. now rewrite <-!elements_spec, E.
+  * intros. by rewrite <-!elements_spec, E.
 Qed.
 Global Instance collection_size_proper: Proper ((≡) ==> (=)) size.
-Proof. intros ?? E. apply Permutation_length. now rewrite E. Qed.
+Proof. intros ?? E. apply Permutation_length. by rewrite E. Qed.
 
-Lemma size_empty : size ∅ = 0.
+Lemma size_empty : size (∅ : C) = 0.
 Proof.
-  unfold size, collection_size. rewrite (in_nil_inv (elements ∅)).
-  * easy.
+  unfold size, collection_size. simpl.
+  rewrite (elem_of_nil_inv (elements ∅)).
+  * done.
   * intro. rewrite <-elements_spec. solve_elem_of.
 Qed.
-Lemma size_empty_inv X : size X = 0 → X ≡ ∅.
+Lemma size_empty_inv (X : C) : size X = 0 → X ≡ ∅.
 Proof.
   intros. apply equiv_empty. intro. rewrite elements_spec.
-  rewrite (nil_length (elements X)); intuition.
+  rewrite (nil_length (elements X)). by rewrite elem_of_nil. done.
 Qed.
-Lemma size_empty_iff X : size X = 0 ↔ X ≡ ∅.
-Proof. split. apply size_empty_inv. intros E. now rewrite E, size_empty. Qed.
+Lemma size_empty_iff (X : C) : size X = 0 ↔ X ≡ ∅.
+Proof. split. apply size_empty_inv. intros E. by rewrite E, size_empty. Qed.
+Lemma size_non_empty_iff (X : C) : size X ≠ 0 ↔ X ≢ ∅.
+Proof. by rewrite size_empty_iff. Qed.
 
-Lemma size_singleton x : size {[ x ]} = 1.
+Lemma size_singleton (x : A) : size {[ x ]} = 1.
 Proof.
   change (length (elements {[ x ]}) = length [x]).
   apply Permutation_length, NoDup_Permutation.
   * apply elements_nodup.
   * apply NoDup_singleton.
-  * intros. rewrite <-elements_spec. esolve_elem_of firstorder.
+  * intros.
+    by rewrite <-elements_spec, elem_of_singleton, elem_of_list_singleton.
 Qed.
 Lemma size_singleton_inv X x y : size X = 1 → x ∈ X → y ∈ X → x = y.
 Proof.
-  unfold size, collection_size. rewrite !elements_spec.
+  unfold size, collection_size. simpl. rewrite !elements_spec.
   generalize (elements X). intros [|? l].
-  * discriminate.
-  * injection 1. intro. rewrite (nil_length l) by easy.
-    simpl. intuition congruence.
+  * done.
+  * injection 1. intro. rewrite (nil_length l) by done.
+    simpl. rewrite !elem_of_list_singleton. congruence.
 Qed.
 
-Lemma choose X : X ≢ ∅ → { x | x ∈ X }.
+Lemma elem_of_or_empty X : (∃ x, x ∈ X) ∨ X ≡ ∅.
 Proof.
-  destruct (elements X) as [|x l] eqn:E.
-  * intros []. apply equiv_empty.
-    intros x. rewrite elements_spec, E. contradiction.
-  * exists x. rewrite elements_spec, E. now left.
+  destruct (elements X) as [|x xs] eqn:E.
+  * right. apply equiv_empty. intros x Ex.
+    by rewrite elements_spec, E, elem_of_nil in Ex.
+  * left. exists x. rewrite elements_spec, E.
+    by constructor.
 Qed.
-Lemma size_pos_choose X : 0 < size X → { x | x ∈ X }.
+
+Lemma choose X : X ≢ ∅ → ∃ x, x ∈ X.
+Proof.
+  destruct (elem_of_or_empty X) as [[x ?]|?].
+  * by exists x.
+  * done.
+Qed.
+Lemma size_pos_choose X : 0 < size X → ∃ x, x ∈ X.
 Proof.
-  intros E. apply choose.
-  intros E2. rewrite E2, size_empty in E.
-  now destruct (Lt.lt_n_0 0).
+  intros E1. apply choose.
+  intros E2. rewrite E2, size_empty in E1.
+  by apply (Lt.lt_n_0 0).
 Qed.
-Lemma size_1_choose X : size X = 1 → { x | X ≡ {[ x ]} }.
+Lemma size_1_choose X : size X = 1 → ∃ x, X ≡ {[ x ]}.
 Proof.
   intros E. destruct (size_pos_choose X).
   * rewrite E. auto with arith.
   * exists x. apply elem_of_equiv. split.
-    + intro. rewrite elem_of_singleton. eauto using size_singleton_inv.
+    + intro. rewrite elem_of_singleton.
+      eauto using size_singleton_inv.
     + solve_elem_of.
 Qed.
 
-Program Instance collection_car_eq_dec_slow (x y : A) :
-    Decision (x = y) | 100 :=
-  match Compare_dec.zerop (size ({[ x ]} ∩ {[ y ]})) with
-  | left _ => right _
-  | right _ => left _
-  end.
-Next Obligation.
-  intro. apply empty_ne_singleton with x.
-  transitivity ({[ x ]} ∩ {[ y ]}).
-  * symmetry. now apply size_empty_iff.
-  * solve_elem_of.
-Qed.
-Next Obligation. edestruct size_pos_choose; esolve_elem_of. Qed.
-
-Instance elem_of_dec_slow (x : A) (X : C) : Decision (x ∈ X) | 100 :=
-  match decide_rel In x (elements X) with
-  | left Hx => left (proj2 (elements_spec _ _) Hx)
-  | right Hx => right (Hx ∘ proj1 (elements_spec _ _))
-  end.
-
-Lemma union_difference X Y : X ∪ Y ∖ X ≡ X ∪ Y.
-Proof. split; intros x; destruct (decide (x ∈ X)); solve_elem_of. Qed.
-
 Lemma size_union X Y : X ∩ Y ≡ ∅ → size (X ∪ Y) = size X + size Y.
 Proof.
-  intros [E _]. unfold size, collection_size. rewrite <-app_length.
+  intros [E _]. unfold size, collection_size. simpl. rewrite <-app_length.
   apply Permutation_length, NoDup_Permutation.
   * apply elements_nodup.
-  * apply NoDup_app; try apply elements_nodup.
+  * apply NoDup_app; repeat split; try apply elements_nodup.
     intros x. rewrite <-!elements_spec. esolve_elem_of.
-  * intros. rewrite in_app_iff, <-!elements_spec. solve_elem_of.
+  * intros. rewrite elem_of_app, <-!elements_spec. solve_elem_of.
 Qed.
-Lemma size_union_alt X Y : size (X ∪ Y) = size X + size (Y ∖ X).
+
+Instance elem_of_dec_slow (x : A) (X : C) : Decision (x ∈ X) | 100.
 Proof.
-  rewrite <-size_union. now rewrite union_difference. solve_elem_of.
+  refine (cast_if (decide_rel (∈) x (elements X)));
+    by rewrite (elements_spec _).
+Defined.
+Global Program Instance collection_subseteq_dec_slow (X Y : C) :
+    Decision (X ⊆ Y) | 100 :=
+  match decide_rel (=) (size (X ∖ Y)) 0 with
+  | left E1 => left _
+  | right E1 => right _
+  end.
+Next Obligation.
+  intros x Ex; apply dec_stable; intro.
+  destruct (proj1 (elem_of_empty x)).
+  apply (size_empty_inv _ E1).
+  by rewrite elem_of_difference.
 Qed.
-Lemma size_add X x : x ∉ X → size ({[ x ]} ∪ X) = S (size X).
-Proof.
-  intros. rewrite size_union. now rewrite size_singleton. solve_elem_of.
+Next Obligation.
+  intros E2. destruct E1.
+  apply size_empty_iff, equiv_empty. intros x.
+  rewrite elem_of_difference. intros [E3 ?].
+  by apply E2 in E3.
 Qed.
-Lemma size_difference X Y : X ⊆ Y → size X + size (Y ∖ X) = size Y.
-Proof. intros. now rewrite <-size_union_alt, subseteq_union_1. Qed.
-Lemma size_remove X x : x ∈ X → S (size (X ∖ {[ x ]})) = size X.
+
+Lemma size_union_alt X Y : size (X ∪ Y) = size X + size (Y ∖ X).
 Proof.
-  intros. rewrite <-(size_difference {[ x ]} X).
-  * rewrite size_singleton. auto with arith.
-  * solve_elem_of.
+  rewrite <-size_union by solve_elem_of.
+  setoid_replace (Y ∖ X) with ((Y ∪ X) ∖ X) by esolve_elem_of.
+  rewrite <-union_difference, (commutative (∪)); solve_elem_of.
 Qed.
 
 Lemma subseteq_size X Y : X ⊆ Y → size X ≤ size Y.
 Proof.
-  intros. rewrite <-(subseteq_union_1 X Y) by easy.
-  rewrite <-(union_difference X Y), size_union by solve_elem_of.
-  auto with arith.
+  intros. rewrite (union_difference X Y), size_union_alt by done. lia.
 Qed.
-
-Lemma collection_wf_ind (P : C → Prop) :
-  (∀ X, (∀ Y, size Y < size X → P Y) → P X) →
-  ∀ X, P X.
+Lemma subset_size X Y : X ⊂ Y → size X < size Y.
 Proof.
-  intros Hind. cut (∀ n X, size X < n → P X).
-  { intros help X. apply help with (S (size X)). auto with arith. }
-  induction n; intros.
-  * now destruct (Lt.lt_n_0 (size X)).
-  * apply Hind. intros. apply IHn. eauto with arith.
+  intros. rewrite (union_difference X Y) by solve_elem_of.
+  rewrite size_union_alt, difference_twice.
+  cut (size (Y ∖ X) ≠ 0); [lia |].
+  by apply size_non_empty_iff, non_empty_difference.
 Qed.
 
+Lemma collection_wf : wf (@subset C _).
+Proof. apply well_founded_lt_compat with size, subset_size. Qed.
+
 Lemma collection_ind (P : C → Prop) :
   Proper ((≡) ==> iff) P →
   P ∅ →
   (∀ x X, x ∉ X → P X → P ({[ x ]} ∪ X)) →
   ∀ X, P X.
 Proof.
-  intros ? Hemp Hadd. apply collection_wf_ind.
-  intros X IH. destruct (Compare_dec.zerop (size X)).
-  * now rewrite size_empty_inv.
-  * destruct (size_pos_choose X); auto.
-    rewrite <-(subseteq_union_1 {[ x ]} X) by solve_elem_of.
-    rewrite <-union_difference.
-    apply Hadd; [solve_elem_of |]. apply IH.
-    rewrite <-(size_remove X x); auto with arith.
+  intros ? Hemp Hadd. apply well_founded_induction with (⊂).
+  { apply collection_wf. }
+  intros X IH. destruct (elem_of_or_empty X) as [[x ?]|HX].
+  * rewrite (union_difference {[ x ]} X) by solve_elem_of.
+    apply Hadd. solve_elem_of. apply IH. esolve_elem_of.
+  * by rewrite HX.
 Qed.
 
 Lemma collection_fold_ind {B} (P : B → C → Prop) (f : A → B → B) (b : B) :
@@ -166,55 +166,44 @@ Lemma collection_fold_ind {B} (P : B → C → Prop) (f : A → B → B) (b : B)
   ∀ X, P (collection_fold f b X) X.
 Proof.
   intros ? Hemp Hadd.
-  cut (∀ l, NoDup l → ∀ X, (∀ x, x ∈ X ↔ In x l) → P (fold_right f b l) X).
+  cut (∀ l, NoDup l → ∀ X, (∀ x, x ∈ X ↔ x ∈ l) → P (foldr f b l) X).
   { intros help ?. apply help. apply elements_nodup. apply elements_spec. }
-  induction 1 as [|x l ?? IHl]; simpl.
-  * intros X HX. rewrite equiv_empty. easy. intros ??. firstorder.
-  * intros X HX.
-    rewrite <-(subseteq_union_1 {[ x ]} X) by esolve_elem_of.
-    rewrite <-union_difference.
-    apply Hadd. solve_elem_of. apply IHl.
-    intros y. split.
-    + intros. destruct (proj1 (HX y)); solve_elem_of.
-    + esolve_elem_of.
-Qed.
-
-Lemma collection_fold_proper {B} (f : A → B → B) (b : B) :
-  (∀ a1 a2 b, f a1 (f a2 b) = f a2 (f a1 b)) →
-  Proper ((≡) ==> (=)) (collection_fold f b).
-Proof. intros ??? E. apply fold_right_permutation. auto. now rewrite E. Qed.
-
-Global Program Instance cforall_dec `(P : A → Prop)
-    `{∀ x, Decision (P x)} X : Decision (cforall P X) | 100 :=
-  match decide (Forall P (elements X)) with
-  | left Hall => left _
-  | right Hall => right _
-  end.
-Next Obligation.
-  red. setoid_rewrite elements_spec. now apply Forall_forall.
-Qed.
-Next Obligation.
-  intro. apply Hall, Forall_forall. setoid_rewrite <-elements_spec. auto.
+  induction 1 as [|x l ?? IH]; simpl.
+  * intros X HX. setoid_rewrite elem_of_nil in HX.
+    rewrite equiv_empty. done. esolve_elem_of.
+  * intros X HX. setoid_rewrite elem_of_cons in HX.
+    rewrite (union_difference {[ x ]} X) by esolve_elem_of.
+    apply Hadd. solve_elem_of. apply IH. esolve_elem_of.
 Qed.
 
-Global Program Instance cexists_dec `(P : A → Prop)
-    `{∀ x, Decision (P x)} X : Decision (cexists P X) | 100 :=
-  match decide (Exists P (elements X)) with
-  | left Hex => left _
-  | right Hex => right _
-  end.
-Next Obligation.
-  red. setoid_rewrite elements_spec. now apply Exists_exists.
-Qed.
-Next Obligation.
-  intro. apply Hex, Exists_exists. setoid_rewrite <-elements_spec. auto.
+Lemma collection_fold_proper {B} (R : relation B)
+    `{!Equivalence R}
+    (f : A → B → B) (b : B)
+    `{!Proper ((=) ==> R ==> R) f}
+    (Hf : ∀ a1 a2 b, R (f a1 (f a2 b)) (f a2 (f a1 b))) :
+  Proper ((≡) ==> R) (collection_fold f b).
+Proof.
+  intros ?? E. apply (foldr_permutation R f b).
+  * auto.
+  * by rewrite E.
 Qed.
 
+Global Instance cforall_dec `(P : A → Prop)
+  `{∀ x, Decision (P x)} X : Decision (cforall P X) | 100.
+Proof.
+  refine (cast_if (decide (Forall P (elements X))));
+    abstract (unfold cforall; setoid_rewrite elements_spec;
+      by rewrite <-Forall_forall).
+Defined.
+
+Global Instance cexists_dec `(P : A → Prop) `{∀ x, Decision (P x)} X :
+  Decision (cexists P X) | 100.
+Proof.
+  refine (cast_if (decide (Exists P (elements X))));
+    abstract (unfold cexists; setoid_rewrite elements_spec;
+      by rewrite <-Exists_exists).
+Defined.
+
 Global Instance rel_elem_of_dec `{∀ x y, Decision (R x y)} x X :
   Decision (elem_of_upto R x X) | 100 := decide (cexists (R x) X).
-
-Lemma not_elem_of_intersection x X Y : x ∉ X ∩ Y ↔ x ∉ X ∨ x ∉ Y.
-Proof. destruct (decide (x ∈ X)); solve_elem_of. Qed.
-Lemma not_elem_of_difference x X Y : x ∉ X ∖ Y ↔ x ∉ X ∨ x ∈ Y.
-Proof. destruct (decide (x ∈ Y)); solve_elem_of. Qed.
 End fin_collection.
diff --git a/theories/fin_maps.v b/theories/fin_maps.v
index 073b46b0b5117e229c0baed52f1c70462272cd16..3f6adef99b6d3fc0ac04f93e94f9ca0dc48f221b 100644
--- a/theories/fin_maps.v
+++ b/theories/fin_maps.v
@@ -4,8 +4,8 @@
 finite maps and collects some theory on it. Most importantly, it proves useful
 induction principles for finite maps and implements the tactic [simplify_map]
 to simplify goals involving finite maps. *)
+Require Import Permutation.
 Require Export prelude.
-
 (** * Axiomatization of finite maps *)
 (** We require Leibniz equality to be extensional on finite maps. This of
 course limits the space of finite map implementations, but since we are mainly
@@ -13,14 +13,24 @@ interested in finite maps with numbers as indexes, we do not consider this to
 be a serious limitation. The main application of finite maps is to implement
 the memory, where extensionality of Leibniz equality is very important for a
 convenient use in the assertions of our axiomatic semantics. *)
-(** Finiteness is axiomatized by requiring each map to have a finite domain.
-Since we may have multiple implementations of finite sets, the [dom] function is
-parametrized by an implementation of finite sets over the map's key type. *)
+
+(** Finiteness is axiomatized by requiring that each map can be translated
+to an association list. The translation to association lists is used to
+implement the [dom] function, and for well founded recursion on finite maps. *)
+
 (** Finite map implementations are required to implement the [merge] function
 which enables us to give a generic implementation of [union_with],
 [intersection_with], and [difference_with]. *)
-Class FinMap K M `{∀ A, Empty (M A)} `{Lookup K M} `{FMap M}
-    `{PartialAlter K M} `{Dom K M} `{Merge M} := {
+
+Class FinMapToList K A M := finmap_to_list: M → list (K * A).
+
+Class FinMap K M `{!FMap M}
+    `{∀ A, Lookup K A (M A)}
+    `{∀ A, Empty (M A)}
+    `{∀ A, PartialAlter K A (M A)}
+    `{∀ A, Merge A (M A)}
+    `{∀ A, FinMapToList K A (M A)}
+    `{∀ i j : K, Decision (i = j)} := {
   finmap_eq {A} (m1 m2 : M A) :
     (∀ i, m1 !! i = m2 !! i) → m1 = m2;
   lookup_empty {A} i :
@@ -31,8 +41,10 @@ Class FinMap K M `{∀ A, Empty (M A)} `{Lookup K M} `{FMap M}
     i ≠ j → partial_alter f i m !! j = m !! j;
   lookup_fmap {A B} (f : A → B) (m : M A) i :
     (f <$> m) !! i = f <$> m !! i;
-  elem_of_dom C {A} `{Collection K C} (m : M A) i :
-    i ∈ dom C m ↔ is_Some (m !! i);
+  finmap_to_list_nodup {A} (m : M A) :
+    NoDup (finmap_to_list m);
+  elem_of_finmap_to_list {A} (m : M A) i x :
+    (i,x) ∈ finmap_to_list m ↔ m !! i = Some x;
   merge_spec {A} f `{!PropHolds (f None None = None)}
     (m1 m2 : M A) i : merge f m1 m2 !! i = f (m1 !! i) (m2 !! i)
 }.
@@ -40,431 +52,1246 @@ Class FinMap K M `{∀ A, Empty (M A)} `{Lookup K M} `{FMap M}
 (** * Derived operations *)
 (** All of the following functions are defined in a generic way for arbitrary
 finite map implementations. These generic implementations do not cause a
-significant enough performance loss to make including them in the finite map
-axiomatization worthwhile. *)
-Instance finmap_alter `{PartialAlter K M} : Alter K M := λ A f,
+significant performance loss to make including them in the finite map interface
+worthwhile. *)
+Instance finmap_insert `{PartialAlter K A M} : Insert K A M := λ i x,
+  partial_alter (λ _, Some x) i.
+Instance finmap_alter `{PartialAlter K A M} : Alter K A M := λ f,
   partial_alter (fmap f).
-Instance finmap_insert `{PartialAlter K M} : Insert K M := λ A k x,
-  partial_alter (λ _, Some x) k.
-Instance finmap_delete `{PartialAlter K M} : Delete K M := λ A,
+Instance finmap_delete `{PartialAlter K A M} : Delete K M :=
   partial_alter (λ _, None).
-Instance finmap_singleton `{PartialAlter K M} {A}
-  `{Empty (M A)} : Singleton (K * A) (M A) := λ p, <[fst p:=snd p]>∅.
+Instance finmap_singleton `{PartialAlter K A M}
+  `{Empty M} : Singleton (K * A) M := λ p, <[fst p:=snd p]>∅.
 
-Definition list_to_map `{Insert K M} {A} `{Empty (M A)}
-  (l : list (K * A)) : M A := insert_list l ∅.
+Definition finmap_of_list `{Insert K A M} `{Empty M}
+  (l : list (K * A)) : M := insert_list l ∅.
+Instance finmap_dom `{FinMapToList K A M} : Dom K M := λ C _ _ _,
+  foldr ((∪) ∘ singleton ∘ fst) ∅ ∘ finmap_to_list.
 
-Instance finmap_union_with `{Merge M} : UnionWith M := λ A f,
+Instance finmap_union_with `{Merge A M} : UnionWith A M := λ f,
   merge (union_with f).
-Instance finmap_intersection_with `{Merge M} : IntersectionWith M := λ A f,
+Instance finmap_intersection_with `{Merge A M} : IntersectionWith A M := λ f,
   merge (intersection_with f).
-Instance finmap_difference_with `{Merge M} : DifferenceWith M := λ A f,
+Instance finmap_difference_with `{Merge A M} : DifferenceWith A M := λ f,
   merge (difference_with f).
 
-(** Two finite maps are disjoint if they do not have overlapping cells. *)
-Instance finmap_disjoint `{Lookup K M} {A} : Disjoint (M A) := λ m1 m2,
-  ∀ i, m1 !! i = None ∨ m2 !! i = None.
+(** The relation [intersection_forall R] on finite maps describes that the
+relation [R] holds for each pair in the intersection. *)
+Definition finmap_forall `{Lookup K A M} (P : K → A → Prop) : M → Prop :=
+  λ m, ∀ i x, m !! i = Some x → P i x.
+Definition finmap_intersection_forall `{Lookup K A M}
+    (R : relation A) : relation M := λ m1 m2,
+  ∀ i x1 x2, m1 !! i = Some x1 → m2 !! i = Some x2 → R x1 x2.
+Instance finmap_disjoint `{∀ A, Lookup K A (M A)} : Disjoint (M A) := λ A,
+  finmap_intersection_forall (λ _ _, False).
 
 (** The union of two finite maps only has a meaningful definition for maps
 that are disjoint. However, as working with partial functions is inconvenient
 in Coq, we define the union as a total function. In case both finite maps
 have a value at the same index, we take the value of the first map. *)
-Instance finmap_union `{Merge M} {A} : Union (M A) := union_with (λ x _ , x).
+Instance finmap_union `{Merge A M} : Union M :=
+  union_with (λ x _, Some x).
+Instance finmap_intersection `{Merge A M} : Intersection M :=
+  union_with (λ x _, Some x).
+(** The difference operation removes all values from the first map whose
+index contains a value in the second map as well. *)
+Instance finmap_difference `{Merge A M} : Difference M :=
+  difference_with (λ _ _, None).
 
 (** * General theorems *)
-Section finmap.
-Context `{FinMap K M} `{∀ i j : K, Decision (i = j)} {A : Type}.
-
-Global Instance finmap_subseteq: SubsetEq (M A) := λ m n,
-  ∀ i x, m !! i = Some x → n !! i = Some x.
-Global Instance: BoundedPreOrder (M A).
-Proof. split. firstorder. intros m i x. rewrite lookup_empty. discriminate. Qed.
-
-Lemma lookup_subseteq_Some (m1 m2 : M A) i x :
-  m1 ⊆ m2 → m1 !! i = Some x → m2 !! i = Some x.
-Proof. auto. Qed.
-Lemma lookup_subseteq_None (m1 m2 : M A) i :
-  m1 ⊆ m2 → m2 !! i = None → m1 !! i = None.
-Proof. rewrite !eq_None_not_Some. firstorder. Qed.
-Lemma lookup_ne (m : M A) i j : m !! i ≠ m !! j → i ≠ j.
-Proof. congruence. Qed.
-
-Lemma not_elem_of_dom C `{Collection K C} (m : M A) i :
-  i ∉ dom C m ↔ m !! i = None.
-Proof. now rewrite (elem_of_dom C), eq_None_not_Some. Qed.
-
-Lemma finmap_empty (m : M A) : (∀ i, m !! i = None) → m = ∅.
-Proof. intros Hm. apply finmap_eq. intros. now rewrite Hm, lookup_empty. Qed.
-Lemma dom_empty C `{Collection K C} : dom C (∅ : M A) ≡ ∅.
-Proof.
-  split; intro.
-  * rewrite (elem_of_dom C), lookup_empty. simplify_is_Some.
-  * solve_elem_of.
-Qed.
-Lemma dom_empty_inv C `{Collection K C} (m : M A) : dom C m ≡ ∅ → m = ∅.
-Proof.
-  intros E. apply finmap_empty. intros. apply (not_elem_of_dom C).
-  rewrite E. solve_elem_of.
-Qed.
-
-Lemma lookup_empty_not i : ¬is_Some ((∅ : M A) !! i).
-Proof. rewrite lookup_empty. simplify_is_Some. Qed.
-Lemma lookup_empty_Some i (x : A) : ¬∅ !! i = Some x.
-Proof. rewrite lookup_empty. discriminate. Qed.
-
-Lemma partial_alter_compose (m : M A) i f g :
-  partial_alter (f ∘ g) i m = partial_alter f i (partial_alter g i m).
-Proof.
-  intros. apply finmap_eq. intros ii. case (decide (i = ii)).
-  * intros. subst. now rewrite !lookup_partial_alter.
-  * intros. now rewrite !lookup_partial_alter_ne.
-Qed.
-Lemma partial_alter_comm (m : M A) i j f g :
-  i ≠ j →
- partial_alter f i (partial_alter g j m) = partial_alter g j (partial_alter f i m).
-Proof.
-  intros. apply finmap_eq. intros jj.
-  destruct (decide (jj = j)).
-  * subst. now rewrite lookup_partial_alter_ne,
-     !lookup_partial_alter, lookup_partial_alter_ne.
-  * destruct (decide (jj = i)).
-    + subst. now rewrite lookup_partial_alter,
-       !lookup_partial_alter_ne, lookup_partial_alter by congruence.
-    + now rewrite !lookup_partial_alter_ne by congruence.
-Qed.
-Lemma partial_alter_self_alt (m : M A) i x :
-  x = m !! i → partial_alter (λ _, x) i m = m.
-Proof.
-  intros. apply finmap_eq. intros ii.
-  destruct (decide (i = ii)).
-  * subst. now rewrite lookup_partial_alter.
-  * now rewrite lookup_partial_alter_ne.
-Qed.
-Lemma partial_alter_self (m : M A) i : partial_alter (λ _, m !! i) i m = m.
-Proof. now apply partial_alter_self_alt. Qed.
-
-Lemma lookup_insert (m : M A) i x : <[i:=x]>m !! i = Some x.
-Proof. unfold insert. apply lookup_partial_alter. Qed.
-Lemma lookup_insert_rev (m : M A) i x y : <[i:= x ]>m !! i = Some y → x = y.
-Proof. rewrite lookup_insert. congruence. Qed.
-Lemma lookup_insert_ne (m : M A) i j x : i ≠ j → <[i:=x]>m !! j = m !! j.
-Proof. unfold insert. apply lookup_partial_alter_ne. Qed.
-Lemma insert_comm (m : M A) i j x y :
-  i ≠ j → <[i:=x]>(<[j:=y]>m) = <[j:=y]>(<[i:=x]>m).
-Proof. apply partial_alter_comm. Qed.
-
-Lemma lookup_insert_Some (m : M A) i j x y :
-  <[i:=x]>m !! j = Some y ↔ (i = j ∧ x = y) ∨ (i ≠ j ∧ m !! j = Some y).
-Proof.
-  split.
-  * destruct (decide (i = j)); subst;
-      rewrite ?lookup_insert, ?lookup_insert_ne; intuition congruence.
-  * intros [[??]|[??]].
-    + subst. apply lookup_insert.
-    + now rewrite lookup_insert_ne.
-Qed.
-Lemma lookup_insert_None (m : M A) i j x :
-  <[i:=x]>m !! j = None ↔ m !! j = None ∧ i ≠ j.
-Proof.
-  split.
-  * destruct (decide (i = j)); subst;
-      rewrite ?lookup_insert, ?lookup_insert_ne; intuition congruence.
-  * intros [??]. now rewrite lookup_insert_ne.
-Qed.
-
-Lemma lookup_singleton_Some i j (x y : A) :
-  {[(i, x)]} !! j = Some y ↔ i = j ∧ x = y.
-Proof.
-  unfold singleton, finmap_singleton.
-  rewrite lookup_insert_Some, lookup_empty. simpl.
-  intuition congruence.
-Qed.
-Lemma lookup_singleton_None i j (x : A) :
-  {[(i, x)]} !! j = None ↔ i ≠ j.
-Proof.
-  unfold singleton, finmap_singleton.
-  rewrite lookup_insert_None, lookup_empty. simpl. tauto.
-Qed.
-
-Lemma lookup_singleton i (x : A) : {[(i, x)]} !! i = Some x.
-Proof. rewrite lookup_singleton_Some. tauto. Qed.
-Lemma lookup_singleton_ne i j (x : A) : i ≠ j → {[(i, x)]} !! j = None.
-Proof. now rewrite lookup_singleton_None. Qed.
-
-Lemma lookup_delete (m : M A) i : delete i m !! i = None.
-Proof. apply lookup_partial_alter. Qed.
-Lemma lookup_delete_ne (m : M A) i j : i ≠ j → delete i m !! j = m !! j.
-Proof. apply lookup_partial_alter_ne. Qed.
-
-Lemma lookup_delete_Some (m : M A) i j y :
-  delete i m !! j = Some y ↔ i ≠ j ∧ m !! j = Some y.
-Proof.
-  split.
-  * destruct (decide (i = j)); subst;
-      rewrite ?lookup_delete, ?lookup_delete_ne; intuition congruence.
-  * intros [??]. now rewrite lookup_delete_ne.
-Qed.
-Lemma lookup_delete_None (m : M A) i j :
-  delete i m !! j = None ↔ i = j ∨ m !! j = None.
-Proof.
-  destruct (decide (i = j)).
-  * subst. rewrite lookup_delete. tauto.
-  * rewrite lookup_delete_ne; tauto.
-Qed.
-
-Lemma delete_empty i : delete i (∅ : M A) = ∅.
-Proof. rewrite <-(partial_alter_self ∅) at 2. now rewrite lookup_empty. Qed.
-Lemma delete_singleton i (x : A) : delete i {[(i, x)]} = ∅.
-Proof. setoid_rewrite <-partial_alter_compose. apply delete_empty. Qed.
-Lemma delete_comm (m : M A) i j : delete i (delete j m) = delete j (delete i m).
-Proof. destruct (decide (i = j)). now subst. now apply partial_alter_comm. Qed.
-Lemma delete_insert_comm (m : M A) i j x :
-  i ≠ j → delete i (<[j:=x]>m) = <[j:=x]>(delete i m).
-Proof. intro. now apply partial_alter_comm. Qed.
-
-Lemma delete_notin (m : M A) i : m !! i = None → delete i m = m.
-Proof.
-  intros. apply finmap_eq. intros j.
-  destruct (decide (i = j)).
-  * subst. now rewrite lookup_delete.
-  * now apply lookup_delete_ne.
-Qed.
-
-Lemma delete_partial_alter (m : M A) i f :
-  m !! i = None → delete i (partial_alter f i m) = m.
-Proof.
-  intros. unfold delete, finmap_delete. rewrite <-partial_alter_compose.
-  rapply partial_alter_self_alt. congruence.
-Qed.
-Lemma delete_partial_alter_dom C `{Collection K C} (m : M A) i f :
-  i ∉ dom C m → delete i (partial_alter f i m) = m.
-Proof. rewrite (not_elem_of_dom C). apply delete_partial_alter. Qed.
-Lemma delete_insert (m : M A) i x : m !! i = None → delete i (<[i:=x]>m) = m.
-Proof. apply delete_partial_alter. Qed.
-Lemma delete_insert_dom C `{Collection K C} (m : M A) i x :
-  i ∉ dom C m → delete i (<[i:=x]>m) = m.
-Proof. rewrite (not_elem_of_dom C). apply delete_partial_alter. Qed.
-Lemma insert_delete (m : M A) i x : m !! i = Some x → <[i:=x]>(delete i m) = m.
-Proof.
-  intros Hmi. unfold delete, finmap_delete, insert, finmap_insert.
-  rewrite <-partial_alter_compose. unfold compose. rewrite <-Hmi.
-  now apply partial_alter_self_alt.
-Qed.
-
-Lemma elem_of_dom_delete C `{Collection K C} (m : M A) i j :
-  i ∈ dom C (delete j m) ↔ i ≠ j ∧ i ∈ dom C m.
-Proof.
-  rewrite !(elem_of_dom C). unfold is_Some.
-  setoid_rewrite lookup_delete_Some. firstorder auto.
-Qed.
-Lemma not_elem_of_dom_delete C `{Collection K C} (m : M A) i :
-  i ∉ dom C (delete i m).
-Proof. apply (not_elem_of_dom C), lookup_delete. Qed.
-
-(** * Induction principles *)
-(** We use the induction principle on finite collections to prove the
-following induction principle on finite maps. *)
-Lemma finmap_ind_alt C (P : M A → Prop) `{FinCollection K C} :
-  P ∅ →
-  (∀ i x m, i ∉ dom C m → P m → P (<[i:=x]>m)) →
-  ∀ m, P m.
-Proof.
-  intros Hemp Hinsert m.
-  apply (collection_ind (λ X, ∀ m, dom C m ≡ X → P m)) with (dom C m).
-  * solve_proper.
-  * clear m. intros m Hm. rewrite finmap_empty.
-    + easy.
-    + intros. rewrite <-(not_elem_of_dom C), Hm.
-      now solve_elem_of.
-  * clear m. intros i X Hi IH m Hdom.
-    assert (is_Some (m !! i)) as [x Hx].
-    { apply (elem_of_dom C).
-      rewrite Hdom. clear Hdom.
-      now solve_elem_of. }
-    rewrite <-(insert_delete m i x) by easy.
-    apply Hinsert.
-    { now apply (not_elem_of_dom_delete C). }
-    apply IH. apply elem_of_equiv. intros.
-    rewrite (elem_of_dom_delete C).
-    esolve_elem_of.
-  * easy.
-Qed.
-
-(** We use the [listset] implementation to prove an induction principle that
-does not mention the map's domain. *)
-Lemma finmap_ind (P : M A → Prop) :
-  P ∅ →
-  (∀ i x m, m !! i = None → P m → P (<[i:=x]>m)) →
-  ∀ m, P m.
-Proof.
-  setoid_rewrite <-(not_elem_of_dom (listset _)).
-  apply (finmap_ind_alt (listset _) P).
-Qed.
-
-(** * Deleting and inserting multiple elements *)
-Lemma lookup_delete_list (m : M A) is j :
-  In j is → delete_list is m !! j = None.
-Proof.
-  induction is as [|i is]; simpl; [easy |].
-  intros [?|?].
-  * subst. now rewrite lookup_delete.
-  * destruct (decide (i = j)).
-    + subst. now rewrite lookup_delete.
-    + rewrite lookup_delete_ne; auto.
-Qed.
-Lemma lookup_delete_list_notin (m : M A) is j :
-  ¬In j is → delete_list is m !! j = m !! j.
-Proof.
-  induction is; simpl; [easy |].
-  intros. rewrite lookup_delete_ne; tauto.
-Qed.
-
-Lemma delete_list_notin (m : M A) is :
-  Forall (λ i, m !! i = None) is → delete_list is m = m.
-Proof.
-  induction 1; simpl; [easy |].
-  rewrite delete_notin; congruence.
-Qed.
-Lemma delete_list_insert_comm (m : M A) is j x :
-  ¬In j is → delete_list is (<[j:=x]>m) = <[j:=x]>(delete_list is m).
-Proof.
-  induction is; simpl; [easy |].
-  intros. rewrite IHis, delete_insert_comm; tauto.
-Qed.
-
-Lemma lookup_insert_list (m : M A) l1 l2 i x :
-  (∀y, ¬In (i,y) l1) → insert_list (l1 ++ (i,x) :: l2) m !! i = Some x.
-Proof.
-  induction l1 as [|[j y] l1 IH]; simpl.
-  * intros. now rewrite lookup_insert.
-  * intros Hy. rewrite lookup_insert_ne; naive_solver.
-Qed.
-
-Lemma lookup_insert_list_not_in (m : M A) l i :
-  (∀y, ¬In (i,y) l) → insert_list l m !! i = m !! i.
-Proof.
-  induction l as [|[j y] l IH]; simpl.
-  * easy.
-  * intros Hy. rewrite lookup_insert_ne; naive_solver.
-Qed.
-
-(** * Properties of the merge operation *)
-Section merge.
-  Context (f : option A → option A → option A).
-
-  Global Instance: LeftId (=) None f → LeftId (=) ∅ (merge f).
-  Proof.
-    intros ??. apply finmap_eq. intros.
-    now rewrite !(merge_spec f), lookup_empty, (left_id None f).
-  Qed.
-  Global Instance: RightId (=) None f → RightId (=) ∅ (merge f).
-  Proof.
-    intros ??. apply finmap_eq. intros.
-    now rewrite !(merge_spec f), lookup_empty, (right_id None f).
-  Qed.
-  Global Instance: Idempotent (=) f → Idempotent (=) (merge f).
-  Proof. intros ??. apply finmap_eq. intros. now rewrite !(merge_spec f). Qed.
-
-  Context `{!PropHolds (f None None = None)}.
-
-  Lemma merge_spec_alt m1 m2 m :
-    (∀ i, m !! i = f (m1 !! i) (m2 !! i)) ↔ merge f m1 m2 = m.
-  Proof.
-    split; [| intro; subst; apply (merge_spec _) ].
-    intros Hlookup. apply finmap_eq. intros. rewrite Hlookup.
-    apply (merge_spec _).
-  Qed.
-
-  Lemma merge_comm m1 m2 :
-    (∀ i, f (m1 !! i) (m2 !! i) = f (m2 !! i) (m1 !! i)) →
-    merge f m1 m2 = merge f m2 m1.
-  Proof. intros. apply finmap_eq. intros. now rewrite !(merge_spec f). Qed.
-  Global Instance: Commutative (=) f → Commutative (=) (merge f).
-  Proof. intros ???. apply merge_comm. intros. now apply (commutative f). Qed.
-
-  Lemma merge_assoc m1 m2 m3 :
-    (∀ i, f (m1 !! i) (f (m2 !! i) (m3 !! i)) =
-          f (f (m1 !! i) (m2 !! i)) (m3 !! i)) →
-    merge f m1 (merge f m2 m3) = merge f (merge f m1 m2) m3.
-  Proof. intros. apply finmap_eq. intros. now rewrite !(merge_spec f). Qed.
-  Global Instance: Associative (=) f → Associative (=) (merge f).
-  Proof. intros ????. apply merge_assoc. intros. now apply (associative f). Qed.
-End merge.
-
-(** * Properties of the union and intersection operation *)
-Section union_intersection.
-  Context (f : A → A → A).
-
-  Lemma finmap_union_with_merge m1 m2 i x y :
+Section finmap_common.
+  Context `{FinMap K M} {A : Type}.
+
+  Global Instance finmap_subseteq: SubsetEq (M A) := λ m n,
+    ∀ i x, m !! i = Some x → n !! i = Some x.
+  Global Instance: BoundedPreOrder (M A).
+  Proof. split; [firstorder |]. intros m i x. by rewrite lookup_empty. Qed.
+
+  Lemma lookup_weaken (m1 m2 : M A) i x :
+    m1 !! i = Some x → m1 ⊆ m2 → m2 !! i = Some x.
+  Proof. auto. Qed.
+  Lemma lookup_weaken_is_Some (m1 m2 : M A) i :
+    is_Some (m1 !! i) → m1 ⊆ m2 → is_Some (m2 !! i).
+  Proof. inversion 1. eauto using lookup_weaken. Qed.
+  Lemma lookup_weaken_None (m1 m2 : M A) i :
+    m2 !! i = None → m1 ⊆ m2 → m1 !! i = None.
+  Proof.
+    rewrite eq_None_not_Some. intros Hm2 Hm1m2.
+    specialize (Hm1m2 i). destruct (m1 !! i); naive_solver.
+  Qed.
+
+  Lemma lookup_weaken_inv (m1 m2 : M A) i x y :
     m1 !! i = Some x →
+    m1 ⊆ m2 →
     m2 !! i = Some y →
-    union_with f m1 m2 !! i = Some (f x y).
+    x = y.
   Proof.
-    intros Hx Hy. unfold union_with, finmap_union_with.
-    now rewrite (merge_spec _), Hx, Hy.
+    intros Hm1 ? Hm2. eapply lookup_weaken in Hm1; eauto.
+    congruence.
   Qed.
-  Lemma finmap_union_with_l m1 m2 i x :
-    m1 !! i = Some x → m2 !! i = None → union_with f m1 m2 !! i = Some x.
+
+  Lemma lookup_ne (m : M A) i j : m !! i ≠ m !! j → i ≠ j.
+  Proof. congruence. Qed.
+  Lemma finmap_empty (m : M A) : (∀ i, m !! i = None) → m = ∅.
+  Proof. intros Hm. apply finmap_eq. intros. by rewrite Hm, lookup_empty. Qed.
+  Lemma lookup_empty_is_Some i : ¬is_Some ((∅ : M A) !! i).
+  Proof. rewrite lookup_empty. by inversion 1. Qed.
+  Lemma lookup_empty_Some i (x : A) : ¬∅ !! i = Some x.
+  Proof. by rewrite lookup_empty. Qed.
+
+  Lemma partial_alter_compose (m : M A) i f g :
+    partial_alter (f ∘ g) i m = partial_alter f i (partial_alter g i m).
   Proof.
-    intros Hx Hy. unfold union_with, finmap_union_with.
-    now rewrite (merge_spec _), Hx, Hy.
+    intros. apply finmap_eq. intros ii. case (decide (i = ii)).
+    * intros. subst. by rewrite !lookup_partial_alter.
+    * intros. by rewrite !lookup_partial_alter_ne.
   Qed.
-  Lemma finmap_union_with_r m1 m2 i y :
-    m1 !! i = None → m2 !! i = Some y → union_with f m1 m2 !! i = Some y.
+  Lemma partial_alter_comm (m : M A) i j f g :
+    i ≠ j →
+    partial_alter f i (partial_alter g j m) =
+      partial_alter g j (partial_alter f i m).
   Proof.
-    intros Hx Hy. unfold union_with, finmap_union_with.
-    now rewrite (merge_spec _), Hx, Hy.
+    intros. apply finmap_eq. intros jj.
+    destruct (decide (jj = j)).
+    * subst. by rewrite lookup_partial_alter_ne,
+       !lookup_partial_alter, lookup_partial_alter_ne.
+    * destruct (decide (jj = i)).
+      + subst. by rewrite lookup_partial_alter,
+         !lookup_partial_alter_ne, lookup_partial_alter by congruence.
+      + by rewrite !lookup_partial_alter_ne by congruence.
   Qed.
-  Lemma finmap_union_with_None m1 m2 i :
-    union_with f m1 m2 !! i = None ↔ m1 !! i = None ∧ m2 !! i = None.
+  Lemma partial_alter_self_alt (m : M A) i x :
+    x = m !! i → partial_alter (λ _, x) i m = m.
   Proof.
-    unfold union_with, finmap_union_with. rewrite (merge_spec _).
-    destruct (m1 !! i), (m2 !! i); compute; intuition congruence.
+    intros. apply finmap_eq. intros ii.
+    destruct (decide (i = ii)).
+    * subst. by rewrite lookup_partial_alter.
+    * by rewrite lookup_partial_alter_ne.
+  Qed.
+  Lemma partial_alter_self (m : M A) i : partial_alter (λ _, m !! i) i m = m.
+  Proof. by apply partial_alter_self_alt. Qed.
+
+  Lemma lookup_insert (m : M A) i x : <[i:=x]>m !! i = Some x.
+  Proof. unfold insert. apply lookup_partial_alter. Qed.
+  Lemma lookup_insert_rev (m : M A) i x y : <[i:= x ]>m !! i = Some y → x = y.
+  Proof. rewrite lookup_insert. congruence. Qed.
+  Lemma lookup_insert_ne (m : M A) i j x : i ≠ j → <[i:=x]>m !! j = m !! j.
+  Proof. unfold insert. apply lookup_partial_alter_ne. Qed.
+  Lemma insert_comm (m : M A) i j x y :
+    i ≠ j → <[i:=x]>(<[j:=y]>m) = <[j:=y]>(<[i:=x]>m).
+  Proof. apply partial_alter_comm. Qed.
+
+  Lemma lookup_insert_Some (m : M A) i j x y :
+    <[i:=x]>m !! j = Some y ↔ (i = j ∧ x = y) ∨ (i ≠ j ∧ m !! j = Some y).
+  Proof.
+    split.
+    * destruct (decide (i = j)); subst;
+        rewrite ?lookup_insert, ?lookup_insert_ne; intuition congruence.
+    * intros [[??]|[??]].
+      + subst. apply lookup_insert.
+      + by rewrite lookup_insert_ne.
+  Qed.
+  Lemma lookup_insert_None (m : M A) i j x :
+    <[i:=x]>m !! j = None ↔ m !! j = None ∧ i ≠ j.
+  Proof.
+    split.
+    * destruct (decide (i = j)); subst;
+        rewrite ?lookup_insert, ?lookup_insert_ne; intuition congruence.
+    * intros [??]. by rewrite lookup_insert_ne.
+  Qed.
+
+  Lemma lookup_singleton_Some i j (x y : A) :
+    {[(i, x)]} !! j = Some y ↔ i = j ∧ x = y.
+  Proof.
+    unfold singleton, finmap_singleton.
+    rewrite lookup_insert_Some, lookup_empty. simpl.
+    intuition congruence.
+  Qed.
+  Lemma lookup_singleton_None i j (x : A) :
+    {[(i, x)]} !! j = None ↔ i ≠ j.
+  Proof.
+    unfold singleton, finmap_singleton.
+    rewrite lookup_insert_None, lookup_empty. simpl. tauto.
+  Qed.
+  Lemma insert_singleton i (x y : A) : <[i:=y]>{[(i, x)]} = {[(i, y)]}.
+  Proof.
+    unfold singleton, finmap_singleton, insert, finmap_insert.
+    by rewrite <-partial_alter_compose.
+  Qed.
+
+  Lemma lookup_singleton i (x : A) : {[(i, x)]} !! i = Some x.
+  Proof. by rewrite lookup_singleton_Some. Qed.
+  Lemma lookup_singleton_ne i j (x : A) : i ≠ j → {[(i, x)]} !! j = None.
+  Proof. by rewrite lookup_singleton_None. Qed.
+
+  Lemma lookup_delete (m : M A) i : delete i m !! i = None.
+  Proof. apply lookup_partial_alter. Qed.
+  Lemma lookup_delete_ne (m : M A) i j : i ≠ j → delete i m !! j = m !! j.
+  Proof. apply lookup_partial_alter_ne. Qed.
+
+  Lemma lookup_delete_Some (m : M A) i j y :
+    delete i m !! j = Some y ↔ i ≠ j ∧ m !! j = Some y.
+  Proof.
+    split.
+    * destruct (decide (i = j)); subst;
+        rewrite ?lookup_delete, ?lookup_delete_ne; intuition congruence.
+    * intros [??]. by rewrite lookup_delete_ne.
+  Qed.
+  Lemma lookup_delete_None (m : M A) i j :
+    delete i m !! j = None ↔ i = j ∨ m !! j = None.
+  Proof.
+    destruct (decide (i = j)).
+    * subst. rewrite lookup_delete. tauto.
+    * rewrite lookup_delete_ne; tauto.
   Qed.
 
-  Global Instance: LeftId (=) ∅ (union_with f : M A → M A → M A) := _.
-  Global Instance: RightId (=) ∅ (union_with f : M A → M A → M A) := _.
-  Global Instance:
-    Commutative (=) f → Commutative (=) (union_with f : M A → M A → M A) := _.
-  Global Instance:
-    Associative (=) f → Associative (=) (union_with f : M A → M A → M A) := _.
-  Global Instance:
-    Idempotent (=) f → Idempotent (=) (union_with f : M A → M A → M A) := _.
-End union_intersection.
-
-Lemma finmap_union_Some (m1 m2 : M A) i x :
-  (m1 ∪ m2) !! i = Some x ↔
-    m1 !! i = Some x ∨ (m1 !! i = None ∧ m2 !! i = Some x).
-Proof.
-  unfold union, finmap_union, union_with, finmap_union_with.
-  rewrite (merge_spec _).
-  destruct (m1 !! i), (m2 !! i); compute; try intuition congruence.
-Qed.
-Lemma finmap_union_None (m1 m2 : M A) b :
-  (m1 ∪ m2) !! b = None ↔ m1 !! b = None ∧ m2 !! b = None.
-Proof. apply finmap_union_with_None. Qed.
-End finmap.
+  Lemma delete_empty i : delete i (∅ : M A) = ∅.
+  Proof. rewrite <-(partial_alter_self ∅) at 2. by rewrite lookup_empty. Qed.
+  Lemma delete_singleton i (x : A) : delete i {[(i, x)]} = ∅.
+  Proof. setoid_rewrite <-partial_alter_compose. apply delete_empty. Qed.
+  Lemma delete_comm (m : M A) i j :
+    delete i (delete j m) = delete j (delete i m).
+  Proof. destruct (decide (i = j)). by subst. by apply partial_alter_comm. Qed.
+  Lemma delete_insert_comm (m : M A) i j x :
+    i ≠ j → delete i (<[j:=x]>m) = <[j:=x]>(delete i m).
+  Proof. intro. by apply partial_alter_comm. Qed.
+
+  Lemma delete_notin (m : M A) i :
+    m !! i = None → delete i m = m.
+  Proof.
+    intros. apply finmap_eq. intros j.
+    destruct (decide (i = j)).
+    * subst. by rewrite lookup_delete.
+    * by apply lookup_delete_ne.
+  Qed.
+
+  Lemma delete_partial_alter (m : M A) i f :
+    m !! i = None → delete i (partial_alter f i m) = m.
+  Proof.
+    intros. unfold delete, finmap_delete. rewrite <-partial_alter_compose.
+    rapply partial_alter_self_alt. congruence.
+  Qed.
+  Lemma delete_insert (m : M A) i x :
+    m !! i = None → delete i (<[i:=x]>m) = m.
+  Proof. apply delete_partial_alter. Qed.
+  Lemma insert_delete (m : M A) i x :
+    m !! i = Some x → <[i:=x]>(delete i m) = m.
+  Proof.
+    intros Hmi. unfold delete, finmap_delete, insert, finmap_insert.
+    rewrite <-partial_alter_compose. unfold compose. rewrite <-Hmi.
+    by apply partial_alter_self_alt.
+  Qed.
+
+  Lemma lookup_delete_list (m : M A) is j :
+    j ∈ is → delete_list is m !! j = None.
+  Proof.
+    induction 1 as [|i j is]; simpl.
+    * by rewrite lookup_delete.
+    * destruct (decide (i = j)).
+      + subst. by rewrite lookup_delete.
+      + rewrite lookup_delete_ne; auto.
+  Qed.
+  Lemma lookup_delete_list_not_elem_of (m : M A) is j :
+    j ∉ is → delete_list is m !! j = m !! j.
+  Proof.
+    induction is; simpl; [done |].
+    rewrite elem_of_cons. intros.
+    intros. rewrite lookup_delete_ne; intuition.
+  Qed.
+  Lemma delete_list_notin (m : M A) is :
+    Forall (λ i, m !! i = None) is → delete_list is m = m.
+  Proof.
+    induction 1; simpl; [done |].
+    rewrite delete_notin; congruence.
+  Qed.
+
+  Lemma delete_list_insert_comm (m : M A) is j x :
+    j ∉ is → delete_list is (<[j:=x]>m) = <[j:=x]>(delete_list is m).
+  Proof.
+    induction is; simpl; [done |].
+    rewrite elem_of_cons. intros.
+    rewrite IHis, delete_insert_comm; intuition.
+  Qed.
+
+  Lemma elem_of_dom C `{SimpleCollection K C} (m : M A) i :
+    i ∈ dom C m ↔ is_Some (m !! i).
+  Proof.
+    unfold dom, finmap_dom. simpl. rewrite is_Some_alt.
+    setoid_rewrite <-elem_of_finmap_to_list.
+    induction (finmap_to_list m) as [|[j x] l IH]; simpl.
+    * rewrite elem_of_empty.
+      setoid_rewrite elem_of_nil. naive_solver.
+    * rewrite elem_of_union, elem_of_singleton.
+      setoid_rewrite elem_of_cons. naive_solver.
+  Qed.
+  Lemma not_elem_of_dom C `{SimpleCollection K C} (m : M A) i :
+    i ∉ dom C m ↔ m !! i = None.
+  Proof. by rewrite (elem_of_dom C), eq_None_not_Some. Qed.
+
+  Lemma dom_empty C `{SimpleCollection K C} : dom C (∅ : M A) ≡ ∅.
+  Proof.
+    split; intro.
+    * rewrite (elem_of_dom C), lookup_empty. by inversion 1.
+    * solve_elem_of.
+  Qed.
+  Lemma dom_empty_inv C `{SimpleCollection K C} (m : M A) :
+    dom C m ≡ ∅ → m = ∅.
+  Proof.
+    intros E. apply finmap_empty. intros. apply (not_elem_of_dom C).
+    rewrite E. solve_elem_of.
+  Qed.
+
+  Lemma delete_partial_alter_dom C `{SimpleCollection K C} (m : M A) i f :
+    i ∉ dom C m → delete i (partial_alter f i m) = m.
+  Proof. rewrite (not_elem_of_dom C). apply delete_partial_alter. Qed.
+  Lemma delete_insert_dom C `{SimpleCollection K C} (m : M A) i x :
+    i ∉ dom C m → delete i (<[i:=x]>m) = m.
+  Proof. rewrite (not_elem_of_dom C). apply delete_partial_alter. Qed.
+  Lemma elem_of_dom_delete C `{SimpleCollection K C} (m : M A) i j :
+    i ∈ dom C (delete j m) ↔ i ≠ j ∧ i ∈ dom C m.
+  Proof.
+    rewrite !(elem_of_dom C), <-!not_eq_None_Some.
+    rewrite lookup_delete_None. intuition.
+  Qed.
+  Lemma not_elem_of_dom_delete C `{SimpleCollection K C} (m : M A) i :
+    i ∉ dom C (delete i m).
+  Proof. apply (not_elem_of_dom C), lookup_delete. Qed.
+
+  Lemma subseteq_dom C `{SimpleCollection K C} (m1 m2 : M A) :
+    m1 ⊆ m2 → dom C m1 ⊆ dom C m2.
+  Proof.
+    unfold subseteq, finmap_subseteq, collection_subseteq.
+    intros ??. rewrite !(elem_of_dom C). inversion 1. eauto.
+  Qed.
+  Lemma subset_dom C `{SimpleCollection K C} (m1 m2 : M A) :
+    m1 ⊂ m2 → dom C m1 ⊂ dom C m2.
+  Proof.
+    intros [Hss1 Hss2]. split.
+    * by apply subseteq_dom.
+    * intros Hdom. destruct Hss2. intros i x Hi.
+      specialize (Hdom i). rewrite !(elem_of_dom C) in Hdom.
+      feed inversion Hdom. eauto.
+      by erewrite (Hss1 i) in Hi by eauto.
+  Qed.
+  Lemma finmap_wf : wf (@subset (M A) _).
+  Proof.
+    apply (wf_projected (⊂) (dom (listset K))).
+    * by apply (subset_dom _).
+    * by apply collection_wf.
+  Qed.
+
+  Lemma partial_alter_subseteq (m : M A) i f :
+    m !! i = None →
+    m ⊆ partial_alter f i m.
+  Proof.
+    intros Hi j x Hj. rewrite lookup_partial_alter_ne; congruence.
+  Qed.
+  Lemma partial_alter_subset (m : M A) i f :
+    m !! i = None →
+    is_Some (f (m !! i)) →
+    m ⊂ partial_alter f i m.
+  Proof.
+    intros Hi Hfi. split.
+    * by apply partial_alter_subseteq.
+    * inversion Hfi as [x Hx]. intros Hm.
+      apply (Some_ne_None x). rewrite <-(Hm i x); [done|].
+      by rewrite lookup_partial_alter.
+  Qed.
+  Lemma insert_subseteq (m : M A) i x :
+    m !! i = None →
+    m ⊆ <[i:=x]>m.
+  Proof. apply partial_alter_subseteq. Qed.
+  Lemma insert_subset (m : M A) i x :
+    m !! i = None →
+    m ⊂ <[i:=x]>m.
+  Proof. intro. apply partial_alter_subset; eauto. Qed.
+
+  Lemma delete_subseteq (m : M A) i :
+    delete i m ⊆ m.
+  Proof. intros j x. rewrite lookup_delete_Some. tauto. Qed.
+  Lemma delete_subseteq_compat (m1 m2 : M A) i :
+    m1 ⊆ m2 →
+    delete i m1 ⊆ delete i m2.
+  Proof. intros ? j x. rewrite !lookup_delete_Some. intuition eauto. Qed.
+  Lemma delete_subset_alt (m : M A) i x :
+    m !! i = Some x → delete i m ⊂ m.
+  Proof.
+    split.
+    * apply delete_subseteq.
+    * intros Hi. apply (None_ne_Some x).
+      by rewrite <-(lookup_delete m i), (Hi i x).
+  Qed.
+  Lemma delete_subset (m : M A) i :
+    is_Some (m !! i) → delete i m ⊂ m.
+  Proof. inversion 1. eauto using delete_subset_alt. Qed.
+
+  (** * Induction principles *)
+  (** We use the induction principle on finite collections to prove the
+  following induction principle on finite maps. *)
+  Lemma finmap_ind_alt C (P : M A → Prop) `{FinCollection K C} :
+    P ∅ →
+    (∀ i x m, i ∉ dom C m → P m → P (<[i:=x]>m)) →
+    ∀ m, P m.
+  Proof.
+    intros Hemp Hinsert m.
+    apply (collection_ind (λ X, ∀ m, dom C m ≡ X → P m)) with (dom C m).
+    * solve_proper.
+    * clear m. intros m Hm. rewrite finmap_empty.
+      + done.
+      + intros. rewrite <-(not_elem_of_dom C), Hm.
+        by solve_elem_of.
+    * clear m. intros i X Hi IH m Hdom.
+      assert (∃ x, m !! i = Some x) as [x ?].
+      { apply is_Some_alt, (elem_of_dom C).
+        rewrite Hdom. clear Hdom.
+        by solve_elem_of. }
+      rewrite <-(insert_delete m i x) by done.
+      apply Hinsert.
+      { by apply (not_elem_of_dom_delete C). }
+      apply IH. apply elem_of_equiv. intros.
+      rewrite (elem_of_dom_delete C).
+      esolve_elem_of.
+    * done.
+  Qed.
+
+  (** We use the [listset] implementation to prove an induction principle that
+  does not use the map's domain. *)
+  Lemma finmap_ind (P : M A → Prop) :
+    P ∅ →
+    (∀ i x m, m !! i = None → P m → P (<[i:=x]>m)) →
+    ∀ m, P m.
+  Proof.
+    setoid_rewrite <-(not_elem_of_dom (listset _)).
+    apply (finmap_ind_alt (listset _) P).
+  Qed.
+End finmap_common.
 
 (** * The finite map tactic *)
 (** The tactic [simplify_map by tac] simplifies finite map expressions
-occuring in the conclusion and assumptions. It uses [tac] to discharge generated
+occuring in the conclusion and hypotheses. It uses [tac] to discharge generated
 inequalities. *)
-Tactic Notation "simplify_map" "by" tactic(T) := repeat
+Tactic Notation "simpl_map" "by" tactic3(tac) := repeat
   match goal with
-  | _ => progress simplify_equality
   | H : context[ ∅ !! _ ] |- _ => rewrite lookup_empty in H
   | H : context[ (<[_:=_]>_) !! _ ] |- _ => rewrite lookup_insert in H
-  | H : context[ (<[_:=_]>_) !! _ ] |- _ => rewrite lookup_insert_ne in H by T
-  | H : context[ (delete _ _) !! _ ] |- _ => rewrite lookup_delete in H
-  | H : context[ (delete _ _) !! _ ] |- _ => rewrite lookup_delete_ne in H by T
+  | H : context[ (<[_:=_]>_) !! _ ] |- _ => rewrite lookup_insert_ne in H by tac
+  | H : context[ (delete _ _) !! _] |- _ => rewrite lookup_delete in H
+  | H : context[ (delete _ _) !! _] |- _ => rewrite lookup_delete_ne in H by tac
   | H : context[ {[ _ ]} !! _ ] |- _ => rewrite lookup_singleton in H
-  | H : context[ {[ _ ]} !! _ ] |- _ => rewrite lookup_singleton_ne in H by T
+  | H : context[ {[ _ ]} !! _ ] |- _ => rewrite lookup_singleton_ne in H by tac
   | |- context[ ∅ !! _ ] => rewrite lookup_empty
   | |- context[ (<[_:=_]>_) !! _ ] => rewrite lookup_insert
-  | |- context[ (<[_:=_]>_) !! _ ] => rewrite lookup_insert_ne by T
+  | |- context[ (<[_:=_]>_) !! _ ] => rewrite lookup_insert_ne by tac
   | |- context[ (delete _ _) !! _ ] => rewrite lookup_delete
-  | |- context[ (delete _ _) !! _ ] => rewrite lookup_delete_ne by T
+  | |- context[ (delete _ _) !! _ ] => rewrite lookup_delete_ne by tac
   | |- context[ {[ _ ]} !! _ ] => rewrite lookup_singleton
-  | |- context[ {[ _ ]} !! _ ] => rewrite lookup_singleton_ne by T
+  | |- context[ {[ _ ]} !! _ ] => rewrite lookup_singleton_ne by tac
+  | |- context [ lookup (A:=?A) ?i ?m ] =>
+     let x := fresh in evar (x:A);
+     let x' := eval unfold x in x in clear x;
+     let E := fresh in
+     assert (m !! i = Some x') as E by tac;
+     rewrite E; clear E
+  end.
+
+Create HintDb simpl_map.
+Tactic Notation "simpl_map" := simpl_map by eauto with simpl_map.
+Tactic Notation "simplify_map_equality" "by" tactic3(tac) := repeat
+  match goal with
+  | _ => progress simpl_map by tac
+  | _ => progress simplify_equality
+  | H : {[ _ ]} !! _ = None |- _ => rewrite lookup_singleton_None in H
+  | H : {[ _ ]} !! _ = Some _ |- _ =>
+     rewrite lookup_singleton_Some in H; destruct H
+  | H1 : ?m1 !! ?i = Some ?x, H2 : ?m2 !! ?i = Some ?y |- _ =>
+    let H3 := fresh in
+    feed pose proof (lookup_weaken_inv m1 m2 i x y) as H3;
+      [done | by tac | done | ];
+    clear H2; symmetry in H3
+  | H1 : ?m1 !! ?i = Some ?x, H2 : ?m2 !! ?i = None |- _ =>
+    let H3 := fresh in
+    assert (m1 ⊆ m2) as H3 by tac;
+    apply H3 in H1; congruence
+  end.
+Tactic Notation "simplify_map_equality" :=
+  simplify_map_equality by eauto with simpl_map.
+
+(** * More theorems on finite maps *)
+Section finmap_more.
+  Context `{FinMap K M} {A : Type}.
+
+  (** ** Properties of conversion to lists *)
+  Lemma finmap_to_list_unique (m : M A) i x y :
+    (i,x) ∈ finmap_to_list m →
+    (i,y) ∈ finmap_to_list m →
+    x = y.
+  Proof. rewrite !elem_of_finmap_to_list. congruence. Qed.
+  Lemma finmap_to_list_key_nodup (m : M A) :
+    NoDup (fst <$> finmap_to_list m).
+  Proof.
+    eauto using NoDup_fmap_fst, finmap_to_list_unique, finmap_to_list_nodup.
+  Qed.
+
+  Lemma elem_of_finmap_of_list_1 (l : list (K * A)) i x :
+    NoDup (fst <$> l) → (i,x) ∈ l → finmap_of_list l !! i = Some x.
+  Proof.
+    induction l as [|[j y] l IH]; simpl.
+    * by rewrite elem_of_nil.
+    * rewrite NoDup_cons, elem_of_cons, elem_of_list_fmap.
+      intros [Hl ?] [?|?]; simplify_map_equality; [done |].
+      destruct (decide (i = j)); simplify_map_equality; [|done].
+      destruct Hl. by exists (j,x).
+  Qed.
+  Lemma elem_of_finmap_of_list_2 (l : list (K * A)) i x :
+    finmap_of_list l !! i = Some x → (i,x) ∈ l.
+  Proof.
+    induction l as [|[j y] l IH]; simpl.
+    * by rewrite lookup_empty.
+    * rewrite elem_of_cons. destruct (decide (i = j));
+        simplify_map_equality; intuition congruence.
+  Qed.
+  Lemma elem_of_finmap_of_list (l : list (K * A)) i x :
+    NoDup (fst <$> l) →
+    (i,x) ∈ l ↔ finmap_of_list l !! i = Some x.
+  Proof.
+    split; auto using elem_of_finmap_of_list_1, elem_of_finmap_of_list_2.
+  Qed.
+
+  Lemma not_elem_of_finmap_of_list_1 (l : list (K * A)) i :
+    i ∉ fst <$> l → finmap_of_list l !! i = None.
+  Proof.
+    rewrite elem_of_list_fmap, eq_None_not_Some, is_Some_alt.
+    intros Hi [x ?]. destruct Hi. exists (i,x). simpl.
+    auto using elem_of_finmap_of_list_2.
+  Qed.
+  Lemma not_elem_of_finmap_of_list_2 (l : list (K * A)) i :
+    finmap_of_list l !! i = None → i ∉ fst <$> l.
+  Proof.
+    induction l as [|[j y] l IH]; simpl.
+    * rewrite elem_of_nil. tauto.
+    * rewrite elem_of_cons.
+      destruct (decide (i = j)); simplify_map_equality; by intuition.
+  Qed.
+  Lemma not_elem_of_finmap_of_list (l : list (K * A)) i :
+    i ∉ fst <$> l ↔ finmap_of_list l !! i = None.
+  Proof.
+    split; auto using not_elem_of_finmap_of_list_1,
+      not_elem_of_finmap_of_list_2.
+  Qed.
+
+  Lemma finmap_of_list_proper (l1 l2 : list (K * A)) :
+    NoDup (fst <$> l1) →
+    Permutation l1 l2 →
+    finmap_of_list l1 = finmap_of_list l2.
+  Proof.
+    intros ? Hperm. apply finmap_eq. intros i. apply option_eq. intros x.
+    by rewrite <-!elem_of_finmap_of_list; rewrite <-?Hperm.
+  Qed.
+  Lemma finmap_of_list_inj (l1 l2 : list (K * A)) :
+    NoDup (fst <$> l1) →
+    NoDup (fst <$> l2) →
+    finmap_of_list l1 = finmap_of_list l2 →
+    Permutation l1 l2.
+  Proof.
+    intros ?? Hl1l2.
+    apply NoDup_Permutation; auto using (NoDup_fmap_1 fst).
+    intros [i x]. by rewrite !elem_of_finmap_of_list, Hl1l2.
+  Qed.
+  Lemma finmap_of_to_list (m : M A) :
+    finmap_of_list (finmap_to_list m) = m.
+  Proof.
+    apply finmap_eq. intros i. apply option_eq. intros x.
+    by rewrite <-elem_of_finmap_of_list, elem_of_finmap_to_list
+      by auto using finmap_to_list_key_nodup.
+  Qed.
+  Lemma finmap_to_of_list (l : list (K * A)) :
+    NoDup (fst <$> l) →
+    Permutation (finmap_to_list (finmap_of_list l)) l.
+  Proof.
+    auto using finmap_of_list_inj,
+      finmap_to_list_key_nodup, finmap_of_to_list.
+  Qed.
+  Lemma finmap_to_list_inj (m1 m2 : M A) :
+    Permutation (finmap_to_list m1) (finmap_to_list m2) →
+    m1 = m2.
+  Proof.
+    intros.
+    rewrite <-(finmap_of_to_list m1), <-(finmap_of_to_list m2).
+    auto using finmap_of_list_proper, finmap_to_list_key_nodup.
+  Qed.
+
+  Lemma finmap_to_list_empty :
+    finmap_to_list ∅ = @nil (K * A).
+  Proof.
+    apply elem_of_nil_inv. intros [i x].
+    rewrite elem_of_finmap_to_list. apply lookup_empty_Some.
+  Qed.
+  Lemma finmap_to_list_insert (m : M A) i x :
+    m !! i = None →
+    Permutation (finmap_to_list (<[i:=x]>m)) ((i,x) :: finmap_to_list m).
+  Proof.
+    intros. apply finmap_of_list_inj; simpl.
+    * apply finmap_to_list_key_nodup.
+    * constructor; auto using finmap_to_list_key_nodup.
+      rewrite elem_of_list_fmap.
+      intros [[??] [? Hlookup]]; subst; simpl in *.
+      rewrite elem_of_finmap_to_list in Hlookup. congruence.
+    * by rewrite !finmap_of_to_list.
+  Qed.
+
+  Lemma finmap_of_list_nil :
+    finmap_of_list (@nil (K * A)) = ∅.
+  Proof. done. Qed.
+  Lemma finmap_of_list_cons (l : list (K * A)) i x :
+    finmap_of_list ((i, x) :: l) = <[i:=x]>(finmap_of_list l).
+  Proof. done. Qed.
+
+  Lemma finmap_to_list_empty_inv (m : M A) :
+    Permutation (finmap_to_list m) [] → m = ∅.
+  Proof. rewrite <-finmap_to_list_empty. apply finmap_to_list_inj. Qed.
+  Lemma finmap_to_list_insert_inv (m : M A) l i x :
+    Permutation (finmap_to_list m) ((i,x) :: l) →
+    m = <[i:=x]>(finmap_of_list l).
+  Proof.
+    intros Hperm. apply finmap_to_list_inj.
+    assert (NoDup (fst <$> (i, x) :: l)) as Hnodup.
+    { rewrite <-Hperm. auto using finmap_to_list_key_nodup. }
+    simpl in Hnodup. rewrite NoDup_cons in Hnodup.
+    destruct Hnodup.
+    rewrite Hperm, finmap_to_list_insert, finmap_to_of_list;
+      auto using not_elem_of_finmap_of_list_1.
+  Qed.
+
+  (** ** Properties of the forall predicate *)
+  Section finmap_forall.
+    Context (P : K → A → Prop).
+
+    Lemma finmap_forall_to_list m :
+      finmap_forall P m ↔ Forall (curry P) (finmap_to_list m).
+    Proof.
+      rewrite Forall_forall. split.
+      * intros Hforall [i x].
+        rewrite elem_of_finmap_to_list. by apply (Hforall i x).
+      * intros Hforall i x.
+        rewrite <-elem_of_finmap_to_list. by apply (Hforall (i,x)).
+    Qed.
+
+    Global Instance finmap_forall_dec
+      `{∀ i x, Decision (P i x)} m : Decision (finmap_forall P m).
+    Proof.
+      refine (cast_if (decide (Forall (curry P) (finmap_to_list m))));
+        by rewrite finmap_forall_to_list.
+    Defined.
+  End finmap_forall.
+
+  (** ** Properties of the merge operation *)
+  Section merge_with.
+    Context (f : option A → option A → option A).
+
+    Global Instance: LeftId (=) None f → LeftId (=) ∅ (merge f).
+    Proof.
+      intros ??. apply finmap_eq. intros.
+      by rewrite !(merge_spec f), lookup_empty, (left_id None f).
+    Qed.
+    Global Instance: RightId (=) None f → RightId (=) ∅ (merge f).
+    Proof.
+      intros ??. apply finmap_eq. intros.
+      by rewrite !(merge_spec f), lookup_empty, (right_id None f).
+    Qed.
+
+    Context `{!PropHolds (f None None = None)}.
+
+    Lemma merge_spec_alt m1 m2 m :
+      (∀ i, m !! i = f (m1 !! i) (m2 !! i)) ↔ merge f m1 m2 = m.
+    Proof.
+      split; [| intro; subst; apply (merge_spec _) ].
+      intros Hlookup. apply finmap_eq. intros. rewrite Hlookup.
+      apply (merge_spec _).
+    Qed.
+
+    Lemma merge_commutative m1 m2 :
+      (∀ i, f (m1 !! i) (m2 !! i) = f (m2 !! i) (m1 !! i)) →
+      merge f m1 m2 = merge f m2 m1.
+    Proof. intros. apply finmap_eq. intros. by rewrite !(merge_spec f). Qed.
+    Global Instance: Commutative (=) f → Commutative (=) (merge f).
+    Proof.
+      intros ???. apply merge_commutative. intros. by apply (commutative f).
+    Qed.
+
+    Lemma merge_associative m1 m2 m3 :
+      (∀ i, f (m1 !! i) (f (m2 !! i) (m3 !! i)) =
+            f (f (m1 !! i) (m2 !! i)) (m3 !! i)) →
+      merge f m1 (merge f m2 m3) = merge f (merge f m1 m2) m3.
+    Proof. intros. apply finmap_eq. intros. by rewrite !(merge_spec f). Qed.
+    Global Instance: Associative (=) f → Associative (=) (merge f).
+    Proof.
+      intros ????. apply merge_associative. intros. by apply (associative f).
+    Qed.
+
+    Lemma merge_idempotent m1 :
+      (∀ i, f (m1 !! i) (m1 !! i) = m1 !! i) →
+      merge f m1 m1 = m1.
+    Proof. intros. apply finmap_eq. intros. by rewrite !(merge_spec f). Qed.
+    Global Instance: Idempotent (=) f → Idempotent (=) (merge f).
+    Proof.
+      intros ??. apply merge_idempotent. intros. by apply (idempotent f).
+    Qed.
+  End merge_with.
+
+  (** ** Properties on the intersection forall relation *)
+  Section intersection_forall.
+    Context (R : relation A).
+
+    Global Instance finmap_intersection_forall_sym:
+      Symmetric R → Symmetric (finmap_intersection_forall R).
+    Proof. firstorder auto. Qed.
+    Lemma finmap_intersection_forall_empty_l (m : M A) :
+      finmap_intersection_forall R ∅ m.
+    Proof. intros ???. by simpl_map. Qed.
+    Lemma finmap_intersection_forall_empty_r (m : M A) :
+      finmap_intersection_forall R m ∅.
+    Proof. intros ???. by simpl_map. Qed.
+
+    (** Due to the finiteness of finite maps, we can extract a witness are
+    property does not hold for the intersection. *)
+    Lemma finmap_not_intersection_forall `{∀ x y, Decision (R x y)} (m1 m2 : M A) :
+      ¬finmap_intersection_forall R m1 m2
+        ↔ ∃ i x1 x2, m1 !! i = Some x1 ∧ m2 !! i = Some x2 ∧ ¬R x1 x2.
+    Proof.
+      split.
+      * intros Hdisjoint.
+        set (Pi i := ∀ x1 x2, m1 !! i = Some x1 → m2 !! i = Some x2 → ¬R x1 x2).
+        assert (∀ i, Decision (Pi i)).
+        { intros i. unfold Decision, Pi.
+          destruct (m1 !! i) as [x1|], (m2 !! i) as [x2|]; try (by left).
+          destruct (decide (R x1 x2)).
+          * naive_solver.
+          * intuition congruence. }
+        destruct (decide (cexists Pi (dom (listset _) m1 ∩ dom (listset _) m2)))
+          as [[i [Hdom Hi]] | Hi].
+        + rewrite elem_of_intersection in Hdom.
+          rewrite !(elem_of_dom (listset _)), !is_Some_alt in Hdom.
+          destruct Hdom as [[x1 ?] [x2 ?]]. exists i x1 x2; auto.
+        + destruct Hdisjoint. intros i x1 x2 Hx1 Hx2.
+          apply dec_stable. intros HP.
+          destruct Hi. exists i.
+          rewrite elem_of_intersection, !(elem_of_dom (listset _)).
+          intuition eauto; congruence.
+      * intros (i & x1 & x2 & Hx1 & Hx2 & Hx1x2) Hdisjoint.
+        by apply Hx1x2, (Hdisjoint i x1 x2).
+    Qed.
+  End intersection_forall.
+
+  (** ** Properties on the disjoint maps *)
+  Lemma finmap_disjoint_alt (m1 m2 : M A) :
+    m1 ⊥ m2 ↔ ∀ i, m1 !! i = None ∨ m2 !! i = None.
+  Proof.
+    split; intros Hm1m2 i; specialize (Hm1m2 i);
+      destruct (m1 !! i), (m2 !! i); naive_solver.
+  Qed.    
+  Lemma finmap_not_disjoint (m1 m2 : M A) :
+    ¬m1 ⊥ m2 ↔ ∃ i x1 x2, m1 !! i = Some x1 ∧ m2 !! i = Some x2.
+  Proof.
+    unfold disjoint, finmap_disjoint.
+    rewrite finmap_not_intersection_forall.
+    * naive_solver.
+    * right. auto.
+  Qed.
+
+  Global Instance: Symmetric (@disjoint (M A) _).
+  Proof. apply finmap_intersection_forall_sym. auto. Qed.
+  Lemma finmap_disjoint_empty_l (m : M A) : ∅ ⊥ m.
+  Proof. apply finmap_intersection_forall_empty_l. Qed.
+  Lemma finmap_disjoint_empty_r (m : M A) : m ⊥ ∅.
+  Proof. apply finmap_intersection_forall_empty_r. Qed.
+
+  Lemma finmap_disjoint_weaken (m1 m1' m2 m2' : M A) :
+    m1' ⊥ m2' →
+    m1 ⊆ m1' → m2 ⊆ m2' →
+    m1 ⊥ m2.
+  Proof.
+    intros Hdisjoint Hm1 Hm2 i x1 x2 Hx1 Hx2.
+    destruct (Hdisjoint i x1 x2); auto.
+  Qed.
+  Lemma finmap_disjoint_weaken_l (m1 m1' m2  : M A) :
+    m1' ⊥ m2 → m1 ⊆ m1' → m1 ⊥ m2.
+  Proof. eauto using finmap_disjoint_weaken. Qed.
+  Lemma finmap_disjoint_weaken_r (m1 m2 m2' : M A) :
+    m1 ⊥ m2' → m2 ⊆ m2' → m1 ⊥ m2.
+  Proof. eauto using finmap_disjoint_weaken. Qed.
+
+  Lemma finmap_disjoint_Some_l (m1 m2 : M A) i x:
+    m1 ⊥ m2 →
+    m1 !! i = Some x →
+    m2 !! i = None.
+  Proof.
+    intros Hdisjoint ?. rewrite eq_None_not_Some, is_Some_alt.
+    intros [x2 ?]. by apply (Hdisjoint i x x2).
+  Qed.
+  Lemma finmap_disjoint_Some_r (m1 m2 : M A) i x:
+    m1 ⊥ m2 →
+    m2 !! i = Some x →
+    m1 !! i = None.
+  Proof. rewrite (symmetry_iff (⊥)). apply finmap_disjoint_Some_l. Qed.
+
+  Lemma finmap_disjoint_singleton_l (m : M A) i x :
+    {[(i, x)]} ⊥ m ↔ m !! i = None.
+  Proof.
+    split.
+    * intro. apply (finmap_disjoint_Some_l {[(i, x)]} _ _ x); by simpl_map.
+    * intros ? j y1 y2 ??.
+      destruct (decide (i = j)); simplify_map_equality; congruence.
+  Qed.
+  Lemma finmap_disjoint_singleton_r (m : M A) i x :
+    m ⊥ {[(i, x)]} ↔ m !! i = None.
+  Proof. by rewrite (symmetry_iff (⊥)), finmap_disjoint_singleton_l. Qed.
+
+  Lemma finmap_disjoint_singleton_l_2 (m : M A) i x :
+    m !! i = None → {[(i, x)]} ⊥ m.
+  Proof. by rewrite finmap_disjoint_singleton_l. Qed.
+  Lemma finmap_disjoint_singleton_r_2 (m : M A) i x :
+    m !! i = None → m ⊥ {[(i, x)]}.
+  Proof. by rewrite finmap_disjoint_singleton_r. Qed.
+
+  (** ** Properties of the union and intersection operation *)
+  Section union_intersection_with.
+    Context (f : A → A → option A).
+
+    Lemma finmap_union_with_Some m1 m2 i x y :
+      m1 !! i = Some x →
+      m2 !! i = Some y →
+      union_with f m1 m2 !! i = f x y.
+    Proof.
+      intros Hx Hy. unfold union_with, finmap_union_with.
+      by rewrite (merge_spec _), Hx, Hy.
+    Qed.
+    Lemma finmap_union_with_Some_l m1 m2 i x :
+      m1 !! i = Some x →
+      m2 !! i = None →
+      union_with f m1 m2 !! i = Some x.
+    Proof.
+      intros Hx Hy. unfold union_with, finmap_union_with.
+      by rewrite (merge_spec _), Hx, Hy.
+    Qed.
+    Lemma finmap_union_with_Some_r m1 m2 i y :
+      m1 !! i = None →
+      m2 !! i = Some y →
+      union_with f m1 m2 !! i = Some y.
+    Proof.
+      intros Hx Hy. unfold union_with, finmap_union_with.
+      by rewrite (merge_spec _), Hx, Hy.
+    Qed.
+
+    Global Instance: LeftId (=) ∅ (@union_with _ (M A) _ f).
+    Proof. unfold union_with, finmap_union_with. apply _. Qed.
+    Global Instance: RightId (=) ∅ (@union_with _ (M A) _ f).
+    Proof. unfold union_with, finmap_union_with. apply _. Qed.
+    Global Instance:
+      Commutative (=) f → Commutative (=) (@union_with _ (M A) _ f).
+    Proof. unfold union_with, finmap_union_with. apply _. Qed.
+  End union_intersection_with.
+
+  Global Instance: LeftId (=) ∅ (@union (M A) _) := _.
+  Global Instance: RightId (=) ∅ (@union (M A) _) := _.
+  Global Instance: Associative (=) (@union (M A) _).
+  Proof.
+    intros m1 m2 m3. unfold union, finmap_union, union_with, finmap_union_with.
+    apply (merge_associative _). intros i.
+    by destruct (m1 !! i), (m2 !! i), (m3 !! i).
+  Qed.
+  Global Instance: Idempotent (=) (@union (M A) _).
+    intros m. unfold union, finmap_union, union_with, finmap_union_with.
+    apply (merge_idempotent _). intros i. by destruct (m !! i).
+  Qed.
+
+  Lemma lookup_union_Some_raw (m1 m2 : M A) i x :
+    (m1 ∪ m2) !! i = Some x ↔
+      m1 !! i = Some x ∨ (m1 !! i = None ∧ m2 !! i = Some x).
+  Proof.
+    unfold union, finmap_union, union_with, finmap_union_with.
+    rewrite (merge_spec _).
+    destruct (m1 !! i), (m2 !! i); compute; intuition congruence.
+  Qed.
+  Lemma lookup_union_None (m1 m2 : M A) i :
+    (m1 ∪ m2) !! i = None ↔ m1 !! i = None ∧ m2 !! i = None.
+  Proof.
+    unfold union, finmap_union, union_with, finmap_union_with.
+    rewrite (merge_spec _).
+    destruct (m1 !! i), (m2 !! i); compute; intuition congruence.
+  Qed.
+
+  Lemma lookup_union_Some (m1 m2 : M A) i x :
+    m1 ⊥ m2 →
+    (m1 ∪ m2) !! i = Some x ↔ m1 !! i = Some x ∨ m2 !! i = Some x.
+  Proof.
+    intros Hdisjoint. rewrite lookup_union_Some_raw.
+    intuition eauto using finmap_disjoint_Some_r.
+  Qed.
+
+  Lemma lookup_union_Some_l (m1 m2 : M A) i x :
+    m1 !! i = Some x →
+    (m1 ∪ m2) !! i = Some x.
+  Proof. intro. rewrite lookup_union_Some_raw; intuition. Qed.
+  Lemma lookup_union_Some_r (m1 m2 : M A) i x :
+    m1 ⊥ m2 →
+    m2 !! i = Some x →
+    (m1 ∪ m2) !! i = Some x.
+  Proof. intro. rewrite lookup_union_Some; intuition. Qed.
+
+  Lemma finmap_union_comm (m1 m2 : M A) :
+    m1 ⊥ m2 →
+    m1 ∪ m2 = m2 ∪ m1.
+  Proof.
+    intros Hdisjoint. apply (merge_commutative (union_with (λ x _, Some x))).
+    intros i. specialize (Hdisjoint i).
+    destruct (m1 !! i), (m2 !! i); compute; naive_solver.
+  Qed.
+
+  Lemma finmap_subseteq_union (m1 m2 : M A) :
+    m1 ⊆ m2 →
+    m1 ∪ m2 = m2.
+  Proof.
+    intros Hm1m2.
+    apply finmap_eq. intros i. apply option_eq. intros x.
+    rewrite lookup_union_Some_raw. split; [by intuition |].
+    intros Hm2. specialize (Hm1m2 i).
+    destruct (m1 !! i) as [y|]; [| by auto].
+    rewrite (Hm1m2 y eq_refl) in Hm2. intuition congruence.
+  Qed.
+  Lemma finmap_subseteq_union_l (m1 m2 : M A) :
+    m1 ⊆ m1 ∪ m2.
+  Proof. intros ? i x. rewrite lookup_union_Some_raw. intuition. Qed.
+  Lemma finmap_subseteq_union_r (m1 m2 : M A) :
+    m1 ⊥ m2 →
+    m2 ⊆ m1 ∪ m2.
+  Proof.
+    intros. rewrite finmap_union_comm by done.
+    by apply finmap_subseteq_union_l.
+  Qed.
+
+  Lemma finmap_disjoint_union_l (m1 m2 m3 : M A) :
+    m1 ∪ m2 ⊥ m3 ↔ m1 ⊥ m3 ∧ m2 ⊥ m3.
+  Proof.
+    rewrite !finmap_disjoint_alt.
+    setoid_rewrite lookup_union_None. naive_solver.
+  Qed.
+  Lemma finmap_disjoint_union_r (m1 m2 m3 : M A) :
+    m1 ⊥ m2 ∪ m3 ↔ m1 ⊥ m2 ∧ m1 ⊥ m3.
+  Proof.
+    rewrite !finmap_disjoint_alt.
+    setoid_rewrite lookup_union_None. naive_solver.
+  Qed.
+  Lemma finmap_disjoint_union_l_2 (m1 m2 m3 : M A) :
+    m1 ⊥ m3 → m2 ⊥ m3 → m1 ∪ m2 ⊥ m3.
+  Proof. by rewrite finmap_disjoint_union_l. Qed.
+  Lemma finmap_disjoint_union_r_2 (m1 m2 m3 : M A) :
+    m1 ⊥ m2 → m1 ⊥ m3 → m1 ⊥ m2 ∪ m3.
+  Proof. by rewrite finmap_disjoint_union_r. Qed.
+  Lemma finmap_union_cancel_l (m1 m2 m3 : M A) :
+    m1 ⊥ m3 →
+    m2 ⊥ m3 →
+    m1 ∪ m3 = m2 ∪ m3 →
+    m1 = m2.
+  Proof.
+    revert m1 m2 m3.
+    cut (∀ (m1 m2 m3 : M A) i x,
+      m1 ⊥ m3 →
+      m2 ⊥ m3 →
+      m1 ∪ m3 = m2 ∪ m3 →
+      m1 !! i = Some x → m2 !! i = Some x).
+    { intros. apply finmap_eq. intros i.
+      apply option_eq. naive_solver. }
+    intros m1 m2 m3 b v Hm1m3 Hm2m3 E ?.
+    destruct (proj1 (lookup_union_Some m2 m3 b v Hm2m3)) as [E2|E2].
+    * rewrite <-E. by apply lookup_union_Some_l.
+    * done.
+    * contradict E2. by apply eq_None_ne_Some, finmap_disjoint_Some_l with m1 v.
+  Qed.
+  Lemma finmap_union_cancel_r (m1 m2 m3 : M A) :
+    m1 ⊥ m3 →
+    m2 ⊥ m3 →
+    m3 ∪ m1 = m3 ∪ m2 →
+    m1 = m2.
+  Proof.
+    intros ??. rewrite !(finmap_union_comm m3) by done.
+    by apply finmap_union_cancel_l.
+  Qed.
+
+  Lemma insert_union_singleton_l (m : M A) i x :
+    <[i:=x]>m = {[(i,x)]} ∪ m.
+  Proof.
+    apply finmap_eq. intros j. apply option_eq. intros y.
+    rewrite lookup_union_Some_raw.
+    destruct (decide (i = j)); simplify_map_equality; intuition congruence.
+  Qed.
+  Lemma insert_union_singleton_r (m : M A) i x :
+    m !! i = None →
+    <[i:=x]>m = m ∪ {[(i,x)]}.
+  Proof.
+    intro. rewrite insert_union_singleton_l, finmap_union_comm; [done |].
+    by apply finmap_disjoint_singleton_l.
+  Qed.
+
+  Lemma finmap_disjoint_insert_l (m1 m2 : M A) i x :
+    <[i:=x]>m1 ⊥ m2 ↔ m2 !! i = None ∧ m1 ⊥ m2.
+  Proof.
+    rewrite insert_union_singleton_l.
+    by rewrite finmap_disjoint_union_l, finmap_disjoint_singleton_l.
+  Qed.
+  Lemma finmap_disjoint_insert_r (m1 m2 : M A) i x :
+    m1 ⊥ <[i:=x]>m2 ↔ m1 !! i = None ∧ m1 ⊥ m2.
+  Proof.
+    rewrite insert_union_singleton_l.
+    by rewrite finmap_disjoint_union_r, finmap_disjoint_singleton_r.
+  Qed.
+
+  Lemma finmap_disjoint_insert_l_2 (m1 m2 : M A) i x :
+    m2 !! i = None → m1 ⊥ m2 → <[i:=x]>m1 ⊥ m2.
+  Proof. by rewrite finmap_disjoint_insert_l. Qed.
+  Lemma finmap_disjoint_insert_r_2 (m1 m2 : M A) i x :
+    m1 !! i = None → m1 ⊥ m2 → m1 ⊥ <[i:=x]>m2.
+  Proof. by rewrite finmap_disjoint_insert_r. Qed.
+
+  Lemma insert_union_l (m1 m2 : M A) i x :
+    <[i:=x]>(m1 ∪ m2) = <[i:=x]>m1 ∪ m2.
+  Proof. by rewrite !insert_union_singleton_l, (associative (∪)). Qed.
+  Lemma insert_union_r (m1 m2 : M A) i x :
+    m1 !! i = None →
+    <[i:=x]>(m1 ∪ m2) = m1 ∪ <[i:=x]>m2.
+  Proof.
+    intro. rewrite !insert_union_singleton_l, !(associative (∪)).
+    rewrite (finmap_union_comm m1); [done |].
+    by apply finmap_disjoint_singleton_r.
+  Qed.
+
+  Lemma insert_list_union l (m : M A) :
+    insert_list l m = finmap_of_list l ∪ m.
+  Proof.
+    induction l; simpl.
+    * by rewrite (left_id _ _).
+    * by rewrite IHl, insert_union_l.
+  Qed.
+
+  Lemma insert_subseteq_r (m1 m2 : M A) i x :
+    m1 !! i = None → m1 ⊆ m2 → m1 ⊆ <[i:=x]>m2.
+  Proof.
+    intros ?? j. by destruct (decide (j = i)); intros; simplify_map_equality.
+  Qed.
+
+  (** ** Properties of the delete operation *)
+  Lemma finmap_disjoint_delete_l (m1 m2 : M A) i :
+    m1 ⊥ m2 → delete i m1 ⊥ m2.
+  Proof.
+    rewrite !finmap_disjoint_alt.
+    intros Hdisjoint j. destruct (Hdisjoint j); auto.
+    rewrite lookup_delete_None. tauto.
+  Qed.
+  Lemma finmap_disjoint_delete_r (m1 m2 : M A) i :
+    m1 ⊥ m2 → m1 ⊥ delete i m2.
+  Proof. symmetry. by apply finmap_disjoint_delete_l. Qed.
+
+  Lemma finmap_disjoint_delete_list_l (m1 m2 : M A) is :
+    m1 ⊥ m2 → delete_list is m1 ⊥ m2.
+  Proof. induction is; simpl; auto using finmap_disjoint_delete_l. Qed.
+  Lemma finmap_disjoint_delete_list_r (m1 m2 : M A) is :
+    m1 ⊥ m2 → m1 ⊥ delete_list is m2.
+  Proof. induction is; simpl; auto using finmap_disjoint_delete_r. Qed.
+
+  Lemma finmap_union_delete (m1 m2 : M A) i :
+    delete i (m1 ∪ m2) = delete i m1 ∪ delete i m2.
+  Proof.
+    intros. apply finmap_eq. intros j. apply option_eq. intros y.
+    destruct (decide (i = j)); simplify_map_equality;
+     rewrite ?lookup_union_Some_raw; simpl_map; intuition congruence.
+  Qed.
+  Lemma finmap_union_delete_list (m1 m2 : M A) is :
+    delete_list is (m1 ∪ m2) = delete_list is m1 ∪ delete_list is m2.
+  Proof.
+    induction is; simpl; [done |].
+    by rewrite IHis, finmap_union_delete.
+  Qed.
+
+  Lemma finmap_disjoint_union_list_l (ms : list (M A)) (m : M A) :
+    ⋃ ms ⊥ m ↔ Forall (⊥ m) ms.
+  Proof.
+    split.
+    * induction ms; simpl; rewrite ?finmap_disjoint_union_l; intuition.
+    * induction 1; simpl.
+      + apply finmap_disjoint_empty_l.
+      + by rewrite finmap_disjoint_union_l.
+  Qed.
+  Lemma finmap_disjoint_union_list_r (ms : list (M A)) (m : M A) :
+    m ⊥ ⋃ ms ↔ Forall (⊥ m) ms.
+  Proof. by rewrite (symmetry_iff (⊥)), finmap_disjoint_union_list_l. Qed.
+
+  Lemma finmap_disjoint_union_list_l_2 (ms : list (M A)) (m : M A) :
+    Forall (⊥ m) ms → ⋃ ms ⊥ m.
+  Proof. by rewrite finmap_disjoint_union_list_l. Qed.
+  Lemma finmap_disjoint_union_list_r_2 (ms : list (M A)) (m : M A) :
+    Forall (⊥ m) ms → m ⊥ ⋃ ms.
+  Proof. by rewrite finmap_disjoint_union_list_r. Qed.
+
+  (** ** Properties of the conversion from lists to maps *)
+  Lemma finmap_disjoint_of_list_l (m : M A) ixs :
+    finmap_of_list ixs ⊥ m ↔ Forall (λ ix, m !! fst ix = None) ixs.
+  Proof.
+    split.
+    * induction ixs; simpl; rewrite ?finmap_disjoint_insert_l in *; intuition.
+    * induction 1; simpl.
+      + apply finmap_disjoint_empty_l.
+      + rewrite finmap_disjoint_insert_l. auto.
+  Qed.
+  Lemma finmap_disjoint_of_list_r (m : M A) ixs :
+    m ⊥ finmap_of_list ixs ↔ Forall (λ ix, m !! fst ix = None) ixs.
+  Proof. by rewrite (symmetry_iff (⊥)), finmap_disjoint_of_list_l. Qed.
+
+  Lemma finmap_disjoint_of_list_zip_l (m : M A) is xs :
+    same_length is xs →
+    finmap_of_list (zip is xs) ⊥ m ↔ Forall (λ i, m !! i = None) is.
+  Proof.
+    intro. rewrite finmap_disjoint_of_list_l.
+    rewrite <-(zip_fst is xs) at 2 by done.
+    by rewrite Forall_fmap.
+  Qed.
+  Lemma finmap_disjoint_of_list_zip_r (m : M A) is xs :
+    same_length is xs →
+    m ⊥ finmap_of_list (zip is xs) ↔ Forall (λ i, m !! i = None) is.
+  Proof.
+    intro. by rewrite (symmetry_iff (⊥)), finmap_disjoint_of_list_zip_l.
+  Qed.
+  Lemma finmap_disjoint_of_list_zip_l_2 (m : M A) is xs :
+    same_length is xs →
+    Forall (λ i, m !! i = None) is →
+    finmap_of_list (zip is xs) ⊥ m.
+  Proof. intro. by rewrite finmap_disjoint_of_list_zip_l. Qed.
+  Lemma finmap_disjoint_of_list_zip_r_2 (m : M A) is xs :
+    same_length is xs →
+    Forall (λ i, m !! i = None) is →
+    m ⊥ finmap_of_list (zip is xs).
+  Proof. intro. by rewrite finmap_disjoint_of_list_zip_r. Qed.
+
+  (** ** Properties with respect to vectors *)
+  Lemma union_delete_vec {n} (ms : vec (M A) n) (i : fin n) :
+    list_disjoint ms →
+    ms !!! i ∪ ⋃ delete (fin_to_nat i) (vec_to_list ms) = ⋃ ms.
+  Proof.
+    induction ms as [|m ? ms]; inversion_clear 1;
+      inv_fin i; simpl; [done | intros i].
+    rewrite (finmap_union_comm m), (associative_eq _ _), IHms.
+    * by apply finmap_union_comm, finmap_disjoint_union_list_l.
+    * done.
+    * by apply finmap_disjoint_union_list_r, Forall_delete.
+  Qed.
+
+  Lemma union_insert_vec {n} (ms : vec (M A) n) (i : fin n) m :
+    m ⊥ ⋃ delete (fin_to_nat i) (vec_to_list ms) →
+    ⋃ vinsert i m ms = m ∪ ⋃ delete (fin_to_nat i) (vec_to_list ms).
+  Proof.
+    induction ms as [|m' ? ms IH];
+      inv_fin i; simpl; [done | intros i Hdisjoint].
+    rewrite finmap_disjoint_union_r in Hdisjoint.
+    rewrite IH, !(associative_eq (∪)), (finmap_union_comm m); intuition.
+  Qed.
+
+  Lemma finmap_list_disjoint_delete_vec {n} (ms : vec (M A) n) (i : fin n) :
+    list_disjoint ms →
+    Forall (⊥ ms !!! i) (delete (fin_to_nat i) (vec_to_list ms)).
+  Proof.
+    induction ms; inversion_clear 1; inv_fin i; simpl.
+    * done.
+    * constructor. symmetry. by apply Forall_vlookup. auto.
+  Qed.
+
+  Lemma finmap_list_disjoint_insert_vec {n} (ms : vec (M A) n) (i : fin n) m :
+    list_disjoint ms →
+    Forall (⊥ m) (delete (fin_to_nat i) (vec_to_list ms)) →
+    list_disjoint (vinsert i m ms).
+  Proof.
+    induction ms as [|m' ? ms]; inversion_clear 1; inv_fin i; simpl.
+    { constructor; auto. }
+    intros i. inversion_clear 1. constructor; [| by auto].
+    apply Forall_vlookup_2. intros j.
+    destruct (decide (i = j)); subst;
+      rewrite ?vlookup_insert, ?vlookup_insert_ne by done.
+    * done.
+    * by apply Forall_vlookup.
+  Qed.
+
+  (** ** Properties of the difference operation *)
+  Lemma finmap_difference_Some (m1 m2 : M A) i x :
+    (m1 ∖ m2) !! i = Some x ↔ m1 !! i = Some x ∧ m2 !! i = None.
+  Proof.
+    unfold difference, finmap_difference,
+      difference_with, finmap_difference_with.
+    rewrite (merge_spec _).
+    destruct (m1 !! i), (m2 !! i); compute; intuition congruence.
+  Qed.
+
+  Lemma finmap_disjoint_difference_l (m1 m2 m3 : M A) :
+    m2 ⊆ m3 → m1 ∖ m3 ⊥ m2.
+  Proof.
+    intros E i. specialize (E i).
+    unfold difference, finmap_difference,
+      difference_with, finmap_difference_with.
+    rewrite (merge_spec _).
+    destruct (m1 !! i), (m2 !! i), (m3 !! i); compute;
+      try intuition congruence.
+    ediscriminate E; eauto.
+  Qed.
+  Lemma finmap_disjoint_difference_r (m1 m2 m3 : M A) :
+    m2 ⊆ m3 → m2 ⊥ m1 ∖ m3.
+  Proof. intros. symmetry. by apply finmap_disjoint_difference_l. Qed.
+
+  Lemma finmap_union_difference (m1 m2 : M A) :
+    m1 ⊆ m2 → m2 = m1 ∪ m2 ∖ m1.
+  Proof.
+    intro Hm1m2. apply finmap_eq. intros i.
+    apply option_eq. intros v. specialize (Hm1m2 i).
+    unfold difference, finmap_difference,
+      difference_with, finmap_difference_with.
+    rewrite lookup_union_Some_raw, (merge_spec _).
+    destruct (m1 !! i) as [v'|], (m2 !! i);
+      try specialize (Hm1m2 v'); compute; intuition congruence.
+  Qed.
+End finmap_more.
+
+Hint Extern 80 ((_ ∪ _) !! _ = Some _) =>
+  apply lookup_union_Some_l : simpl_map.
+Hint Extern 81 ((_ ∪ _) !! _ = Some _) =>
+  apply lookup_union_Some_r : simpl_map.
+Hint Extern 80 ({[ _ ]} !! _ = Some _) =>
+  apply lookup_singleton : simpl_map.
+Hint Extern 80 (<[_:=_]> _ !! _ = Some _) =>
+  apply lookup_insert : simpl_map.
+
+(** * Tactic to decompose disjoint assumptions *)
+(** The tactic [decompose_map_disjoint] simplifies occurences of [disjoint]
+in the conclusion and hypotheses that involve the union, insert, or singleton
+operation. *)
+Ltac decompose_map_disjoint := repeat
+  match goal with
+  | H : _ ∪ _ ⊥ _ |- _ =>
+    apply finmap_disjoint_union_l in H; destruct H
+  | H : _ ⊥ _ ∪ _ |- _ =>
+    apply finmap_disjoint_union_r in H; destruct H
+  | H : {[ _ ]} ⊥ _ |- _ => apply finmap_disjoint_singleton_l in H
+  | H : _ ⊥ {[ _ ]} |- _ =>  apply finmap_disjoint_singleton_r in H
+  | H : <[_:=_]>_ ⊥ _ |- _ =>
+    apply finmap_disjoint_insert_l in H; destruct H
+  | H : _ ⊥ <[_:=_]>_ |- _ =>
+    apply finmap_disjoint_insert_r in H; destruct H
+  | H : ⋃ _ ⊥ _ |- _ => apply finmap_disjoint_union_list_l in H
+  | H : _ ⊥ ⋃ _ |- _ => apply finmap_disjoint_union_list_r in H
+  | H : ∅ ⊥_ |- _ => clear H
+  | H : _ ⊥ ∅ |- _ => clear H
+  | H : list_disjoint [] |- _ => clear H
+  | H : list_disjoint [_] |- _ => clear H
+  | H : list_disjoint (_ :: _) |- _ =>
+    apply list_disjoint_cons_inv in H; destruct H
+  | H : Forall (⊥ _) _ |- _ => rewrite Forall_vlookup in H
+  | H : Forall (⊥ _) [] |- _ => clear H
+  | H : Forall (⊥ _) (_ :: _) |- _ =>
+    rewrite Forall_cons in H; destruct H
+  | H : Forall (⊥ _) (_ :: _) |- _ =>
+    rewrite Forall_app in H; destruct H
   end.
-Tactic Notation "simplify_map" := simplify_map by auto.
diff --git a/theories/fresh_numbers.v b/theories/fresh_numbers.v
new file mode 100644
index 0000000000000000000000000000000000000000..0be3b5577a06346374b738643ae9e64bba1a7070
--- /dev/null
+++ b/theories/fresh_numbers.v
@@ -0,0 +1,36 @@
+(* Copyright (c) 2012, Robbert Krebbers. *)
+(* This file is distributed under the terms of the BSD license. *)
+(** Given a finite set of binary naturals [N], we generate a fresh element by
+taking the maximum, and adding one to it. We declare this operation as an
+instance of the type class [Fresh]. *)
+Require Export numbers fin_collections.
+
+Definition Nmax `{Elements N C} : C → N := collection_fold Nmax 0%N.
+
+Instance Nmax_proper `{FinCollection N C} : Proper ((≡) ==> (=)) Nmax.
+Proof.
+  apply (collection_fold_proper (=)).
+  * solve_proper.
+  * intros. rewrite !N.max_assoc. f_equal. apply N.max_comm.
+Qed.
+
+Lemma Nmax_max `{FinCollection N C} X x : x ∈ X → (x ≤ Nmax X)%N.
+Proof.
+  apply (collection_fold_ind (λ b X, x ∈ X → (x ≤ b)%N)).
+  * solve_proper.
+  * solve_elem_of.
+  * solve_elem_of (eauto using N.le_max_l, N.le_max_r, N.le_trans).
+Qed.
+
+Instance Nfresh `{FinCollection N C} : Fresh N C := λ l, (1 + Nmax l)%N.
+Instance Nfresh_spec `{FinCollection N C} : FreshSpec N C.
+Proof.
+  split.
+  * apply _.
+  * intros. unfold fresh, Nfresh.
+    setoid_replace X with Y; [done |].
+    by apply elem_of_equiv.
+  * intros X E. assert (1 ≤ 0)%N as []; [| done].
+    apply N.add_le_mono_r with (Nmax X).
+    by apply Nmax_max.
+Qed.
diff --git a/theories/list.v b/theories/list.v
index 34dc58d6e5c19389775e0cb64bc56671ca5602b1..0a5141079df66eb52f90ab7000efe52e0479a763 100644
--- a/theories/list.v
+++ b/theories/list.v
@@ -2,15 +2,25 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects general purpose definitions and theorems on lists that
 are not in the Coq standard library. *)
+
 Require Import Permutation.
-Require Export base decidable option.
+Require Export base decidable option numbers.
 
+Arguments length {_} _.
 Arguments cons {_} _ _.
 Arguments app {_} _ _.
-
-Arguments In {_} _ _.
-Arguments NoDup_nil {_}.
 Arguments Permutation {_} _ _.
+Arguments Forall_cons {_} _ _ _ _ _.
+
+Notation Forall_nil_2 := Forall_nil.
+Notation Forall_cons_2 := Forall_cons.
+
+Notation tail := tl.
+Notation take := firstn.
+Notation drop := skipn.
+Notation take_drop := firstn_skipn.
+Arguments take {_} !_ !_ /.
+Arguments drop {_} !_ !_ /.
 
 Notation "(::)" := cons (only parsing) : C_scope.
 Notation "( x ::)" := (cons x) (only parsing) : C_scope.
@@ -20,350 +30,1322 @@ Notation "( l ++)" := (app l) (only parsing) : C_scope.
 Notation "(++ k )" := (λ l, app l k) (only parsing) : C_scope.
 
 (** * General definitions *)
-(** Looking up elements and updating elements in a list. *)
-Global Instance list_lookup: Lookup nat list :=
-  fix list_lookup A (i : nat) (l : list A) {struct l} : option A :=
+(** Looking up, updating, and deleting elements of a list. *)
+Instance list_lookup {A} : Lookup nat A (list A) :=
+  fix go (i : nat) (l : list A) {struct l} : option A :=
   match l with
   | [] => None
   | x :: l =>
     match i with
     | 0 => Some x
-    | S i => @lookup _ _ list_lookup _ i l
+    | S i => @lookup _ _ _ go i l
     end
   end.
-
-Global Instance list_alter: Alter nat list :=
-  fix list_alter A (f : A → A) (i : nat) (l : list A) {struct l} :=
+Instance list_alter {A} (f : A → A) : AlterD nat A (list A) f :=
+  fix go (i : nat) (l : list A) {struct l} :=
   match l with
   | [] => []
   | x :: l =>
     match i with
     | 0 => f x :: l
-    | S i => x :: @alter _ _ list_alter _ f i l
+    | S i => x :: @alter _ _ _ f go i l
     end
   end.
-
-(** The [simpl] tactic does not simplify [list_lookup] as it is wrapped into
-an operational type class and we cannot let it unfold on a per instance basis.
-Therefore we use the [simplify_list_lookup] tactic to perform these
-simplifications. Bug: it does not unfold under binders. *)
-Ltac simplify_list_lookup := repeat
-  match goal with
-  | |- context C [@nil ?A !! _] =>
-    let X := (context C [@None A]) in change X
-  | |- context C [(?x :: _) !! 0] =>
-    let X := (context C [Some x]) in change X
-  | |- context C [(_ :: ?l) !! S ?i] =>
-    let X := (context C [l !! i]) in change X
-  | H : context C [@nil ?A !! _] |- _ =>
-    let X := (context C [@None A]) in change X in H
-  | H : context C [(?x :: _) !! 0] |- _ =>
-    let X := (context C [Some x]) in change X in H
-  | H : context C [(_ :: ?l) !! S ?i] |- _  =>
-    let X := (context C [l !! i]) in change X in H
+Instance list_delete {A} : Delete nat (list A) :=
+  fix go (i : nat) (l : list A) {struct l} : list A :=
+  match l with
+  | [] => []
+  | x :: l =>
+    match i with
+    | 0 => l
+    | S i => x :: @delete _ _ go i l
+    end
   end.
+Instance list_insert {A} : Insert nat A (list A) := λ i x,
+  alter (λ _, x) i.
 
 (** The function [option_list] converts an element of the option type into
 a list. *)
 Definition option_list {A} : option A → list A := option_rect _ (λ x, [x]) [].
 
+(** The function [filter P l] returns the list of elements of [l] that
+satisfies [P]. The order remains unchanged. *)
+Instance list_filter {A} : Filter A (list A) :=
+  fix go P _ l :=
+  match l with
+  | [] => []
+  | x :: l =>
+     if decide (P x)
+     then x :: @filter _ _ (@go) _ _ l
+     else @filter _ _ (@go) _ _ l
+  end.
+
+(** The function [replicate n x] generates a list with length [n] of elements
+with value [x]. *)
+Fixpoint replicate {A} (n : nat) (x : A) : list A :=
+  match n with
+  | 0 => []
+  | S n => x :: replicate n x
+  end.
+
+(** The function [reverse l] returns the elements of [l] in reverse order. *)
+Definition reverse {A} (l : list A) : list A := rev_append l [].
+
+(** The function [resize n y l] takes the first [n] elements of [l] in case
+[length l ≤ n], and otherwise appends elements with value [x] to [l] to obtain
+a list of length [n]. *)
+Fixpoint resize {A} (n : nat) (y : A) (l : list A) : list A :=
+  match l with
+  | [] => replicate n y
+  | x :: l =>
+    match n with
+    | 0 => []
+    | S n => x :: resize n y l
+    end
+  end.
+Arguments resize {_} !_ _ !_.
+
 (** The predicate [prefix_of] holds if the first list is a prefix of the second.
 The predicate [suffix_of] holds if the first list is a suffix of the second. *)
 Definition prefix_of {A} (l1 l2 : list A) : Prop := ∃ k, l2 = l1 ++ k.
 Definition suffix_of {A} (l1 l2 : list A) : Prop := ∃ k, l2 = k ++ l1.
 
+(** * Tactics on lists *)
+Lemma cons_inv {A} (l1 l2 : list A) x1 x2 :
+  x1 :: l1 = x2 :: l2 → x1 = x2 ∧ l1 = l2.
+Proof. by injection 1. Qed.
+
+(** The tactic [discriminate_list_equality] discharges goals containing
+invalid list equalities as an assumption. *)
+Tactic Notation "discriminate_list_equality" hyp(H) :=
+  apply (f_equal length) in H;
+  repeat (simpl in H || rewrite app_length in H);
+  exfalso; lia.
+Tactic Notation "discriminate_list_equality" :=
+  repeat_on_hyps (fun H => discriminate_list_equality H).
+
+(** The tactic [simplify_list_equality] simplifies assumptions involving
+equalities on lists. *)
+Ltac simplify_list_equality := repeat
+  match goal with
+  | H : _ :: _ = _ :: _ |- _ =>
+     apply cons_inv in H; destruct H
+     (* to circumvent bug #2939 in some situations *)
+  | H : _ ++ _ = _ ++ _ |- _ => first
+     [ apply app_inj_tail in H; destruct H
+     | apply app_inv_head in H
+     | apply app_inv_tail in H ]
+  | H : [?x] !! ?i = Some ?y |- _ =>
+     destruct i; [change (Some x = Some y) in H|discriminate]
+  | _ => progress simplify_equality
+  | H : _ |- _ => discriminate_list_equality H
+  end.
+
 (** * General theorems *)
 Section general_properties.
 Context {A : Type}.
 
+Global Instance: ∀ x : A, Injective (=) (=) (x ::).
+Proof. by injection 1. Qed.
+Global Instance: ∀ l : list A, Injective (=) (=) (:: l).
+Proof. by injection 1. Qed.
+Global Instance: ∀ k : list A, Injective (=) (=) (k ++).
+Proof. intros ???. apply app_inv_head. Qed.
+Global Instance: ∀ k : list A, Injective (=) (=) (++ k).
+Proof. intros ???. apply app_inv_tail. Qed.
+
+Lemma app_inj (l1 k1 l2 k2 : list A) :
+  length l1 = length k1 →
+  l1 ++ l2 = k1 ++ k2 → l1 = k1 ∧ l2 = k2.
+Proof. revert k1. induction l1; intros [|??]; naive_solver. Qed.
+
 Lemma list_eq (l1 l2 : list A) : (∀ i, l1 !! i = l2 !! i) → l1 = l2.
 Proof.
   revert l2. induction l1; intros [|??] H.
-  * easy.
+  * done.
   * discriminate (H 0).
   * discriminate (H 0).
-  * f_equal. now injection (H 0). apply IHl1. intro. apply (H (S _)).
+  * f_equal; [by injection (H 0) |].
+    apply IHl1. intro. apply (H (S _)).
 Qed.
+Lemma list_eq_nil (l : list A) : (∀ i, l !! i = None) → l = nil.
+Proof. intros. by apply list_eq. Qed.
 
-Lemma list_lookup_nil i : @nil A !! i = None.
-Proof. now destruct i. Qed.
-Lemma list_lookup_tail (l : list A) i : tail l !! i = l !! S i.
-Proof. now destruct l. Qed.
+Global Instance list_eq_dec {dec : ∀ x y : A, Decision (x = y)} : ∀ l k,
+  Decision (l = k) := list_eq_dec dec.
+Definition list_singleton_dec (l : list A) : { x | l = [x] } + { length l ≠ 1 }.
+Proof.
+ by refine (
+  match l with
+  | [x] => inleft (x ↾ _)
+  | _ => inright _
+  end).
+Defined.
 
-Lemma list_lookup_Some_In (l : list A) i x : l !! i = Some x → In x l.
+Lemma nil_or_length_pos (l : list A) : l = [] ∨ length l ≠ 0.
+Proof. destruct l; simpl; auto with lia. Qed.
+Lemma nil_length (l : list A) : length l = 0 → l = [].
+Proof. by destruct l. Qed.
+Lemma lookup_nil i : @nil A !! i = None.
+Proof. by destruct i. Qed.
+Lemma lookup_tail (l : list A) i : tail l !! i = l !! S i.
+Proof. by destruct l. Qed.
+
+Lemma lookup_lt_length (l : list A) i :
+  is_Some (l !! i) ↔ i < length l.
 Proof.
-  revert i. induction l; intros [|i] ?;
-    simplify_list_lookup; simplify_equality; constructor (solve [eauto]).
+  revert i. induction l.
+  * split; by inversion 1.
+  * intros [|?]; simpl.
+    + split; eauto with arith.
+    + by rewrite <-NPeano.Nat.succ_lt_mono.
+Qed.
+Lemma lookup_lt_length_1 (l : list A) i :
+  is_Some (l !! i) → i < length l.
+Proof. apply lookup_lt_length. Qed.
+Lemma lookup_lt_length_alt (l : list A) i x :
+  l !! i = Some x → i < length l.
+Proof. intros Hl. by rewrite <-lookup_lt_length, Hl. Qed.
+Lemma lookup_lt_length_2 (l : list A) i :
+  i < length l → is_Some (l !! i).
+Proof. apply lookup_lt_length. Qed.
+
+Lemma lookup_ge_length (l : list A) i :
+  l !! i = None ↔ length l ≤ i.
+Proof. rewrite eq_None_not_Some, lookup_lt_length. lia. Qed.
+Lemma lookup_ge_length_1 (l : list A) i :
+  l !! i = None → length l ≤ i.
+Proof. by rewrite lookup_ge_length. Qed.
+Lemma lookup_ge_length_2 (l : list A) i :
+  length l ≤ i → l !! i = None.
+Proof. by rewrite lookup_ge_length. Qed.
+
+Lemma list_eq_length_eq (l1 l2 : list A) :
+  length l2 = length l1 →
+  (∀ i x y, l1 !! i = Some x → l2 !! i = Some y → x = y) →
+  l1 = l2.
+Proof.
+  intros Hlength Hlookup. apply list_eq. intros i.
+  destruct (l2 !! i) as [x|] eqn:E.
+  * feed inversion (lookup_lt_length_2 l1 i) as [y].
+    { pose proof (lookup_lt_length_alt l2 i x E). lia. }
+    f_equal. eauto.
+  * rewrite lookup_ge_length in E |- *. lia.
+Qed.
+
+Lemma lookup_app_l (l1 l2 : list A) i :
+  i < length l1 →
+  (l1 ++ l2) !! i = l1 !! i.
+Proof. revert i. induction l1; intros [|?]; simpl; auto with lia. Qed.
+Lemma lookup_app_l_Some (l1 l2 : list A) i x :
+  l1 !! i = Some x →
+  (l1 ++ l2) !! i = Some x.
+Proof. intros. rewrite lookup_app_l; eauto using lookup_lt_length_alt. Qed.
+
+Lemma lookup_app_r (l1 l2 : list A) i :
+  (l1 ++ l2) !! (length l1 + i) = l2 !! i.
+Proof.
+  revert i.
+  induction l1; intros [|i]; simpl in *; simplify_equality; auto.
+Qed.
+Lemma lookup_app_r_alt (l1 l2 : list A) i :
+  length l1 ≤ i →
+  (l1 ++ l2) !! i = l2 !! (i - length l1).
+Proof.
+  intros. assert (i = length l1 + (i - length l1)) as Hi by lia.
+  rewrite Hi at 1. by apply lookup_app_r.
 Qed.
+Lemma lookup_app_r_Some (l1 l2 : list A) i x :
+  l2 !! i = Some x →
+  (l1 ++ l2) !! (length l1 + i) = Some x.
+Proof. by rewrite lookup_app_r. Qed.
+Lemma lookup_app_r_Some_alt (l1 l2 : list A) i x :
+  length l1 ≤ i →
+  l2 !! (i - length l1) = Some x →
+  (l1 ++ l2) !! i = Some x.
+Proof. intro. by rewrite lookup_app_r_alt. Qed.
 
-Lemma list_lookup_In_Some (l : list A) x : In x l → ∃ i, l !! i = Some x.
+Lemma lookup_app_inv (l1 l2 : list A) i x :
+  (l1 ++ l2) !! i = Some x →
+  l1 !! i = Some x ∨ l2 !! (i - length l1) = Some x.
 Proof.
-  induction l; destruct 1; subst.
-  * now exists 0.
-  * destruct IHl as [i ?]; auto. now exists (S i).
+  revert i.
+  induction l1; intros [|i] ?; simpl in *; simplify_equality; auto.
 Qed.
-Lemma list_lookup_In (l : list A) x : In x l ↔ ∃ i, l !! i = Some x.
-Proof. firstorder eauto using list_lookup_Some_In, list_lookup_In_Some. Qed.
 
-Lemma list_lookup_app_length (l1 l2 : list A) (x : A) :
+Lemma list_lookup_middle (l1 l2 : list A) (x : A) :
   (l1 ++ x :: l2) !! length l1 = Some x.
-Proof. induction l1; simpl; now simplify_list_lookup. Qed.
+Proof. by induction l1; simpl. Qed.
 
-Lemma list_lookup_lt_length (l : list A) i : is_Some (l !! i) ↔ i < length l.
+Lemma alter_length (f : A → A) l i :
+  length (alter f i l) = length l.
+Proof. revert i. induction l; intros [|?]; simpl; auto with lia. Qed.
+Lemma insert_length (l : list A) i x :
+  length (<[i:=x]>l) = length l.
+Proof. apply alter_length. Qed.
+
+Lemma list_lookup_alter (f : A → A) l i :
+  alter f i l !! i = f <$> l !! i.
+Proof. revert i. induction l. done. intros [|i]. done. apply (IHl i). Qed.
+Lemma list_lookup_alter_ne (f : A → A) l i j :
+  i ≠ j → alter f i l !! j = l !! j.
 Proof.
-  revert i. induction l.
-  * split; now inversion 1.
-  * intros [|?]; simplify_list_lookup; simpl.
-    + split; auto with arith.
-    + now rewrite <-NPeano.Nat.succ_lt_mono.
+  revert i j. induction l; [done|].
+  intros [|i] [|j] ?; try done. apply (IHl i). congruence.
+Qed.
+Lemma list_lookup_insert (l : list A) i x :
+  i < length l →
+  <[i:=x]>l !! i = Some x.
+Proof.
+  intros Hi. unfold insert, list_insert.
+  rewrite list_lookup_alter.
+  by feed inversion (lookup_lt_length_2 l i).
 Qed.
-Lemma list_lookup_weaken (l l' : list A) i x :
-  l !! i = Some x → (l ++ l') !! i = Some x.
-Proof. revert i. induction l. discriminate. now intros []. Qed.
+Lemma list_lookup_insert_ne (l : list A) i j x :
+  i ≠ j → <[i:=x]>l !! j = l !! j.
+Proof. apply list_lookup_alter_ne. Qed.
 
-Lemma fold_right_permutation {B} (f : A → B → B) (b : B) :
-  (∀ a1 a2 b, f a1 (f a2 b) = f a2 (f a1 b)) →
-  Proper (Permutation ==> (=)) (fold_right f b).
-Proof. intro. induction 1; simpl; congruence. Qed.
+Lemma list_lookup_other (l : list A) i x :
+  length l ≠ 1 →
+  l !! i = Some x →
+  ∃ j y, j ≠ i ∧ l !! j = Some y.
+Proof.
+  intros Hl Hi.
+  destruct i; destruct l as [|x0 [|x1 l]]; simpl in *; simplify_equality.
+  * by exists 1 x1.
+  * by exists 0 x0.
+Qed.
 
-Lemma Forall_impl (P Q : A → Prop) l :
-  Forall P l → (∀ x, P x → Q x) → Forall Q l.
-Proof. induction 1; auto. Qed.
+Lemma alter_app_l (f : A → A) (l1 l2 : list A) i :
+  i < length l1 →
+  alter f i (l1 ++ l2) = alter f i l1 ++ l2.
+Proof.
+  revert i.
+  induction l1; intros [|?] ?; simpl in *; f_equal; auto with lia.
+Qed.
+Lemma alter_app_r (f : A → A) (l1 l2 : list A) i :
+  alter f (length l1 + i) (l1 ++ l2) = l1 ++ alter f i l2.
+Proof.
+  revert i.
+  induction l1; intros [|?]; simpl in *; f_equal; auto.
+Qed.
+Lemma alter_app_r_alt (f : A → A) (l1 l2 : list A) i :
+  length l1 ≤ i →
+  alter f i (l1 ++ l2) = l1 ++ alter f (i - length l1) l2.
+Proof.
+  intros. assert (i = length l1 + (i - length l1)) as Hi by lia.
+  rewrite Hi at 1. by apply alter_app_r.
+Qed.
 
-Lemma Forall2_length {B} (P : A → B → Prop) l1 l2 :
-  Forall2 P l1 l2 → length l1 = length l2.
-Proof. induction 1; simpl; auto. Qed.
+Lemma insert_app_l (l1 l2 : list A) i x :
+  i < length l1 →
+  <[i:=x]>(l1 ++ l2) = <[i:=x]>l1 ++ l2.
+Proof. apply alter_app_l. Qed.
+Lemma insert_app_r (l1 l2 : list A) i x :
+  <[length l1 + i:=x]>(l1 ++ l2) = l1 ++ <[i:=x]>l2.
+Proof. apply alter_app_r. Qed.
+Lemma insert_app_r_alt (l1 l2 : list A) i x :
+  length l1 ≤ i →
+  <[i:=x]>(l1 ++ l2) = l1 ++ <[i - length l1:=x]>l2.
+Proof. apply alter_app_r_alt. Qed.
 
-Lemma Forall2_impl {B} (P Q : A → B → Prop) l1 l2 :
-  Forall2 P l1 l2 → (∀ x y, P x y → Q x y) → Forall2 Q l1 l2.
-Proof. induction 1; auto. Qed.
+Lemma insert_consecutive_length (l : list A) i k :
+  length (insert_consecutive i k l) = length l.
+Proof. revert i. by induction k; intros; simpl; rewrite ?insert_length. Qed.
 
-Lemma Forall2_unique {B} (P : A → B → Prop) l k1 k2 :
-  Forall2 P l k1 →
-  Forall2 P l k2 →
-  (∀ x y1 y2, P x y1 → P x y2 → y1 = y2) →
-  k1 = k2.
+Lemma not_elem_of_nil (x : A) : x ∉ [].
+Proof. by inversion 1. Qed.
+Lemma elem_of_nil (x : A) : x ∈ [] ↔ False.
+Proof. intuition. by destruct (not_elem_of_nil x). Qed.
+Lemma elem_of_nil_inv (l : list A) : (∀ x, x ∉ l) → l = [].
+Proof. destruct l. done. by edestruct 1; constructor. Qed.
+Lemma elem_of_cons (l : list A) x y :
+  x ∈ y :: l ↔ x = y ∨ x ∈ l.
+Proof.
+  split.
+  * inversion 1; subst. by left. by right.
+  * intros [?|?]; subst. by left. by right.
+Qed.
+Lemma not_elem_of_cons (l : list A) x y :
+  x ∉ y :: l ↔ x ≠ y ∧ x ∉ l.
+Proof. rewrite elem_of_cons. tauto. Qed.
+Lemma elem_of_app (l1 l2 : list A) x :
+  x ∈ l1 ++ l2 ↔ x ∈ l1 ∨ x ∈ l2.
 Proof.
-  intros H. revert k2. induction H; inversion_clear 1; intros; f_equal; eauto.
+  induction l1.
+  * split; [by right|]. intros [Hx|]; [|done].
+    by destruct (elem_of_nil x).
+  * simpl. rewrite !elem_of_cons, IHl1. tauto.
 Qed.
+Lemma not_elem_of_app (l1 l2 : list A) x :
+  x ∉ l1 ++ l2 ↔ x ∉ l1 ∧ x ∉ l2.
+Proof. rewrite elem_of_app. tauto. Qed.
 
+Lemma elem_of_list_singleton (x y : A) : x ∈ [y] ↔ x = y.
+Proof. rewrite elem_of_cons, elem_of_nil. tauto. Qed.
+
+Global Instance elem_of_list_permutation_proper (x : A) :
+  Proper (Permutation ==> iff) (x ∈).
+Proof. induction 1; rewrite ?elem_of_nil, ?elem_of_cons; intuition. Qed.
+
+Lemma elem_of_list_split (l : list A) x :
+  x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2.
+Proof.
+  induction 1 as [x l|x y l ? [l1 [l2 ?]]].
+  * by eexists [], l.
+  * subst. by exists (y :: l1) l2.
+Qed.
+
+Global Instance elem_of_list_dec {dec : ∀ x y : A, Decision (x = y)} :
+  ∀ (x : A) l, Decision (x ∈ l).
+Proof.
+ intros x. refine (
+  fix go l :=
+  match l return Decision (x ∈ l) with
+  | [] => right (not_elem_of_nil _)
+  | y :: l => cast_if_or (decide_rel (=) x y) (go l)
+  end); clear go dec; subst; try (by constructor); by inversion 1.
+Defined.
+
+Lemma elem_of_list_lookup_1 (l : list A) x :
+  x ∈ l → ∃ i, l !! i = Some x.
+Proof.
+  induction 1 as [|???? IH].
+  * by exists 0.
+  * destruct IH as [i ?]; auto. by exists (S i).
+Qed.
+Lemma elem_of_list_lookup_2 (l : list A) i x :
+  l !! i = Some x → x ∈ l.
+Proof.
+  revert i. induction l; intros [|i] ?;
+    simpl; simplify_equality; constructor; eauto.
+Qed.
+Lemma elem_of_list_lookup (l : list A) x :
+  x ∈ l ↔ ∃ i, l !! i = Some x.
+Proof.
+  firstorder eauto using
+    elem_of_list_lookup_1, elem_of_list_lookup_2.
+Qed.
+
+Lemma NoDup_nil : NoDup (@nil A) ↔ True.
+Proof. split; constructor. Qed.
+Lemma NoDup_cons (x : A) l : NoDup (x :: l) ↔ x ∉ l ∧ NoDup l.
+Proof. split. by inversion 1. intros [??]. by constructor. Qed.
+Lemma NoDup_cons_11 (x : A) l : NoDup (x :: l) → x ∉ l.
+Proof. rewrite NoDup_cons. by intros [??]. Qed.
+Lemma NoDup_cons_12 (x : A) l : NoDup (x :: l) → NoDup l.
+Proof. rewrite NoDup_cons. by intros [??]. Qed.
 Lemma NoDup_singleton (x : A) : NoDup [x].
-Proof. constructor. easy. constructor. Qed.
+Proof. constructor. apply not_elem_of_nil. constructor. Qed.
+
 Lemma NoDup_app (l k : list A) :
-  NoDup l → NoDup k → (∀ x, In x l → ¬In x k) → NoDup (l ++ k).
+  NoDup (l ++ k) ↔ NoDup l ∧ (∀ x, x ∈ l → x ∉ k) ∧ NoDup k.
 Proof.
-  induction 1; simpl.
-  * easy.
-  * constructor. rewrite in_app_iff. firstorder. firstorder.
+  induction l; simpl.
+  * rewrite NoDup_nil.
+    setoid_rewrite elem_of_nil. naive_solver.
+  * rewrite !NoDup_cons.
+    setoid_rewrite elem_of_cons. setoid_rewrite elem_of_app.
+    naive_solver.
 Qed.
 
-Global Instance: ∀ k : list A, Injective (=) (=) (k ++).
-Proof. intros ???. apply app_inv_head. Qed.
-Global Instance: ∀ k : list A, Injective (=) (=) (++ k).
-Proof. intros ???. apply app_inv_tail. Qed.
-
-Lemma in_nil_inv (l : list A) : (∀ x, ¬In x l) → l = [].
-Proof. destruct l. easy. now edestruct 1; constructor. Qed.
-Lemma nil_length (l : list A) : length l = 0 → l = [].
-Proof. destruct l. easy. discriminate. Qed.
-Lemma NoDup_inv_1 (x : A) (l : list A) : NoDup (x :: l) → ¬In x l.
-Proof. now inversion_clear 1. Qed.
-Lemma NoDup_inv_2 (x : A) (l : list A) : NoDup (x :: l) → NoDup l.
-Proof. now inversion_clear 1. Qed.
-Lemma Forall_inv_2 (P : A → Prop) x l : Forall P (x :: l) → Forall P l.
-Proof. now inversion 1. Qed.
-Lemma Exists_inv (P : A → Prop) x l : Exists P (x :: l) → P x ∨ Exists P l.
-Proof. inversion 1; intuition. Qed.
+Global Instance NoDup_proper:
+  Proper (Permutation ==> iff) (@NoDup A).
+Proof.
+  induction 1 as [|x l k Hlk IH | |].
+  * by rewrite !NoDup_nil.
+  * by rewrite !NoDup_cons, IH, Hlk.
+  * rewrite !NoDup_cons, !elem_of_cons. intuition.
+  * intuition.
+Qed.
 
-Global Instance list_eq_dec {dec : ∀ x y : A, Decision (x = y)} : ∀ l k,
-  Decision (l = k) := list_eq_dec dec.
-Global Instance list_in_dec {dec : ∀ x y : A, Decision (x = y)} : ∀ x l,
-  Decision (In x l) := in_dec dec.
+Lemma NoDup_Permutation (l k : list A) :
+  NoDup l → NoDup k → (∀ x, x ∈ l ↔ x ∈ k) → Permutation l k.
+Proof.
+  intros Hl. revert k. induction Hl as [|x l Hin ? IH].
+  * intros k _ Hk.
+    rewrite (elem_of_nil_inv k); [done |].
+    intros x. rewrite <-Hk, elem_of_nil. intros [].
+  * intros k Hk Hlk.
+    destruct (elem_of_list_split k x) as [l1 [l2 ?]]; subst.
+    { rewrite <-Hlk. by constructor. }
+    rewrite <-Permutation_middle, NoDup_cons in Hk.
+    destruct Hk as [??].
+    apply Permutation_cons_app, IH; [done |].
+    intros y. specialize (Hlk y).
+    rewrite <-Permutation_middle, !elem_of_cons in Hlk.
+    naive_solver.
+Qed.
 
 Global Instance NoDup_dec {dec : ∀ x y : A, Decision (x = y)} :
     ∀ (l : list A), Decision (NoDup l) :=
   fix NoDup_dec l :=
   match l return Decision (NoDup l) with
-  | [] => left NoDup_nil
+  | [] => left NoDup_nil_2
   | x :: l =>
-    match In_dec dec x l with
-    | left Hin => right (λ H, NoDup_inv_1 _ _ H Hin)
+    match decide_rel (∈) x l with
+    | left Hin => right (λ H, NoDup_cons_11 _ _ H Hin)
     | right Hin =>
       match NoDup_dec l with
-      | left H => left (NoDup_cons _ Hin H)
-      | right H => right (H ∘ NoDup_inv_2 _ _)
+      | left H => left (NoDup_cons_2 _ _ Hin H)
+      | right H => right (H ∘ NoDup_cons_12 _ _)
       end
     end
   end.
 
-Global Instance Forall_dec (P : A → Prop) {dec : ∀ x, Decision (P x)} :
-    ∀ l, Decision (Forall P l) :=
-  fix go (l : list A) :=
-  match l return {Forall P l} + {¬Forall P l} with
-  | [] => left (Forall_nil _)
-  | x :: l =>
-    match dec x with
-    | left Hx =>
-      match go l with
-      | left Hl => left (Forall_cons _ Hx Hl)
-      | right Hl => right (Hl ∘ Forall_inv_2 _ _ _)
-      end
-    | right Hx => right (Hx ∘ @Forall_inv _ _ _ _)
-    end
-  end.
+Section remove_dups.
+  Context `{!∀ x y : A, Decision (x = y)}.
 
-Global Instance Exists_dec (P : A → Prop) {dec : ∀ x, Decision (P x)} :
-    ∀ l, Decision (Exists P l) :=
-  fix go (l : list A) :=
-  match l return {Exists P l} + {¬Exists P l} with
-  | [] => right (proj1 (Exists_nil _))
-  | x :: l =>
-    match dec x with
-    | left Hx => left (Exists_cons_hd _ _ _ Hx)
-    | right Hx =>
-      match go l with
-      | left Hl => left (Exists_cons_tl _ Hl)
-      | right Hl => right (or_ind Hx Hl ∘ Exists_inv _ _ _)
-      end
-    end
-  end.
-End general_properties.
+  Fixpoint remove_dups (l : list A) : list A :=
+    match l with
+    | [] => []
+    | x :: l =>
+      if decide_rel (∈) x l then remove_dups l else x :: remove_dups l
+    end.
 
-(** * Theorems on the prefix and suffix predicates *)
-Section prefix_postfix.
-Context {A : Type}.
+  Lemma elem_of_remove_dups l x :
+    x ∈ remove_dups l ↔ x ∈ l.
+  Proof.
+    split; induction l; simpl; repeat case_decide;
+      rewrite ?elem_of_cons; intuition (simplify_equality; auto).
+  Qed.
+
+  Lemma remove_dups_nodup l : NoDup (remove_dups l).
+  Proof.
+    induction l; simpl; repeat case_decide; try constructor; auto.
+    by rewrite elem_of_remove_dups.
+  Qed.
+End remove_dups.
 
-Global Instance: PreOrder (@prefix_of A).
+Lemma elem_of_list_filter `{∀ x : A, Decision (P x)} l x :
+  x ∈ filter P l ↔ P x ∧ x ∈ l.
 Proof.
-  split.
-   intros ?. eexists []. now rewrite app_nil_r.
-  intros ??? [k1 ?] [k2 ?]. exists (k1 ++ k2). subst. now rewrite app_assoc.
-Qed.
-
-Lemma prefix_of_nil (l : list A) : prefix_of [] l.
-Proof. now exists l. Qed.
-Lemma prefix_of_nil_not x (l : list A) : ¬prefix_of (x :: l) [].
-Proof. intros [k E]. discriminate. Qed.
-Lemma prefix_of_cons x y (l1 l2 : list A) :
-  x = y → prefix_of l1 l2 → prefix_of (x :: l1) (y :: l2).
-Proof. intros ? [k E]. exists k. now subst. Qed.
-Lemma prefix_of_cons_inv_1 x y (l1 l2 : list A) :
-  prefix_of (x :: l1) (y :: l2) → x = y.
-Proof. intros [k E]. now injection E. Qed.
-Lemma prefix_of_cons_inv_2 x y (l1 l2 : list A) :
-  prefix_of (x :: l1) (y :: l2) → prefix_of l1 l2.
-Proof. intros [k E]. exists k. now injection E. Qed.
-
-Lemma prefix_of_app_l (l1 l2 l3 : list A) :
-  prefix_of (l1 ++ l3) l2 → prefix_of l1 l2.
-Proof. intros [k ?]. red. exists (l3 ++ k). subst. now rewrite <-app_assoc. Qed.
-Lemma prefix_of_app_r (l1 l2 l3 : list A) :
-  prefix_of l1 l2 → prefix_of l1 (l2 ++ l3).
-Proof. intros [k ?]. exists (k ++ l3). subst. now rewrite app_assoc. Qed.
-
-Global Instance prefix_of_dec `{∀ x y : A, Decision (x = y)} :
-    ∀ l1 l2, Decision (prefix_of l1 l2) :=
-  fix prefix_of_dec l1 l2 :=
-  match l1, l2 return { prefix_of l1 l2 } + { ¬prefix_of l1 l2 } with
-  | [], _ => left (prefix_of_nil _)
-  | _, [] => right (prefix_of_nil_not _ _)
-  | x :: l1, y :: l2 =>
-    match decide_rel (=) x y with
-    | left Exy =>
-      match prefix_of_dec l1 l2 with
-      | left Hl1l2 => left (prefix_of_cons _ _ _ _ Exy Hl1l2)
-      | right Hl1l2 => right (Hl1l2 ∘ prefix_of_cons_inv_2 _ _ _ _)
-      end
-    | right Exy => right (Exy ∘ prefix_of_cons_inv_1 _ _ _ _)
-    end
-  end.
+  unfold filter. induction l; simpl; repeat case_decide;
+     rewrite ?elem_of_nil, ?elem_of_cons; naive_solver.
+Qed.
+Lemma filter_nodup P `{∀ x : A, Decision (P x)} l :
+  NoDup l → NoDup (filter P l).
+Proof.
+  unfold filter. induction 1; simpl; repeat case_decide;
+    rewrite ?NoDup_nil, ?NoDup_cons, ?elem_of_list_filter; tauto.
+Qed.
+
+Lemma reverse_nil : reverse [] = @nil A.
+Proof. done. Qed.
+Lemma reverse_cons (l : list A) x : reverse (x :: l) = reverse l ++ [x].
+Proof. unfold reverse. by rewrite <-!rev_alt. Qed.
+Lemma reverse_snoc (l : list A) x : reverse (l ++ [x]) = x :: reverse l.
+Proof. unfold reverse. by rewrite <-!rev_alt, rev_unit. Qed.
+Lemma reverse_app (l1 l2 : list A) :
+  reverse (l1 ++ l2) = reverse l2 ++ reverse l1.
+Proof. unfold reverse. rewrite <-!rev_alt. apply rev_app_distr. Qed.
+Lemma reverse_length (l : list A) : length (reverse l) = length l.
+Proof. unfold reverse. rewrite <-!rev_alt. apply rev_length. Qed.
+Lemma reverse_involutive (l : list A) : reverse (reverse l) = l.
+Proof. unfold reverse. rewrite <-!rev_alt. apply rev_involutive. Qed. 
 
-Global Instance: PreOrder (@suffix_of A).
+Lemma take_nil n :
+  take n (@nil A) = [].
+Proof. by destruct n. Qed.
+Lemma take_app (l k : list A) :
+  take (length l) (l ++ k) = l.
+Proof. induction l; simpl; f_equal; auto. Qed.
+Lemma take_app_alt (l k : list A) n :
+  n = length l →
+  take n (l ++ k) = l.
+Proof. intros Hn. by rewrite Hn, take_app. Qed.
+Lemma take_app_le (l k : list A) n :
+  n ≤ length l →
+  take n (l ++ k) = take n l.
 Proof.
-  split.
-  * intros ?. now eexists [].
-  * intros ??? [k1 ?] [k2 ?].
-    exists (k2 ++ k1). subst. now rewrite app_assoc.
+  revert n;
+  induction l; intros [|?] ?; simpl in *; f_equal; auto with lia.
+Qed.
+Lemma take_app_ge (l k : list A) n :
+  length l ≤ n →
+  take n (l ++ k) = l ++ take (n - length l) k.
+Proof.
+  revert n;
+  induction l; intros [|?] ?; simpl in *; f_equal; auto with lia.
+Qed.
+Lemma take_ge (l : list A) n :
+  length l ≤ n →
+  take n l = l.
+Proof.
+  revert n.
+  induction l; intros [|?] ?; simpl in *; f_equal; auto with lia.
 Qed.
 
-Lemma prefix_suffix_rev (l1 l2 : list A) :
-  prefix_of l1 l2 ↔ suffix_of (rev l1) (rev l2).
+Lemma take_take (l : list A) n m :
+  take n (take m l) = take (min n m) l.
+Proof. revert n m. induction l; intros [|?] [|?]; simpl; f_equal; auto. Qed.
+Lemma take_idempotent (l : list A) n :
+  take n (take n l) = take n l.
+Proof. by rewrite take_take, Min.min_idempotent. Qed.
+
+Lemma take_length (l : list A) n :
+  length (take n l) = min n (length l).
+Proof. revert n. induction l; intros [|?]; simpl; f_equal; done. Qed.
+Lemma take_length_alt (l : list A) n :
+  n ≤ length l →
+  length (take n l) = n.
+Proof. rewrite take_length. apply Min.min_l. Qed.
+
+Lemma lookup_take (l : list A) n i :
+  i < n → take n l !! i = l !! i.
+Proof.
+  revert n i. induction l; intros [|n] i ?; trivial.
+  * auto with lia.
+  * destruct i; simpl; auto with arith.
+Qed.
+Lemma lookup_take_ge (l : list A) n i :
+  n ≤ i → take n l !! i = None.
+Proof.
+  revert n i.
+  induction l; intros [|?] [|?] ?; simpl; auto with lia.
+Qed.
+Lemma take_alter (f : A → A) l n i :
+  n ≤ i → take n (alter f i l) = take n l.
+Proof.
+  intros. apply list_eq. intros j. destruct (le_lt_dec n j).
+  * by rewrite !lookup_take_ge.
+  * by rewrite !lookup_take, !list_lookup_alter_ne by lia.
+Qed.
+Lemma take_insert (l : list A) n i x :
+  n ≤ i → take n (<[i:=x]>l) = take n l.
+Proof take_alter _ _ _ _.
+
+Lemma drop_nil n :
+  drop n (@nil A) = [].
+Proof. by destruct n. Qed.
+Lemma drop_app (l k : list A) :
+  drop (length l) (l ++ k) = k.
+Proof. induction l; simpl; f_equal; auto. Qed.
+Lemma drop_app_alt (l k : list A) n :
+  n = length l →
+  drop n (l ++ k) = k.
+Proof. intros Hn. by rewrite Hn, drop_app. Qed.
+Lemma drop_length (l : list A) n :
+  length (drop n l) = length l - n.
 Proof.
-  split; intros [k E]; exists (rev k).
-  * now rewrite E, rev_app_distr.
-  * now rewrite <-(rev_involutive l2), E, rev_app_distr, rev_involutive.
+  revert n. by induction l; intros [|i]; simpl; f_equal.
 Qed.
-Lemma suffix_prefix_rev (l1 l2 : list A) :
-  suffix_of l1 l2 ↔ prefix_of (rev l1) (rev l2).
-Proof. now rewrite prefix_suffix_rev, !rev_involutive. Qed.
+Lemma drop_all (l : list A) :
+  drop (length l) l = [].
+Proof. induction l; simpl; auto. Qed.
+Lemma drop_all_alt (l : list A) n :
+  n = length l →
+  drop n l = [].
+Proof. intros. subst. by rewrite drop_all. Qed.
 
-Lemma suffix_of_nil (l : list A) : suffix_of [] l.
-Proof. exists l. now rewrite app_nil_r. Qed.
-Lemma suffix_of_nil_not x (l : list A) : ¬suffix_of (x :: l) [].
-Proof. intros [[] ?]; discriminate. Qed.
-Lemma suffix_of_cons x y (l1 l2 : list A) :
-  x = y → suffix_of l1 l2 → suffix_of (l1 ++ [x]) (l2 ++ [y]).
-Proof. intros ? [k E]. exists k. subst. now rewrite app_assoc. Qed.
-Lemma suffix_of_snoc_inv_1 x y (l1 l2 : list A) :
-  suffix_of (l1 ++ [x]) (l2 ++ [y]) → x = y.
+Lemma lookup_drop (l : list A) n i :
+  drop n l !! i = l !! (n + i).
+Proof. revert n i. induction l; intros [|i] ?; simpl; auto. Qed.
+Lemma drop_alter (f : A → A) l n i  :
+  i < n → drop n (alter f i l) = drop n l.
 Proof.
-  rewrite suffix_prefix_rev, !rev_unit. now apply prefix_of_cons_inv_1.
+  intros. apply list_eq. intros j.
+  by rewrite !lookup_drop, !list_lookup_alter_ne by lia.
 Qed.
-Lemma suffix_of_snoc_inv_2 x y (l1 l2 : list A) :
-  suffix_of (l1 ++ [x]) (l2 ++ [y]) → suffix_of l1 l2.
+Lemma drop_insert (l : list A) n i x :
+  i < n → drop n (<[i:=x]>l) = drop n l.
+Proof drop_alter _ _ _ _.
+
+Lemma replicate_length n (x : A) : length (replicate n x) = n.
+Proof. induction n; simpl; auto. Qed.
+Lemma lookup_replicate n (x : A) i :
+  i < n →
+  replicate n x !! i = Some x.
 Proof.
-  rewrite !suffix_prefix_rev, !rev_unit. now apply prefix_of_cons_inv_2.
+  revert i.
+  induction n; intros [|?]; naive_solver auto with lia.
 Qed.
+Lemma lookup_replicate_inv n (x y : A) i :
+  replicate n x !! i = Some y → y = x ∧ i < n.
+Proof.
+  revert i.
+  induction n; intros [|?]; naive_solver auto with lia.
+Qed.
+Lemma replicate_plus n m (x : A) :
+  replicate (n + m) x = replicate n x ++ replicate m x.
+Proof. induction n; simpl; f_equal; auto. Qed.
 
-Lemma suffix_of_cons_l (l1 l2 : list A) x :
-  suffix_of (x :: l1) l2 → suffix_of l1 l2.
-Proof. intros [k ?]. exists (k ++ [x]). subst. now rewrite <-app_assoc. Qed.
-Lemma suffix_of_app_l (l1 l2 l3 : list A) :
-  suffix_of (l3 ++ l1) l2 → suffix_of l1 l2.
-Proof. intros [k ?]. exists (k ++ l3). subst. now rewrite <-app_assoc. Qed.
-Lemma suffix_of_cons_r (l1 l2 : list A) x :
-  suffix_of l1 l2 → suffix_of l1 (x :: l2).
-Proof. intros [k ?]. exists (x :: k). now subst. Qed.
-Lemma suffix_of_app_r (l1 l2 l3 : list A) :
-  suffix_of l1 l2 → suffix_of l1 (l3 ++ l2).
-Proof. intros [k ?]. exists (l3 ++ k). subst. now rewrite app_assoc. Qed.
+Lemma take_replicate n m (x : A) :
+  take n (replicate m x) = replicate (min n m) x.
+Proof. revert m. by induction n; intros [|?]; simpl; f_equal. Qed.
+Lemma take_replicate_plus n m (x : A) :
+  take n (replicate (n + m) x) = replicate n x.
+Proof. by rewrite take_replicate, min_l by lia. Qed.
+Lemma drop_replicate n m (x : A) :
+  drop n (replicate m x) = replicate (m - n) x.
+Proof. revert m. by induction n; intros [|?]; simpl; f_equal. Qed.
+Lemma drop_replicate_plus n m (x : A) :
+  drop n (replicate (n + m) x) = replicate m x.
+Proof. rewrite drop_replicate. f_equal. lia. Qed.
 
-Lemma suffix_of_cons_inv (l1 l2 : list A) x y :
-  suffix_of (x :: l1) (y :: l2) → x :: l1 = y :: l2 ∨ suffix_of (x :: l1) l2.
+Lemma resize_spec (l : list A) n x :
+  resize n x l = take n l ++ replicate (n - length l) x.
 Proof.
-  intros [[|? k] E].
-   now left.
-  right. simplify_equality. now apply suffix_of_app_r.
+  revert n.
+  induction l; intros [|?]; simpl; f_equal; auto.
 Qed.
-Lemma suffix_of_cons_not x (l : list A) : ¬suffix_of (x :: l) l.
+Lemma resize_0 (l : list A) x :
+  resize 0 x l = [].
+Proof. by destruct l. Qed.
+Lemma resize_nil n (x : A) :
+  resize n x [] = replicate n x.
+Proof. rewrite resize_spec. rewrite take_nil. simpl. f_equal. lia. Qed.
+Lemma resize_ge (l : list A) n x :
+  length l ≤ n →
+  resize n x l = l ++ replicate (n - length l) x.
+Proof. intros. by rewrite resize_spec, take_ge. Qed.
+Lemma resize_le (l : list A) n x :
+  n ≤ length l →
+  resize n x l = take n l.
 Proof.
-  intros [k E]. change ([] ++ l = k ++ [x] ++ l) in E.
-  rewrite app_assoc in E. apply app_inv_tail in E.
-  destruct k; simplify_equality.
+  intros. rewrite resize_spec, (proj2 (NPeano.Nat.sub_0_le _ _)) by done.
+  simpl. by rewrite app_nil_r.
 Qed.
 
-Global Program Instance suffix_of_dec `{∀ x y : A, Decision (x = y)} l1 l2 :
-    Decision (suffix_of l1 l2) :=
-  match decide_rel prefix_of (rev_append l1 []) (rev_append l2 []) with
-  | left Hpre => left _
-  | right Hpre => right _
-  end.
-Next Obligation. apply suffix_prefix_rev. now rewrite !rev_alt. Qed.
-Next Obligation.
-  intro. destruct Hpre. rewrite <-!rev_alt.
-  now apply suffix_prefix_rev.
+Lemma resize_all (l : list A) x :
+  resize (length l) x l = l.
+Proof. intros. by rewrite resize_le, take_ge. Qed.
+Lemma resize_all_alt (l : list A) n x :
+  n = length l →
+  resize n x l = l.
+Proof. intros. subst. by rewrite resize_all. Qed.
+
+Lemma resize_plus (l : list A) n m x :
+  resize (n + m) x l = resize n x l ++ resize m x (drop n l).
+Proof.
+  revert n m.
+  induction l; intros [|?] [|?]; simpl; f_equal; auto.
+  * by rewrite plus_0_r, app_nil_r.
+  * by rewrite replicate_plus.
+Qed.
+Lemma resize_plus_eq (l : list A) n m x :
+  length l = n →
+  resize (n + m) x l = l ++ replicate m x.
+Proof.
+  intros. subst.
+  by rewrite resize_plus, resize_all, drop_all, resize_nil.
+Qed.
+
+Lemma resize_app_le (l1 l2 : list A) n x :
+  n ≤ length l1 →
+  resize n x (l1 ++ l2) = resize n x l1.
+Proof.
+  intros.
+  by rewrite !resize_le, take_app_le by (rewrite ?app_length; lia).
+Qed.
+Lemma resize_app_ge (l1 l2 : list A) n x :
+  length l1 ≤ n →
+  resize n x (l1 ++ l2) = l1 ++ resize (n - length l1) x l2.
+Proof.
+  intros.
+  rewrite !resize_spec, take_app_ge, app_assoc by done.
+  do 2 f_equal. rewrite app_length. lia.
+Qed.
+
+Lemma resize_length (l : list A) n x : length (resize n x l) = n.
+Proof.
+  rewrite resize_spec, app_length, replicate_length, take_length. lia.
+Qed.
+Lemma resize_replicate (x : A) n m :
+  resize n x (replicate m x) = replicate n x.
+Proof. revert m. induction n; intros [|?]; simpl; f_equal; auto. Qed.
+
+Lemma resize_resize (l : list A) n m x :
+  n ≤ m →
+  resize n x (resize m x l) = resize n x l.
+Proof.
+  revert n m. induction l; simpl.
+  * intros. by rewrite !resize_nil, resize_replicate.
+  * intros [|?] [|?] ?; simpl; f_equal; auto with lia.
 Qed.
+Lemma resize_idempotent (l : list A) n x :
+  resize n x (resize n x l) = resize n x l.
+Proof. by rewrite resize_resize. Qed.
+
+Lemma resize_take_le (l : list A) n m x :
+  n ≤ m →
+  resize n x (take m l) = resize n x l.
+Proof.
+  revert n m.
+  induction l; intros [|?] [|?] ?; simpl; f_equal; auto with lia.
+Qed.
+Lemma resize_take_eq (l : list A) n x :
+  resize n x (take n l) = resize n x l.
+Proof. by rewrite resize_take_le. Qed.
+
+Lemma take_resize (l : list A) n m x :
+  take n (resize m x l) = resize (min n m) x l.
+Proof.
+  revert n m.
+  induction l; intros [|?] [|?]; simpl; f_equal; auto using take_replicate.
+Qed.
+Lemma take_resize_le (l : list A) n m x :
+  n ≤ m →
+  take n (resize m x l) = resize n x l.
+Proof. intros. by rewrite take_resize, Min.min_l. Qed.
+Lemma take_resize_eq (l : list A) n x :
+  take n (resize n x l) = resize n x l.
+Proof. intros. by rewrite take_resize, Min.min_l. Qed.
+Lemma take_length_resize (l : list A) n x :
+  length l ≤ n →
+  take (length l) (resize n x l) = l.
+Proof. intros. by rewrite take_resize_le, resize_all. Qed.
+Lemma take_length_resize_alt (l : list A) n m x :
+  m = length l →
+  m ≤ n →
+  take m (resize n x l) = l.
+Proof. intros. subst. by apply take_length_resize. Qed.
+Lemma take_resize_plus (l : list A) n m x :
+  take n (resize (n + m) x l) = resize n x l.
+Proof. by rewrite take_resize, min_l by lia. Qed.
+
+Lemma drop_resize_le (l : list A) n m x :
+  n ≤ m →
+  drop n (resize m x l) = resize (m - n) x (drop n l).
+Proof.
+  revert n m. induction l; simpl.
+  * intros. by rewrite drop_nil, !resize_nil, drop_replicate.
+  * intros [|?] [|?] ?; simpl; try case_match; auto with lia.
+Qed.
+Lemma drop_resize_plus (l : list A) n m x :
+  drop n (resize (n + m) x l) = resize m x (drop n l).
+Proof. rewrite drop_resize_le by lia. f_equal. lia. Qed.
+
+Section Forall_Exists.
+  Context (P : A → Prop).
+
+  Lemma Forall_forall l :
+    Forall P l ↔ ∀ x, x ∈ l → P x.
+  Proof.
+    split.
+    * induction 1; inversion 1; subst; auto.
+    * intros Hin. induction l; constructor.
+      + apply Hin. constructor.
+      + apply IHl. intros ??. apply Hin. by constructor.
+  Qed.
+
+  Lemma Forall_nil : Forall P [] ↔ True.
+  Proof. done. Qed.
+  Lemma Forall_cons_1 x l : Forall P (x :: l) → P x ∧ Forall P l.
+  Proof. by inversion 1. Qed.
+  Lemma Forall_cons x l : Forall P (x :: l) ↔ P x ∧ Forall P l.
+  Proof. split. by inversion 1. intros [??]. by constructor. Qed.
+  Lemma Forall_singleton x : Forall P [x] ↔ P x.
+  Proof. rewrite Forall_cons, Forall_nil; tauto. Qed.
+  Lemma Forall_app l1 l2 : Forall P (l1 ++ l2) ↔ Forall P l1 ∧ Forall P l2.
+  Proof.
+    split.
+    * induction l1; inversion 1; intuition.
+    * intros [H ?]. induction H; simpl; intuition.
+  Qed.
+  Lemma Forall_true l : (∀ x, P x) → Forall P l.
+  Proof. induction l; auto. Qed.
+  Lemma Forall_impl l (Q : A → Prop) :
+    Forall P l → (∀ x, P x → Q x) → Forall Q l.
+  Proof. intros H ?. induction H; auto. Defined.
+
+  Lemma Forall_delete l i : Forall P l → Forall P (delete i l).
+  Proof.
+    intros H. revert i.
+    by induction H; intros [|i]; try constructor.
+  Qed.
+  Lemma Forall_lookup l :
+    Forall P l ↔ ∀ i x, l !! i = Some x → P x.
+  Proof.
+    rewrite Forall_forall.
+    setoid_rewrite elem_of_list_lookup.
+    naive_solver.
+  Qed.
+  Lemma Forall_lookup_1 l i x :
+    Forall P l → l !! i = Some x → P x.
+  Proof. rewrite Forall_lookup. eauto. Qed.
+  Lemma Forall_alter f l i :
+    Forall P l →
+    (∀ x, l !! i = Some x → P x → P (f x)) →
+    Forall P (alter f i l).
+  Proof.
+    intros Hl. revert i.
+    induction Hl; simpl; intros [|i]; constructor; auto.
+  Qed.
+
+  Lemma Forall_replicate n x :
+    P x → Forall P (replicate n x).
+  Proof. induction n; simpl; constructor; auto. Qed.
+  Lemma Forall_replicate_eq n (x : A) :
+    Forall (=x) (replicate n x).
+  Proof. induction n; simpl; constructor; auto. Qed.
+
+  Lemma Exists_exists l :
+    Exists P l ↔ ∃ x, x ∈ l ∧ P x.
+  Proof.
+    split.
+    * induction 1 as [x|y ?? IH].
+      + exists x. split. constructor. done.
+      + destruct IH as [x [??]]. exists x. split. by constructor. done. 
+    * intros [x [Hin ?]]. induction l.
+      + by destruct (not_elem_of_nil x).
+      + inversion Hin; subst. by left. right; auto.
+  Qed.
+  Lemma Exists_inv x l : Exists P (x :: l) → P x ∨ Exists P l.
+  Proof. inversion 1; intuition trivial. Qed.
+  Lemma Exists_app l1 l2 : Exists P (l1 ++ l2) ↔ Exists P l1 ∨ Exists P l2.
+  Proof.
+    split.
+    * induction l1; inversion 1; intuition.
+    * intros [H|H].
+      + induction H; simpl; intuition.
+      + induction l1; simpl; intuition. 
+  Qed.
+
+  Lemma Exists_not_Forall l : Exists (not ∘ P) l → ¬Forall P l.
+  Proof. induction 1; inversion_clear 1; contradiction. Qed.
+  Lemma Forall_not_Exists l : Forall (not ∘ P) l → ¬Exists P l.
+  Proof. induction 1; inversion_clear 1; contradiction. Qed.
+
+  Context {dec : ∀ x, Decision (P x)}.
+
+  Fixpoint Forall_Exists_dec l : {Forall P l} + {Exists (not ∘ P) l}.
+  Proof.
+   refine (
+    match l with
+    | [] => left _
+    | x :: l => cast_if_and (dec x) (Forall_Exists_dec l)
+    end); clear Forall_Exists_dec; abstract intuition.
+  Defined.
+
+  Lemma not_Forall_Exists l : ¬Forall P l → Exists (not ∘ P) l.
+  Proof. intro. destruct (Forall_Exists_dec l); intuition. Qed.
+
+  Global Instance Forall_dec l : Decision (Forall P l) :=
+    match Forall_Exists_dec l with
+    | left H => left H
+    | right H => right (Exists_not_Forall _ H)
+    end.
+
+  Fixpoint Exists_Forall_dec l : {Exists P l} + {Forall (not ∘ P) l}.
+  Proof.
+   refine (
+    match l with
+    | [] => right _
+    | x :: l => cast_if_or (dec x) (Exists_Forall_dec l)
+    end); clear Exists_Forall_dec; abstract intuition.
+  Defined.
+
+  Lemma not_Exists_Forall l : ¬Exists P l → Forall (not ∘ P) l.
+  Proof. intro. destruct (Exists_Forall_dec l); intuition. Qed.
+
+  Global Instance Exists_dec l : Decision (Exists P l) :=
+    match Exists_Forall_dec l with
+    | left H => left H
+    | right H => right (Forall_not_Exists _ H)
+    end.
+End Forall_Exists.
+End general_properties.
+
+Section Forall2.
+  Context {A B} (P : A → B → Prop).
+
+  Lemma Forall2_nil_inv_l k :
+    Forall2 P [] k → k = [].
+  Proof. by inversion 1. Qed.
+  Lemma Forall2_nil_inv_r k :
+    Forall2 P k [] → k = [].
+  Proof. by inversion 1. Qed.
+
+  Lemma Forall2_cons_inv l1 l2 x1 x2 :
+    Forall2 P (x1 :: l1) (x2 :: l2) → P x1 x2 ∧ Forall2 P l1 l2.
+  Proof. by inversion 1. Qed.
+  Lemma Forall2_cons_inv_l l1 k x1 :
+    Forall2 P (x1 :: l1) k → ∃ x2 l2,
+      P x1 x2 ∧ Forall2 P l1 l2 ∧ k = x2 :: l2.
+  Proof. inversion 1; subst; eauto. Qed.
+  Lemma Forall2_cons_inv_r k l2 x2 :
+    Forall2 P k (x2 :: l2) → ∃ x1 l1,
+      P x1 x2 ∧ Forall2 P l1 l2 ∧ k = x1 :: l1.
+  Proof. inversion 1; subst; eauto. Qed.
+  Lemma Forall2_cons_nil_inv l1 x1 :
+    Forall2 P (x1 :: l1) [] → False.
+  Proof. by inversion 1. Qed.
+  Lemma Forall2_nil_cons_inv l2 x2 :
+    Forall2 P [] (x2 :: l2) → False.
+  Proof. by inversion 1. Qed.
+
+  Lemma Forall2_flip l1 l2 :
+    Forall2 P l1 l2 ↔ Forall2 (flip P) l2 l1.
+  Proof. split; induction 1; constructor; auto. Qed.
+
+  Lemma Forall2_length l1 l2 :
+    Forall2 P l1 l2 → length l1 = length l2.
+  Proof. induction 1; simpl; auto. Qed.
+
+  Lemma Forall2_impl (Q : A → B → Prop) l1 l2 :
+    Forall2 P l1 l2 → (∀ x y, P x y → Q x y) → Forall2 Q l1 l2.
+  Proof. intros H ?. induction H; auto. Defined.
+
+  Lemma Forall2_unique l k1 k2 :
+    Forall2 P l k1 →
+    Forall2 P l k2 →
+    (∀ x y1 y2, P x y1 → P x y2 → y1 = y2) →
+    k1 = k2.
+  Proof.
+    intros H. revert k2.
+    induction H; inversion_clear 1; intros; f_equal; eauto.
+  Qed.
+
+  Lemma Forall2_Forall_l (Q : A → Prop) l k :
+    Forall2 P l k →
+    Forall (λ y, ∀ x, P x y → Q x) k →
+    Forall Q l.
+  Proof. induction 1; inversion_clear 1; eauto. Qed.
+  Lemma Forall2_Forall_r (Q : B → Prop) l k :
+    Forall2 P l k →
+    Forall (λ x, ∀ y, P x y → Q y) l →
+    Forall Q k.
+  Proof. induction 1; inversion_clear 1; eauto. Qed.
+
+  Lemma Forall2_lookup l1 l2 i x y :
+    Forall2 P l1 l2 →
+      l1 !! i = Some x → l2 !! i = Some y → P x y.
+  Proof.
+    intros H. revert i. induction H.
+    * discriminate.
+    * intros [|?] ??; simpl in *; simplify_equality; eauto.
+  Qed.
+  Lemma Forall2_lookup_l l1 l2 i x :
+    Forall2 P l1 l2 → l1 !! i = Some x → ∃ y,
+      l2 !! i = Some y ∧ P x y.
+  Proof.
+    intros H. revert i. induction H.
+    * discriminate.
+    * intros [|?] ?; simpl in *; simplify_equality; eauto.
+  Qed.
+  Lemma Forall2_lookup_r l1 l2 i y :
+    Forall2 P l1 l2 → l2 !! i = Some y → ∃ x,
+      l1 !! i = Some x ∧ P x y.
+  Proof.
+    intros H. revert i. induction H.
+    * discriminate.
+    * intros [|?] ?; simpl in *; simplify_equality; eauto.
+  Qed.
+
+  Lemma Forall2_alter_l f l1 l2 i :
+    Forall2 P l1 l2 →
+    (∀ x1 x2,
+      l1 !! i = Some x1 → l2 !! i = Some x2 → P x1 x2 → P (f x1) x2) →
+    Forall2 P (alter f i l1) l2.
+  Proof.
+    intros Hl. revert i.
+    induction Hl; simpl; intros [|i]; constructor; auto.
+  Qed.
+  Lemma Forall2_alter_r f l1 l2 i :
+    Forall2 P l1 l2 →
+    (∀ x1 x2,
+      l1 !! i = Some x1 → l2 !! i = Some x2 → P x1 x2 → P x1 (f x2)) →
+    Forall2 P l1 (alter f i l2).
+  Proof.
+    intros Hl. revert i.
+    induction Hl; simpl; intros [|i]; constructor; auto.
+  Qed.
+  Lemma Forall2_alter f g l1 l2 i :
+    Forall2 P l1 l2 →
+    (∀ x1 x2,
+      l1 !! i = Some x1 → l2 !! i = Some x2 → P x1 x2 → P (f x1) (g x2)) →
+    Forall2 P (alter f i l1) (alter g i l2).
+  Proof.
+    intros Hl. revert i.
+    induction Hl; simpl; intros [|i]; constructor; auto.
+  Qed.
+
+  Lemma Forall2_replicate_l l n x :
+    Forall (P x) l →
+    length l = n →
+    Forall2 P (replicate n x) l.
+  Proof.
+    intros Hl. revert n.
+    induction Hl; intros [|?] ?; simplify_equality; constructor; auto.
+  Qed.
+  Lemma Forall2_replicate_r l n x :
+    Forall (flip P x) l →
+    length l = n →
+    Forall2 P l (replicate n x).
+  Proof.
+    intros Hl. revert n.
+    induction Hl; intros [|?] ?; simplify_equality; constructor; auto.
+  Qed.
+  Lemma Forall2_replicate n x1 x2 :
+    P x1 x2 →
+    Forall2 P (replicate n x1) (replicate n x2).
+  Proof. induction n; simpl; constructor; auto. Qed.
+
+  Lemma Forall2_take l1 l2 n :
+    Forall2 P l1 l2 →
+    Forall2 P (take n l1) (take n l2).
+  Proof.
+    intros Hl1l2. revert n.
+    induction Hl1l2; intros [|?]; simpl; auto.
+  Qed.
+  Lemma Forall2_drop l1 l2 n :
+    Forall2 P l1 l2 →
+    Forall2 P (drop n l1) (drop n l2).
+  Proof.
+    intros Hl1l2. revert n.
+    induction Hl1l2; intros [|?]; simpl; auto.
+  Qed.
+  Lemma Forall2_resize l1 l2 x1 x2 n :
+    P x1 x2 →
+    Forall2 P l1 l2 →
+    Forall2 P (resize n x1 l1) (resize n x2 l2).
+  Proof.
+    intros. rewrite !resize_spec, (Forall2_length l1 l2) by done.
+    auto using Forall2_app, Forall2_take, Forall2_replicate.
+  Qed.
+
+  Lemma Forall2_resize_ge_l l1 l2 x1 x2 n m :
+    (∀ x, P x x2) →
+    n ≤ m →
+    Forall2 P (resize n x1 l1) l2 →
+    Forall2 P (resize m x1 l1) (resize m x2 l2).
+  Proof.
+    intros. assert (n = length l2).
+    { by rewrite <-(Forall2_length (resize n x1 l1) l2), resize_length. }
+    rewrite (le_plus_minus n m) by done. subst.
+    rewrite !resize_plus, resize_all, drop_all, resize_nil.
+    apply Forall2_app; [done |].
+    apply Forall2_replicate_r; [| by rewrite resize_length].
+    by apply Forall_true.
+  Qed.
+  Lemma Forall2_resize_ge_r l1 l2 x1 x2 n m :
+    (∀ x3, P x1 x3) →
+    n ≤ m →
+    Forall2 P l1 (resize n x2 l2) →
+    Forall2 P (resize m x1 l1) (resize m x2 l2).
+  Proof.
+    intros. assert (n = length l1).
+    { by rewrite (Forall2_length l1 (resize n x2 l2)), resize_length. }
+    rewrite (le_plus_minus n m) by done. subst.
+    rewrite !resize_plus, resize_all, drop_all, resize_nil.
+    apply Forall2_app; [done |].
+    apply Forall2_replicate_l; [| by rewrite resize_length].
+    by apply Forall_true.
+  Qed.
+
+  Lemma Forall2_trans {C} (Q : B → C → Prop) (R : A → C → Prop) l1 l2 l3 :
+    (∀ x1 x2 x3, P x1 x2 → Q x2 x3 → R x1 x3) →
+    Forall2 P l1 l2 →
+    Forall2 Q l2 l3 →
+    Forall2 R l1 l3.
+  Proof.
+    intros ? Hl1l2. revert l3.
+    induction Hl1l2; inversion_clear 1; eauto.
+  Qed.
+
+  Lemma Forall2_Forall (Q : A → A → Prop) l :
+    Forall (λ x, Q x x) l → Forall2 Q l l.
+  Proof. induction 1; constructor; auto. Qed.
+
+  Global Instance Forall2_dec `{∀ x1 x2, Decision (P x1 x2)} :
+    ∀ l1 l2, Decision (Forall2 P l1 l2).
+  Proof.
+   refine (
+    fix go l1 l2 : Decision (Forall2 P l1 l2) :=
+    match l1, l2 with
+    | [], [] => left _
+    | x1 :: l1, x2 :: l2 => cast_if_and (decide (P x1 x2)) (go l1 l2)
+    | _, _ => right _
+    end); clear go; abstract first [by constructor | by inversion 1].
+  Defined.
+End Forall2.
+
+Section Forall2_order.
+  Context  {A} (R : relation A).
+
+  Global Instance: PreOrder R → PreOrder (Forall2 R).
+  Proof.
+    split.
+    * intros l. induction l; by constructor.
+    * intros ???. apply Forall2_trans. apply transitivity.
+  Qed.
+  Global Instance: AntiSymmetric R → AntiSymmetric (Forall2 R).
+  Proof. induction 2; inversion_clear 1; f_equal; auto. Qed.
+End Forall2_order.
+
+Ltac decompose_elem_of_list := repeat
+  match goal with
+  | H : ?x ∈ [] |- _ => by destruct (not_elem_of_nil x)
+  | H : _ ∈ _ :: _ |- _ => apply elem_of_cons in H; destruct H
+  | H : _ ∈ _ ++ _ |- _ => apply elem_of_app in H; destruct H
+  end.
+Ltac decompose_Forall := repeat
+  match goal with
+  | H : Forall _ [] |- _ => clear H
+  | H : Forall _ (_ :: _) |- _ => rewrite Forall_cons in H; destruct H
+  | H : Forall _ (_ ++ _) |- _ => rewrite Forall_app in H; destruct H
+  | H : Forall2 _ [] [] |- _ => clear H
+  | H : Forall2 _ (_ :: _) [] |- _ => destruct (Forall2_cons_nil_inv _ _ _ H)
+  | H : Forall2 _ [] (_ :: _) |- _ => destruct (Forall2_nil_cons_inv _ _ _ H)
+  | H : Forall2 _ [] ?l |- _ => apply Forall2_nil_inv_l in H; subst l
+  | H : Forall2 _ ?l [] |- _ => apply Forall2_nil_inv_r in H; subst l
+  | H : Forall2 _ (_ :: _) (_ :: _) |- _ =>
+     apply Forall2_cons_inv in H; destruct H
+  | H : Forall2 _ (_ :: _) ?l |- _ =>
+     apply Forall2_cons_inv_l in H; destruct H as (? & ? & ? & ? & ?); subst l
+  | H : Forall2 _ ?l (_ :: _) |- _ =>
+     apply Forall2_cons_inv_r in H; destruct H as (? & ? & ? & ? & ?); subst l
+  end.
+
+(** * Theorems on the prefix and suffix predicates *)
+Section prefix_postfix.
+  Context {A : Type}.
+
+  Global Instance: PreOrder (@prefix_of A).
+  Proof.
+    split.
+    * intros ?. eexists []. by rewrite app_nil_r.
+    * intros ??? [k1 ?] [k2 ?].
+      exists (k1 ++ k2). subst. by rewrite app_assoc.
+  Qed.
+
+  Lemma prefix_of_nil (l : list A) : prefix_of [] l.
+  Proof. by exists l. Qed.
+  Lemma prefix_of_nil_not x (l : list A) : ¬prefix_of (x :: l) [].
+  Proof. by intros [k E]. Qed.
+  Lemma prefix_of_cons x y (l1 l2 : list A) :
+    x = y → prefix_of l1 l2 → prefix_of (x :: l1) (y :: l2).
+  Proof. intros ? [k E]. exists k. by subst. Qed.
+  Lemma prefix_of_cons_inv_1 x y (l1 l2 : list A) :
+    prefix_of (x :: l1) (y :: l2) → x = y.
+  Proof. intros [k E]. by injection E. Qed.
+  Lemma prefix_of_cons_inv_2 x y (l1 l2 : list A) :
+    prefix_of (x :: l1) (y :: l2) → prefix_of l1 l2.
+  Proof. intros [k E]. exists k. by injection E. Qed.
+
+  Lemma prefix_of_app_l (l1 l2 l3 : list A) :
+    prefix_of (l1 ++ l3) l2 → prefix_of l1 l2.
+  Proof. intros [k ?]. red. exists (l3 ++ k). subst. by rewrite <-app_assoc. Qed.
+  Lemma prefix_of_app_r (l1 l2 l3 : list A) :
+    prefix_of l1 l2 → prefix_of l1 (l2 ++ l3).
+  Proof. intros [k ?]. exists (k ++ l3). subst. by rewrite app_assoc. Qed.
+
+  Global Instance: PreOrder (@suffix_of A).
+  Proof.
+    split.
+    * intros ?. by eexists [].
+    * intros ??? [k1 ?] [k2 ?].
+      exists (k2 ++ k1). subst. by rewrite app_assoc.
+  Qed.
+
+  Lemma prefix_suffix_reverse (l1 l2 : list A) :
+    prefix_of l1 l2 ↔ suffix_of (reverse l1) (reverse l2).
+  Proof.
+    split; intros [k E]; exists (reverse k).
+    * by rewrite E, reverse_app.
+    * by rewrite <-(reverse_involutive l2), E, reverse_app, reverse_involutive.
+  Qed.
+  Lemma suffix_prefix_reverse (l1 l2 : list A) :
+    suffix_of l1 l2 ↔ prefix_of (reverse l1) (reverse l2).
+  Proof. by rewrite prefix_suffix_reverse, !reverse_involutive. Qed.
+
+  Lemma suffix_of_nil (l : list A) : suffix_of [] l.
+  Proof. exists l. by rewrite app_nil_r. Qed.
+  Lemma suffix_of_nil_inv (l : list A) : suffix_of l [] → l = [].
+  Proof. by intros [[|?] ?]; simplify_list_equality. Qed.
+  Lemma suffix_of_cons_nil_inv x (l : list A) : ¬suffix_of (x :: l) [].
+  Proof. by intros [[] ?]. Qed.
+
+  Lemma suffix_of_app (l1 l2 k : list A) :
+    suffix_of l1 l2 → suffix_of (l1 ++ k) (l2 ++ k).
+  Proof. intros [k' E]. exists k'. subst. by rewrite app_assoc. Qed.
+
+  Lemma suffix_of_snoc_inv_1 x y (l1 l2 : list A) :
+    suffix_of (l1 ++ [x]) (l2 ++ [y]) → x = y.
+  Proof.
+    rewrite suffix_prefix_reverse, !reverse_snoc.
+    by apply prefix_of_cons_inv_1.
+  Qed.
+  Lemma suffix_of_snoc_inv_2 x y (l1 l2 : list A) :
+    suffix_of (l1 ++ [x]) (l2 ++ [y]) → suffix_of l1 l2.
+  Proof.
+    rewrite !suffix_prefix_reverse, !reverse_snoc.
+    by apply prefix_of_cons_inv_2.
+  Qed.
+
+  Lemma suffix_of_cons_l (l1 l2 : list A) x :
+    suffix_of (x :: l1) l2 → suffix_of l1 l2.
+  Proof. intros [k ?]. exists (k ++ [x]). subst. by rewrite <-app_assoc. Qed.
+  Lemma suffix_of_app_l (l1 l2 l3 : list A) :
+
+  suffix_of (l3 ++ l1) l2 → suffix_of l1 l2.
+  Proof. intros [k ?]. exists (k ++ l3). subst. by rewrite <-app_assoc. Qed.
+  Lemma suffix_of_cons_r (l1 l2 : list A) x :
+    suffix_of l1 l2 → suffix_of l1 (x :: l2).
+  Proof. intros [k ?]. exists (x :: k). by subst. Qed.
+  Lemma suffix_of_app_r (l1 l2 l3 : list A) :
+    suffix_of l1 l2 → suffix_of l1 (l3 ++ l2).
+  Proof. intros [k ?]. exists (l3 ++ k). subst. by rewrite app_assoc. Qed.
+
+  Lemma suffix_of_cons_inv (l1 l2 : list A) x y :
+    suffix_of (x :: l1) (y :: l2) →
+      x :: l1 = y :: l2 ∨ suffix_of (x :: l1) l2.
+  Proof.
+    intros [[|? k] E].
+    * by left.
+    * right. simplify_equality. by apply suffix_of_app_r.
+  Qed.
+
+  Lemma suffix_of_cons_not x (l : list A) : ¬suffix_of (x :: l) l.
+  Proof. intros [??]. discriminate_list_equality. Qed.
+
+  Context `{∀ x y : A, Decision (x = y)}.
+
+  Fixpoint strip_prefix (l1 l2 : list A) : list A * (list A * list A) :=
+    match l1, l2 with
+    | [], l2 => ([], ([], l2))
+    | l1, [] => ([], (l1, []))
+    | x :: l1, y :: l2 =>
+      if decide_rel (=) x y
+      then fst_map (x ::) (strip_prefix l1 l2)
+      else ([], (x :: l1, y :: l2))
+    end.
+
+  Global Instance prefix_of_dec: ∀ l1 l2 : list A,
+      Decision (prefix_of l1 l2) :=
+    fix go l1 l2 :=
+    match l1, l2 return { prefix_of l1 l2 } + { ¬prefix_of l1 l2 } with
+    | [], _ => left (prefix_of_nil _)
+    | _, [] => right (prefix_of_nil_not _ _)
+    | x :: l1, y :: l2 =>
+      match decide_rel (=) x y with
+      | left Exy =>
+        match go l1 l2 with
+        | left Hl1l2 => left (prefix_of_cons _ _ _ _ Exy Hl1l2)
+        | right Hl1l2 => right (Hl1l2 ∘ prefix_of_cons_inv_2 _ _ _ _)
+        end
+      | right Exy => right (Exy ∘ prefix_of_cons_inv_1 _ _ _ _)
+      end
+    end.
+
+  Global Instance suffix_of_dec (l1 l2 : list A) :
+    Decision (suffix_of l1 l2).
+  Proof.
+    refine (cast_if (decide_rel prefix_of (reverse l1) (reverse l2)));
+     abstract (by rewrite suffix_prefix_reverse).
+  Defined.
 End prefix_postfix.
 
-(** The [simplify_suffix_of] removes [suffix_of] assumptions that are
-tautologies, and simplifies [suffix_of] assumptions involving [(::)] and
+(** The [simplify_suffix_of] tactic removes [suffix_of] hypotheses that are
+tautologies, and simplifies [suffix_of] hypotheses involving [(::)] and
 [(++)]. *)
 Ltac simplify_suffix_of := repeat
   match goal with
   | H : suffix_of (_ :: _) _ |- _ =>
     destruct (suffix_of_cons_not _ _ H)
   | H : suffix_of (_ :: _) [] |- _ =>
-    destruct (suffix_of_nil_not _ _ H)
+    apply suffix_of_nil_inv in H
   | H : suffix_of (_ :: _) (_ :: _) |- _ =>
     destruct (suffix_of_cons_inv _ _ _ _ H); clear H
   | H : suffix_of ?x ?x |- _ => clear H
@@ -372,81 +1354,376 @@ Ltac simplify_suffix_of := repeat
   | _ => progress simplify_equality
   end.
 
-(** The [solve_suffix_of] tries to solve goals involving [suffix_of]. It uses
-[simplify_suffix_of] to simplify assumptions, tries to solve [suffix_of]
-conclusions, and adds transitive consequences of assumptions to the context. 
-This tactic either fails or proves the goal. *)
-Ltac solve_suffix_of :=
-  let rec go :=
+(** The [solve_suffix_of] tactic tries to solve goals involving [suffix_of]. It
+uses [simplify_suffix_of] to simplify hypotheses and tries to solve [suffix_of]
+conclusions. This tactic either fails or proves the goal. *)
+Ltac solve_suffix_of := solve [intuition (repeat
   match goal with
-  | _ => progress simplify_suffix_of; go
+  | _ => done
+  | _ => progress simplify_suffix_of
   | |- suffix_of [] _ => apply suffix_of_nil
   | |- suffix_of _ _ => reflexivity
-  | |- suffix_of _ _ => solve [auto]
-  | |- suffix_of _ (_ :: _) => apply suffix_of_cons_r; go
-  | |- suffix_of _ (_ ++ _) => apply suffix_of_app_r; go
-  | H : ¬suffix_of _ _ |- _ => destruct H; go
-  | H1 : suffix_of ?x ?y, H2 : suffix_of ?y ?z |- _ =>
-     match goal with
-     | _ : suffix_of x z |- _ => fail 1
-     | _ => assert (suffix_of x z) by (transitivity y; assumption);
-            clear H1; clear H2; go (**i clear to avoid loops *)
-     end
-  end in go.
+  | |- suffix_of _ (_ :: _) => apply suffix_of_cons_r
+  | |- suffix_of _ (_ ++ _) => apply suffix_of_app_r
+  | H : suffix_of _ _ → False |- _ => destruct H
+  end)].
 Hint Extern 0 (PropHolds (suffix_of _ _)) =>
   unfold PropHolds; solve_suffix_of : typeclass_instances.
 
+(** * Folding lists *)
+Notation foldr := fold_right.
+Notation foldr_app := fold_right_app.
+
+Lemma foldr_permutation {A B} (R : relation B)
+   `{!Equivalence R}
+   (f : A → B → B) (b : B)
+   `{!Proper ((=) ==> R ==> R) f}
+   (Hf : ∀ a1 a2 b, R (f a1 (f a2 b)) (f a2 (f a1 b))) :
+  Proper (Permutation ==> R) (foldr f b).
+Proof.
+  induction 1; simpl.
+  * done.
+  * by f_equiv.
+  * apply Hf.
+  * etransitivity; eauto.
+Qed.
+
+(** We redefine [foldl] with the arguments in the same order as in Haskell. *)
+Definition foldl {A B} (f : A → B → A) : A → list B → A :=
+  fix go a l :=
+  match l with
+  | [] => a
+  | x :: l => go (f a x) l
+  end.
+
+Lemma foldl_app {A B} (f : A → B → A) (l k : list B) (a : A) :
+  foldl f a (l ++ k) = foldl f (foldl f a l) k.
+Proof. revert a. induction l; simpl; auto. Qed.
+
 (** * Monadic operations *)
-Global Instance list_ret: MRet list := λ A a, [a].
-Global Instance list_fmap: FMap list :=
-  fix go A B (f : A → B) (l : list A) :=
+Instance list_ret: MRet list := λ A x, x :: @nil A.
+Instance list_fmap {A B} (f : A → B) : FMapD list f :=
+  fix go (l : list A) :=
   match l with
   | [] => []
-  | x :: l => f x :: @fmap _ go _ _ f l
+  | x :: l => f x :: @fmap _ _ _ f go l
+  end.
+Instance list_bind {A B} (f : A → list B) : MBindD list f :=
+  fix go (l : list A) :=
+  match l with
+  | [] => []
+  | x :: l => f x ++ @mbind _ _ _ f go l
+  end.
+Instance list_join: MJoin list :=
+  fix go A (ls : list (list A)) : list A :=
+  match ls with
+  | [] => []
+  | l :: ls => l ++ @mjoin _ go _ ls
   end.
-Global Instance list_join: MJoin list :=
-  fix go A (l : list (list A)) : list A :=
+
+Definition mapM `{!MBind M} `{!MRet M} {A B}
+    (f : A → M B) : list A → M (list B) :=
+  fix go l :=
   match l with
-  | [] =>  []
-  | x :: l => x ++ @mjoin _ go _ l
+  | [] => mret []
+  | x :: l => y ← f x; k ← go l; mret (y :: k)
   end.
-Global Instance list_bind: MBind list := λ A B f l, mjoin (f <$> l).
 
 Section list_fmap.
   Context {A B : Type} (f : A → B).
 
-  Local Arguments fmap _ _ _ _ _ !_ /.
+  Lemma list_fmap_compose {C} (g : B → C) l :
+    g ∘ f <$> l = g <$> f <$> l.
+  Proof. induction l; simpl; f_equal; auto. Qed.
+
+  Lemma list_fmap_ext (g : A → B) (l : list A) :
+    (∀ x, f x = g x) → fmap f l = fmap g l.
+  Proof. intro. induction l; simpl; f_equal; auto. Qed.
+  Lemma list_fmap_ext_alt (g : A → B) (l : list A) :
+    Forall (λ x, f x = g x) l ↔ fmap f l = fmap g l.
+  Proof.
+    split.
+    * induction 1; simpl; f_equal; auto.
+    * induction l; simpl; constructor; simplify_equality; auto.
+  Qed.
+
+  Global Instance: Injective (=) (=) f → Injective (=) (=) (fmap f).
+  Proof.
+    intros ? l1. induction l1 as [|x l1 IH].
+    * by intros [|??].
+    * intros [|??]; simpl; intros; f_equal; simplify_equality; auto.
+  Qed.
+  Lemma fmap_app l1 l2 : f <$> l1 ++ l2 = (f <$> l1) ++ (f <$> l2).
+  Proof. induction l1; simpl; by f_equal. Qed.
+  Lemma fmap_cons_inv y l k :
+    f <$> l = y :: k → ∃ x l', y = f x ∧ l = x :: l'.
+  Proof. intros. destruct l; simpl; simplify_equality; eauto. Qed.
+  Lemma fmap_app_inv l k1 k2 :
+    f <$> l = k1 ++ k2 → ∃ l1 l2,
+      k1 = f <$> l1 ∧ k2 = f <$> l2 ∧ l = l1 ++ l2.
+  Proof.
+    revert l. induction k1 as [|y k1 IH]; simpl.
+    * intros l ?. by eexists [], l.
+    * intros [|x l] ?; simpl; simplify_equality.
+      destruct (IH l) as [l1 [l2 [? [??]]]]; subst; [done |].
+      by exists (x :: l1) l2.
+  Qed.
 
   Lemma fmap_length l : length (f <$> l) = length l.
-  Proof. induction l; simpl; auto. Qed.
+  Proof. induction l; simpl; by f_equal. Qed.
+  Lemma fmap_reverse l : f <$> reverse l = reverse (f <$> l).
+  Proof.
+    induction l; simpl; [done |].
+    by rewrite !reverse_cons, fmap_app, IHl.
+  Qed.
+  Lemma fmap_replicate n x :
+    f <$> replicate n x = replicate n (f x).
+  Proof. induction n; simpl; f_equal; auto. Qed.
 
   Lemma list_lookup_fmap l i : (f <$> l) !! i = f <$> (l !! i).
-  Proof. revert i. induction l; now intros [|]. Qed.
+  Proof. revert i. induction l; by intros [|]. Qed.
+  Lemma list_lookup_fmap_inv l i x :
+    (f <$> l) !! i = Some x → ∃ y, x = f y ∧ l !! i = Some y.
+  Proof.
+    intros Hi. rewrite list_lookup_fmap in Hi.
+    destruct (l !! i) eqn:?; simplify_equality; eauto.
+  Qed.
+   
+  Lemma list_alter_fmap (g : A → A) (h : B → B) l i :
+    Forall (λ x, f (g x) = h (f x)) l →
+    f <$> alter g i l = alter h i (f <$> l).
+  Proof.
+    intros Hl. revert i.
+    induction Hl; intros [|i]; simpl; f_equal; auto.
+  Qed.
+  Lemma elem_of_list_fmap_1 l x : x ∈ l → f x ∈ f <$> l.
+  Proof. induction 1; simpl; rewrite elem_of_cons; intuition. Qed.
+  Lemma elem_of_list_fmap_1_alt l x y : x ∈ l → y = f x → y ∈ f <$> l.
+  Proof. intros. subst. by apply elem_of_list_fmap_1. Qed.
+  Lemma elem_of_list_fmap_2 l x : x ∈ f <$> l → ∃ y, x = f y ∧ y ∈ l.
+  Proof.
+    induction l as [|y l IH]; simpl; intros; decompose_elem_of_list.
+    + exists y. split; [done | by left].
+    + destruct IH as [z [??]]. done. exists z. split; [done | by right].
+  Qed.
+  Lemma elem_of_list_fmap l x : x ∈ f <$> l ↔ ∃ y, x = f y ∧ y ∈  l.
+  Proof.
+    firstorder eauto using elem_of_list_fmap_1_alt, elem_of_list_fmap_2.
+  Qed.
 
-  Lemma in_fmap_1 l x : In x l → In (f x) (f <$> l).
-  Proof. induction l; destruct 1; subst; constructor (solve [auto]). Qed.
-  Lemma in_fmap_1_alt l x y : In x l → y = f x → In y (f <$> l).
-  Proof. intros. subst. now apply in_fmap_1. Qed.
-  Lemma in_fmap_2 l x : In x (f <$> l) → ∃ y, x = f y ∧ In y l.
+  Lemma NoDup_fmap_1 (l : list A) :
+    NoDup (f <$> l) → NoDup l.
   Proof.
-    induction l as [|y l]; destruct 1; subst.
-    * exists y; now intuition.
-    * destruct IHl as [y' [??]]. easy. exists y'; now intuition.
+    induction l; simpl; inversion_clear 1; constructor; auto.
+    rewrite elem_of_list_fmap in *. naive_solver.
   Qed.
-  Lemma in_fmap l x : In x (f <$> l) ↔ ∃ y, x = f y ∧ In y l.
+  Lemma NoDup_fmap_2 `{!Injective (=) (=) f} (l : list A) :
+    NoDup l → NoDup (f <$> l).
   Proof.
-    split.
-    * apply in_fmap_2.
-    * intros [? [??]]. eauto using in_fmap_1_alt.
+    induction 1; simpl; constructor; trivial.
+    rewrite elem_of_list_fmap. intros [y [Hxy ?]].
+    apply (injective f) in Hxy. by subst.
+  Qed.
+  Lemma NoDup_fmap `{!Injective (=) (=) f} (l : list A) :
+    NoDup (f <$> l) ↔ NoDup l.
+  Proof. split; auto using NoDup_fmap_1, NoDup_fmap_2. Qed.
+
+  Global Instance fmap_Permutation_proper:
+    Proper (Permutation ==> Permutation) (fmap f).
+  Proof. induction 1; simpl; econstructor; eauto. Qed.
+
+  Lemma Forall_fmap (l : list A) (P : B → Prop) :
+    Forall P (f <$> l) ↔ Forall (P ∘ f) l.
+  Proof.
+    split; induction l; inversion_clear 1; constructor; auto.
+  Qed.
+
+  Lemma Forall2_fmap_l {C} (P : B → C → Prop) l1 l2 :
+    Forall2 P (f <$> l1) l2 ↔ Forall2 (P ∘ f) l1 l2.
+  Proof.
+    split; revert l2; induction l1; inversion_clear 1; constructor; auto.
+  Qed.
+  Lemma Forall2_fmap_r {C} (P : C → B → Prop) l1 l2 :
+    Forall2 P l1 (f <$> l2) ↔ Forall2 (λ x, P x ∘ f) l1 l2.
+  Proof.
+    split; revert l1; induction l2; inversion_clear 1; constructor; auto.
+  Qed.
+  Lemma Forall2_fmap_1 {C D} (g : C → D) (P : B → D → Prop) l1 l2 :
+    Forall2 P (f <$> l1) (g <$> l2) →
+    Forall2 (λ x1 x2, P (f x1) (g x2)) l1 l2.
+  Proof. revert l2; induction l1; intros [|??]; inversion_clear 1; auto. Qed.
+  Lemma Forall2_fmap_2 {C D} (g : C → D) (P : B → D → Prop) l1 l2 :
+    Forall2 (λ x1 x2, P (f x1) (g x2)) l1 l2 →
+    Forall2 P (f <$> l1) (g <$> l2).
+  Proof. induction 1; simpl; auto. Qed.
+  Lemma Forall2_fmap {C D} (g : C → D) (P : B → D → Prop) l1 l2 :
+    Forall2 P (f <$> l1) (g <$> l2) ↔
+    Forall2 (λ x1 x2, P (f x1) (g x2)) l1 l2.
+  Proof. split; auto using Forall2_fmap_1, Forall2_fmap_2. Qed.
+
+  Lemma mapM_fmap_Some (g : B → option A) (l : list A) :
+    (∀ x, g (f x) = Some x) →
+    mapM g (f <$> l) = Some l.
+  Proof. intros. by induction l; simpl; simplify_option_equality. Qed.
+  Lemma mapM_fmap_Some_inv (g : B → option A) (l : list A) (k : list B) :
+    (∀ x y, g y = Some x → y = f x) →
+    mapM g k = Some l →
+    k = f <$> l.
+  Proof.
+    intros Hgf. revert l; induction k as [|y k]; intros [|x l] ?;
+      simplify_option_equality; f_equiv; eauto.
+  Qed.
+
+  Lemma mapM_Some (g : B → option A) l k :
+    Forall2 (λ x y, g x = Some y) l k →
+    mapM g l = Some k.
+  Proof. by induction 1; simplify_option_equality. Qed.
+
+  Lemma Forall2_impl_mapM (P : B → A → Prop) (g : B → option A) l k :
+    Forall (λ x, ∀ y, g x = Some y → P x y) l →
+    mapM g l = Some k →
+    Forall2 P l k.
+  Proof.
+    intros Hl. revert k. induction Hl; intros [|??] ?;
+      simplify_option_equality; eauto.
   Qed.
 End list_fmap.
 
-Lemma Forall_fst {A B} (l : list (A * B)) (P : A → Prop) :
-  Forall (P ∘ fst) l ↔ Forall P (fst <$> l).
-Proof. induction l; split; inversion 1; subst; constructor; firstorder auto. Qed.
-Lemma Forall_snd {A B} (l : list (A * B)) (P : B → Prop) :
-  Forall (P ∘ snd) l ↔ Forall P (snd <$> l).
-Proof. induction l; split; inversion 1; subst; constructor; firstorder auto. Qed.
+Lemma NoDup_fmap_fst {A B} (l : list (A * B)) :
+  (∀ x y1 y2, (x,y1) ∈ l → (x,y2) ∈ l → y1 = y2) →
+  NoDup l →
+  NoDup (fst <$> l).
+Proof.
+  intros Hunique.
+  induction 1 as [|[x1 y1] l Hin Hnodup IH]; simpl; constructor.
+  * rewrite elem_of_list_fmap.
+    intros [[x2 y2] [??]]; simpl in *; subst. destruct Hin.
+    rewrite (Hunique x2 y1 y2); rewrite ?elem_of_cons; auto.
+  * apply IH. intros.
+    eapply Hunique; rewrite ?elem_of_cons; eauto.
+Qed.
+
+Section list_bind.
+  Context {A B : Type} (f : A → list B).
+
+  Lemma bind_app (l1 l2 : list A) :
+    (l1 ++ l2) ≫= f = (l1 ≫= f) ++ (l2 ≫= f).
+  Proof.
+    induction l1; simpl; [done|].
+    by rewrite <-app_assoc, IHl1.
+  Qed.
+  Lemma elem_of_list_bind (x : B) (l : list A) :
+    x ∈ l ≫= f ↔ ∃ y, x ∈ f y ∧ y ∈ l.
+  Proof.
+    split.
+    * induction l as [|y l IH]; simpl; intros; decompose_elem_of_list.
+      + exists y. split; [done | by left].
+      + destruct IH as [z [??]]. done.
+        exists z. split; [done | by right].
+    * intros [y [Hx Hy]].
+      induction Hy; simpl; rewrite elem_of_app; intuition.
+  Qed.
+
+  Lemma Forall2_bind {C D} (g : C → list D) (P : B → D → Prop) l1 l2 :
+    Forall2 (λ x1 x2, Forall2 P (f x1) (g x2)) l1 l2 →
+    Forall2 P (l1 ≫= f) (l2 ≫= g).
+  Proof. induction 1; simpl; auto using Forall2_app. Qed.
+End list_bind.
+
+Section list_ret_join.
+  Context {A : Type}.
+
+  Lemma list_join_bind (ls : list (list A)) :
+    mjoin ls = ls ≫= id.
+  Proof. induction ls; simpl; f_equal; auto. Qed.
+
+  Lemma elem_of_list_ret (x y : A) :
+    x ∈ @mret list _ A y ↔ x = y.
+  Proof. apply elem_of_list_singleton. Qed.
+  Lemma elem_of_list_join (x : A) (ls : list (list A)) :
+    x ∈ mjoin ls ↔ ∃ l, x ∈ l ∧ l ∈ ls.
+  Proof. by rewrite list_join_bind, elem_of_list_bind. Qed.
+
+  Lemma join_nil (ls : list (list A)) :
+    mjoin ls = [] ↔ Forall (= []) ls.
+  Proof.
+    split.
+    * by induction ls as [|[|??] ?]; constructor; auto.
+    * by induction 1 as [|[|??] ?].
+  Qed.
+  Lemma join_nil_1 (ls : list (list A)) :
+    mjoin ls = [] → Forall (= []) ls.
+  Proof. by rewrite join_nil. Qed.
+  Lemma join_nil_2 (ls : list (list A)) :
+    Forall (= []) ls → mjoin ls = [].
+  Proof. by rewrite join_nil. Qed.
+
+  Lemma join_length (ls : list (list A)) :
+    length (mjoin ls) = foldr (plus ∘ length) 0 ls.
+  Proof. by induction ls; simpl; rewrite ?app_length; f_equal. Qed.
+  Lemma join_length_same (ls : list (list A)) n :
+    Forall (λ l, length l = n) ls →
+    length (mjoin ls) = length ls * n.
+  Proof. rewrite join_length. by induction 1; simpl; f_equal. Qed.
+
+  Lemma lookup_join_same_length (ls : list (list A)) n i :
+    n ≠ 0 →
+    Forall (λ l, length l = n) ls →
+    mjoin ls !! i = ls !! (i `div` n) ≫= (!! (i `mod` n)).
+  Proof.
+    intros Hn Hls. revert i.
+    induction Hls as [|l ls ? Hls IH]; simpl; [done |]. intros i.
+    destruct (decide (i < n)) as [Hin|Hin].
+    * rewrite <-(NPeano.Nat.div_unique i n 0 i) by lia.
+      rewrite <-(NPeano.Nat.mod_unique i n 0 i) by lia.
+      simpl. rewrite lookup_app_l; auto with lia.
+    * replace i with ((i - n) + 1 * n) by lia.
+      rewrite NPeano.Nat.div_add, NPeano.Nat.mod_add by done.
+      replace (i - n + 1 * n) with i by lia.
+      rewrite (plus_comm _ 1), lookup_app_r_alt, IH by lia.
+      by subst.
+  Qed.
+
+  (* This should be provable using the previous lemma in a shorter way *)
+  Lemma alter_join_same_length f (ls : list (list A)) n i :
+    n ≠ 0 →
+    Forall (λ l, length l = n) ls →
+    alter f i (mjoin ls) = mjoin (alter (alter f (i `mod` n)) (i `div` n) ls).
+  Proof.
+    intros Hn Hls. revert i.
+    induction Hls as [|l ls ? Hls IH]; simpl; [done |]. intros i.
+    destruct (decide (i < n)) as [Hin|Hin].
+    * rewrite <-(NPeano.Nat.div_unique i n 0 i) by lia.
+      rewrite <-(NPeano.Nat.mod_unique i n 0 i) by lia.
+      simpl. rewrite alter_app_l; auto with lia.
+    * replace i with ((i - n) + 1 * n) by lia.
+      rewrite NPeano.Nat.div_add, NPeano.Nat.mod_add by done.
+      replace (i - n + 1 * n) with i by lia.
+      rewrite (plus_comm _ 1), alter_app_r_alt, IH by lia.
+      by subst.
+  Qed.
+  Lemma insert_join_same_length (ls : list (list A)) n i x :
+    n ≠ 0 →
+    Forall (λ l, length l = n) ls →
+    <[i:=x]>(mjoin ls) = mjoin (alter <[i `mod` n:=x]> (i `div` n) ls).
+  Proof. apply alter_join_same_length. Qed.
+
+  Lemma Forall2_join {B} (P : A → B → Prop) ls1 ls2 :
+    Forall2 (Forall2 P) ls1 ls2 →
+    Forall2 P (mjoin ls1) (mjoin ls2).
+  Proof. induction 1; simpl; auto using Forall2_app. Qed.
+End list_ret_join.
+
+Ltac simplify_list_fmap_equality := repeat
+  match goal with
+  | _ => progress simplify_equality
+  | H : _ <$> _ = _ :: _ |- _ =>
+    apply fmap_cons_inv in H; destruct H as [? [? [??]]]
+  | H : _ :: _ = _ <$> _ |- _ => symmetry in H
+  | H : _ <$> _ = _ ++ _ |- _ =>
+    apply fmap_app_inv in H; destruct H as [? [? [? [??]]]]
+  | H : _ ++ _ = _ <$> _ |- _ => symmetry in H
+  end.
 
 (** * Indexed folds and maps *)
 (** We define stronger variants of map and fold that also take the index of the
@@ -459,7 +1736,7 @@ Definition imap_go {A B} (f : nat → A → B) : nat → list A → list B :=
   end.
 Definition imap {A B} (f : nat → A → B) : list A → list B := imap_go f 0.
 
-Definition ifold_right {A B} (f : nat → B → A → A)
+Definition ifoldr {A B} (f : nat → B → A → A)
     (a : nat → A) : nat → list B → A :=
   fix go (n : nat) (l : list B) : A :=
   match l with
@@ -467,9 +1744,9 @@ Definition ifold_right {A B} (f : nat → B → A → A)
   | b :: l => f n b (go (S n) l)
   end.
 
-Lemma ifold_right_app {A B} (f : nat → B → A → A) (a : nat → A)
+Lemma ifoldr_app {A B} (f : nat → B → A → A) (a : nat → A)
     (l1 l2 : list B) n :
-  ifold_right f a n (l1 ++ l2) = ifold_right f (λ n, ifold_right f a n l2) n l1.
+  ifoldr f a n (l1 ++ l2) = ifoldr f (λ n, ifoldr f a n l2) n l1.
 Proof.
   revert n a. induction l1 as [| b l1 IH ]; intros; simpl; f_equal; auto.
 Qed.
@@ -485,59 +1762,421 @@ Section same_length.
     | same_length_cons x y l k :
       same_length l k → same_length (x :: l) (y :: k).
 
-  Lemma same_length_length l k :
-    same_length l k ↔ length l = length k.
+  Lemma same_length_length_1 l k :
+    same_length l k → length l = length k.
+  Proof. induction 1; simpl; auto. Qed.
+  Lemma same_length_length_2 l k :
+    length l = length k → same_length l k.
   Proof.
-    split.
-    * induction 1; simpl; auto.
-    * revert k. induction l; intros [|??]; try discriminate;
+    revert k. induction l; intros [|??]; try discriminate;
       constructor; auto with arith.
   Qed.
+  Lemma same_length_length l k :
+    same_length l k ↔ length l = length k.
+  Proof. split; auto using same_length_length_1, same_length_length_2. Qed.
 
   Lemma same_length_lookup l k i :
     same_length l k → is_Some (l !! i) → is_Some (k !! i).
   Proof.
     rewrite same_length_length.
-    setoid_rewrite list_lookup_lt_length.
-    intros E. now rewrite E.
+    setoid_rewrite lookup_lt_length.
+    intros E. by rewrite E.
   Qed.
+
+  Lemma Forall2_same_length (P : A → B → Prop) l1 l2 :
+    Forall2 P l1 l2 →
+    same_length l1 l2.
+  Proof. intro. eapply same_length_length, Forall2_length; eauto. Qed.
+  Lemma Forall2_app_inv (P : A → B → Prop) l1 l2 k1 k2 :
+    same_length l1 k1 →
+    Forall2 P (l1 ++ l2) (k1 ++ k2) → Forall2 P l2 k2.
+  Proof. induction 1. done. inversion 1; subst; auto. Qed.
+
+  Lemma same_length_Forall2 l1 l2 :
+    same_length l1 l2 ↔ Forall2 (λ _ _, True) l1 l2.
+  Proof. split; induction 1; constructor; auto. Qed.
+
+  Lemma same_length_take l1 l2 n :
+    same_length l1 l2 →
+    same_length (take n l1) (take n l2).
+  Proof. rewrite !same_length_Forall2. apply Forall2_take. Qed.
+  Lemma same_length_drop l1 l2 n :
+    same_length l1 l2 →
+    same_length (drop n l1) (drop n l2).
+  Proof. rewrite !same_length_Forall2. apply Forall2_drop. Qed.
+  Lemma same_length_resize l1 l2 x1 x2 n :
+    same_length (resize n x1 l1) (resize n x2 l2).
+  Proof. apply same_length_length. by rewrite !resize_length. Qed.
 End same_length.
 
+Instance: ∀ A, Reflexive (@same_length A A).
+Proof. intros A l. induction l; constructor; auto. Qed.
+
 (** * Zipping lists *)
-(** Since we prefer Haskell style naming, we rename the standard library's
-implementation [combine] into [zip] using a notation. *)
-Notation zip := combine.
+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.
 
-Section zip.
-  Context {A B : Type}.
+Section zip_with.
+  Context {A B C : Type} (f : A → B → C).
 
-  Local Arguments fmap _ _ _ _ _ !_ /.
+  Lemma zip_with_length l1 l2 :
+    length l1 ≤ length l2 →
+    length (zip_with f l1 l2) = length l1.
+  Proof.
+    revert l2.
+    induction l1; intros [|??]; simpl; auto with lia.
+  Qed.
 
-  Lemma zip_fst_le (l1 : list A) (l2 : list B) :
-    length l1 ≤ length l2 → fst <$> zip l1 l2 = l1.
+  Lemma zip_with_fmap_fst_le (g : C → A) l1 l2 :
+    (∀ x y, g (f x y) = x) →
+    length l1 ≤ length l2 →
+    g <$> zip_with f l1 l2 = l1.
   Proof.
     revert l2.
-    induction l1; intros [|??] ?; simpl; f_equal; auto with arith.
-    edestruct Le.le_Sn_0; eauto.
+    induction l1; intros [|??] ??; simpl in *; f_equal; auto with lia.
   Qed.
-  Lemma zip_fst (l1 : list A) (l2 : list B) :
-    same_length l1 l2 → fst <$> zip l1 l2 = l1.
+  Lemma zip_with_fmap_snd_le (g : C → B) l1 l2 :
+    (∀ x y, g (f x y) = y) →
+    length l2 ≤ length l1 →
+    g <$> zip_with f l1 l2 = l2.
   Proof.
-    rewrite same_length_length. intros H.
-    apply zip_fst_le. now rewrite H.
+    revert l1.
+    induction l2; intros [|??] ??; simpl in *; f_equal; auto with lia.
   Qed.
+  Lemma zip_with_fmap_fst (g : C → A) l1 l2 :
+    (∀ x y, g (f x y) = x) →
+    same_length l1 l2 →
+    g <$> zip_with f l1 l2 = l1.
+  Proof. induction 2; simpl; f_equal; auto. Qed.
+  Lemma zip_with_fmap_snd (g : C → B) l1 l2 :
+    (∀ x y, g (f x y) = y) →
+    same_length l1 l2 →
+    g <$> zip_with f l1 l2 = l2.
+  Proof. induction 2; simpl; f_equal; auto. Qed.
 
-  Lemma zip_snd_le (l1 : list A) (l2 : list B) :
-    length l2 ≤ length l1 → snd <$> zip l1 l2 = l2.
+  Lemma Forall_zip_with_fst (P : A → Prop) (Q : C → Prop) l1 l2 :
+    Forall P l1 →
+    Forall (λ y, ∀ x, P x → Q (f x y)) l2 →
+    Forall Q (zip_with f l1 l2).
   Proof.
-    revert l2.
-    induction l1; intros [|??] ?; simpl; f_equal; auto with arith.
-    edestruct Le.le_Sn_0; eauto.
+    intros Hl1. revert l2.
+    induction Hl1; destruct 1; simpl in *; auto.
   Qed.
+  Lemma Forall_zip_with_snd (P : B → Prop) (Q : C → Prop) l1 l2 :
+    Forall (λ x, ∀ y, P y → Q (f x y)) l1 →
+    Forall P l2 →
+    Forall Q (zip_with f l1 l2).
+  Proof.
+    intros Hl1. revert l2.
+    induction Hl1; destruct 1; simpl in *; auto.
+  Qed.
+End zip_with.
+
+Notation zip := (zip_with pair).
+
+Section zip.
+  Context {A B : Type}.
+
+  Lemma zip_length (l1 : list A) (l2 : list B) :
+    length l1 ≤ length l2 →
+    length (zip l1 l2) = length l1.
+  Proof. by apply zip_with_length. Qed.
+
+  Lemma zip_fmap_fst_le (l1 : list A) (l2 : list B) :
+    length l1 ≤ length l2 →
+    fst <$> zip l1 l2 = l1.
+  Proof. by apply zip_with_fmap_fst_le. Qed.
+  Lemma zip_fmap_snd (l1 : list A) (l2 : list B) :
+    length l2 ≤ length l1 →
+    snd <$> zip l1 l2 = l2.
+  Proof. by apply zip_with_fmap_snd_le. Qed.
+
+  Lemma zip_fst (l1 : list A) (l2 : list B) :
+    same_length l1 l2 →
+    fst <$> zip l1 l2 = l1.
+  Proof. by apply zip_with_fmap_fst. Qed.
   Lemma zip_snd (l1 : list A) (l2 : list B) :
     same_length l1 l2 → snd <$> zip l1 l2 = l2.
+  Proof. by apply zip_with_fmap_snd. Qed.
+End zip.
+
+Definition zipped_map {A B} (f : list A → list A → A → B) :
+    list A → list A → list B :=
+  fix go l k :=
+  match k with
+  | [] => []
+  | x :: k => f l k x :: go (x :: l) k
+  end.
+
+Lemma elem_of_zipped_map {A B} (f : list A → list A → A → B) l k x :
+  x ∈ zipped_map f l k ↔
+    ∃ k' k'' y, k = k' ++ [y] ++ k'' ∧ x = f (reverse k' ++ l) k'' y.
+Proof.
+  split.
+  * revert l. induction k as [|z k IH]; simpl;
+      intros l ?; decompose_elem_of_list.
+    + by eexists [], k, z.
+    + destruct (IH (z :: l)) as [k' [k'' [y [??]]]]; [done |]; subst.
+      eexists (z :: k'), k'', y. split; [done |].
+      by rewrite reverse_cons, <-app_assoc.
+  * intros [k' [k'' [y [??]]]]; subst.
+    revert l. induction k' as [|z k' IH]; intros l.
+    + by left.
+    + right. by rewrite reverse_cons, <-!app_assoc.
+Qed.
+
+Section zipped_list_ind.
+  Context {A} (P : list A → list A → Prop).
+  Context (Pnil : ∀ l, P l []).
+  Context (Pcons : ∀ l k x, P (x :: l) k → P l (x :: k)).
+
+  Fixpoint zipped_list_ind l k : P l k :=
+    match k with
+    | [] => Pnil _
+    | x :: k => Pcons _ _ _ (zipped_list_ind (x :: l) k)
+    end.
+End zipped_list_ind.
+
+Inductive zipped_Forall {A} (P : list A → list A → A → Prop) :
+    list A → list A → Prop :=
+  | zipped_Forall_nil l : zipped_Forall P l []
+  | zipped_Forall_cons l k x :
+     P l k x →
+     zipped_Forall P (x :: l) k →
+     zipped_Forall P l (x :: k).
+Arguments zipped_Forall_nil {_ _} _.
+Arguments zipped_Forall_cons {_ _} _ _ _ _ _.
+
+Lemma zipped_Forall_app {A} (P : list A → list A → A → Prop) l k k' :
+  zipped_Forall P l (k ++ k') → zipped_Forall P (reverse k ++ l) k'.
+Proof.
+  revert l. induction k as [|x k IH]; simpl; [done |].
+  inversion_clear 1. rewrite reverse_cons, <-app_assoc.
+  by apply IH.
+Qed.
+
+(** * Permutations *)
+Fixpoint interleave {A} (x : A) (l : list A) : list (list A) :=
+  match l with
+  | [] => [ [x] ]
+  | y :: l => (x :: y :: l) :: ((y ::) <$> interleave x l)
+  end.
+Fixpoint permutations {A} (l : list A) : list (list A) :=
+  match l with
+  | [] => [ [] ]
+  | x :: l => permutations l ≫= interleave x
+  end.
+
+Section permutations.
+  Context {A : Type}.
+
+  Lemma interleave_cons (x : A) (l : list A) :
+    x :: l ∈ interleave x l.
+  Proof. destruct l; simpl; rewrite elem_of_cons; auto. Qed.
+  Lemma interleave_Permutation (x : A) (l l' : list A) :
+    l' ∈ interleave x l → Permutation l' (x :: l).
   Proof.
-    rewrite same_length_length. intros H.
-    apply zip_snd_le. now rewrite H.
+    revert l'. induction l as [|y l IH]; intros l'; simpl.
+    * rewrite elem_of_list_singleton. intros. by subst.
+    * rewrite elem_of_cons, elem_of_list_fmap.
+      intros [?|[? [? H]]]; subst.
+      + by constructor.
+      + rewrite (IH _ H). constructor.
   Qed.
-End zip.
+
+  Lemma permutations_refl (l : list A) :
+    l ∈ permutations l.
+  Proof.
+    induction l; simpl.
+    * by apply elem_of_list_singleton.
+    * apply elem_of_list_bind. eauto using interleave_cons.
+  Qed.
+  Lemma permutations_skip (x : A) (l l' : list A) : 
+    l ∈ permutations l' →
+    x :: l ∈ permutations (x :: l').
+  Proof.
+    intros Hl. simpl. apply elem_of_list_bind.
+    eauto using interleave_cons.
+  Qed.
+  Lemma permutations_swap (x y : A) (l : list A) : 
+    y :: x :: l ∈ permutations (x :: y :: l).
+  Proof.
+    simpl. apply elem_of_list_bind.
+    exists (y :: l). split; simpl.
+    * destruct l; simpl; rewrite !elem_of_cons; auto.
+    * apply elem_of_list_bind. simpl.
+      eauto using interleave_cons, permutations_refl.
+  Qed.
+  Lemma permutations_nil (l : list A) :
+    l ∈ permutations [] ↔ l = [].
+  Proof. simpl. by rewrite elem_of_list_singleton. Qed.
+
+  Lemma interleave_interleave_toggle (x1 x2 : A) (l1 l2 l3 : list A) :
+    l1 ∈ interleave x1 l2 →
+    l2 ∈ interleave x2 l3 → ∃ l4,
+      l1 ∈ interleave x2 l4 ∧ l4 ∈ interleave x1 l3.
+  Proof.
+    revert l1 l2. induction l3 as [|y l3 IH]; intros l1 l2; simpl.
+    { intros Hl1 Hl2.
+      rewrite elem_of_list_singleton in Hl2. subst. simpl in Hl1.
+      rewrite elem_of_cons, elem_of_list_singleton in Hl1.
+      exists [x1]. simpl.
+      rewrite elem_of_cons, !elem_of_list_singleton. tauto. }
+    rewrite elem_of_cons, elem_of_list_fmap.
+    intros Hl1 [? | [l2' [??]]]; subst; simpl in *.
+    * rewrite !elem_of_cons, elem_of_list_fmap in Hl1.
+      destruct Hl1 as [? | [? | [l4 [??]]]]; subst.
+      + exists (x1 :: y :: l3). simpl. rewrite !elem_of_cons. tauto.
+      + exists (x1 :: y :: l3). simpl. rewrite !elem_of_cons. tauto.
+      + exists l4. simpl. rewrite elem_of_cons. auto using interleave_cons.
+    * rewrite elem_of_cons, elem_of_list_fmap in Hl1.
+      destruct Hl1 as [? | [l1' [??]]]; subst.
+      + exists (x1 :: y :: l3). simpl.
+        rewrite !elem_of_cons, !elem_of_list_fmap.
+        split; [| by auto]. right. right. exists (y :: l2').
+        rewrite elem_of_list_fmap. naive_solver.
+      + destruct (IH l1' l2') as [l4 [??]]; auto.
+        exists (y :: l4). simpl.
+        rewrite !elem_of_cons, !elem_of_list_fmap. naive_solver.
+  Qed.
+  Lemma permutations_interleave_toggle (x : A) (l1 l2 l3 : list A) :
+    l1 ∈ permutations l2 →
+    l2 ∈ interleave x l3 → ∃ l4,
+      l1 ∈ interleave x l4 ∧ l4 ∈ permutations l3.
+  Proof.
+    revert l1 l2. induction l3 as [|y l3 IH]; intros l1 l2; simpl.
+    { intros Hl1 Hl2. eexists []. simpl.
+      split; [| by rewrite elem_of_list_singleton].
+      rewrite elem_of_list_singleton in Hl2.
+      by rewrite Hl2 in Hl1. }
+    rewrite elem_of_cons, elem_of_list_fmap.
+    intros Hl1 [? | [l2' [? Hl2']]]; subst; simpl in *.
+    * rewrite elem_of_list_bind in Hl1.
+      destruct Hl1 as [l1' [??]]. by exists l1'.
+    * rewrite elem_of_list_bind in Hl1.
+      setoid_rewrite elem_of_list_bind.
+      destruct Hl1 as [l1' [??]].
+      destruct (IH l1' l2') as [l1'' [??]]; auto.
+      destruct (interleave_interleave_toggle y x l1 l1' l1'') as [? [??]]; eauto.
+  Qed.
+  Lemma permutations_trans (l1 l2 l3 : list A) :
+    l1 ∈ permutations l2 →
+    l2 ∈ permutations l3 →
+    l1 ∈ permutations l3.
+  Proof.
+    revert l1 l2. induction l3 as [|x l3 IH]; intros l1 l2; simpl.
+    * intros Hl1 Hl2. rewrite elem_of_list_singleton in Hl2.
+      by rewrite Hl2 in Hl1.
+    * rewrite !elem_of_list_bind. intros Hl1 [l2' [Hl2 Hl2']].
+      destruct (permutations_interleave_toggle x l1 l2 l2') as [? [??]]; eauto.
+  Qed.
+
+  Lemma permutations_Permutation (l l' : list A) :
+    l' ∈ permutations l ↔ Permutation l l'.
+  Proof.
+    split.
+    * revert l'. induction l; simpl; intros l''.
+      + rewrite elem_of_list_singleton.
+        intros. subst. constructor.
+      + rewrite elem_of_list_bind. intros [l' [Hl'' ?]].
+        rewrite (interleave_Permutation _ _ _ Hl'').
+        constructor; auto.
+    * induction 1; eauto using permutations_refl,
+        permutations_skip, permutations_swap, permutations_trans.
+  Qed.
+
+  Global Instance Permutation_dec `{∀ x y : A, Decision (x = y)}
+    (l1 l2 : list A) : Decision (Permutation l1 l2).
+  Proof.
+    refine (cast_if (decide (l2 ∈ permutations l1)));
+      by rewrite <-permutations_Permutation.
+  Defined.
+End permutations.
+
+(** * Set operations on lists *)
+Section list_set_operations.
+  Context {A} {dec : ∀ x y : A, Decision (x = y)}.
+
+  Fixpoint list_difference (l k : list A) : list A :=
+    match l with
+    | [] => []
+    | x :: l =>
+      if decide_rel (∈) x k
+      then list_difference l k
+      else x :: list_difference l k
+    end.
+  Lemma elem_of_list_difference l k x :
+    x ∈ list_difference l k ↔ x ∈ l ∧ x ∉ k.
+  Proof.
+    split; induction l; simpl; try case_decide;
+      rewrite ?elem_of_nil, ?elem_of_cons; intuition congruence.
+  Qed.
+  Lemma list_difference_nodup l k :
+    NoDup l → NoDup (list_difference l k).
+  Proof.
+    induction 1; simpl; try case_decide.
+    * constructor.
+    * done.
+    * constructor. rewrite elem_of_list_difference; intuition. done.
+  Qed.
+
+  Fixpoint list_intersection (l k : list A) : list A :=
+    match l with
+    | [] => []
+    | x :: l =>
+      if decide_rel (∈) x k
+      then x :: list_intersection l k
+      else list_intersection l k
+    end.
+  Lemma elem_of_list_intersection l k x :
+    x ∈ list_intersection l k ↔ x ∈ l ∧ x ∈ k.
+  Proof.
+    split; induction l; simpl; repeat case_decide;
+      rewrite ?elem_of_nil, ?elem_of_cons; intuition congruence.
+  Qed.
+  Lemma list_intersection_nodup l k :
+    NoDup l → NoDup (list_intersection l k).
+  Proof.
+    induction 1; simpl; try case_decide.
+    * constructor.
+    * constructor. rewrite elem_of_list_intersection; intuition. done.
+    * done.
+  Qed.
+
+  Definition list_intersection_with (f : A → A → option A) :
+      list A → list A → list A :=
+    fix go l k :=
+    match l with
+    | [] => []
+    | x :: l => foldr (λ y,
+       match f x y with None => id | Some z => (z ::) end) (go l k) k
+    end.
+  Lemma elem_of_list_intersection_with f l k x :
+    x ∈ list_intersection_with f l k ↔ ∃ x1 x2,
+      x1 ∈ l ∧ x2 ∈ k ∧ f x1 x2 = Some x.
+  Proof.
+    split.
+    * induction l as [|x1 l IH]; simpl.
+      + by rewrite elem_of_nil.
+      + intros Hx. setoid_rewrite elem_of_cons.
+        cut ((∃ x2, x2 ∈ k ∧ f x1 x2 = Some x)
+          ∨ x ∈ list_intersection_with f l k).
+        { naive_solver. }
+        clear IH. revert Hx. generalize (list_intersection_with f l k).
+        induction k; simpl; [by auto|].
+        case_match; setoid_rewrite elem_of_cons; naive_solver.
+    * intros (x1 & x2 & Hx1 & Hx2 & Hx).
+      induction Hx1 as [x1 | x1 ? l ? IH]; simpl.
+      + generalize (list_intersection_with f l k).
+        induction Hx2; simpl; [by rewrite Hx; left |].
+        case_match; simpl; try setoid_rewrite elem_of_cons; auto.
+      + generalize (IH Hx). clear Hx IH Hx2.
+        generalize (list_intersection_with f l k).
+        induction k; simpl; intros; [done |].
+        case_match; simpl; rewrite ?elem_of_cons; auto.
+  Qed.
+End list_set_operations.
diff --git a/theories/listset.v b/theories/listset.v
index 682319f0fddbf80043c36fe0360266e204e0bd1e..21da87526f4ea6fd31167f0c7b3333f06e41c739 100644
--- a/theories/listset.v
+++ b/theories/listset.v
@@ -1,130 +1,119 @@
 (* Copyright (c) 2012, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-(** This file implements finite as unordered lists without duplicates.
-Although this implementation is slow, it is very useful as decidable equality
-is the only constraint on the carrier set. *)
-Require Export base decidable list collections.
+(** This file implements finite as unordered lists without duplicates
+removed. This implementation forms a monad. *)
+Require Export base decidable collections list.
 
-Definition listset A := sig (@NoDup A).
+Record listset A := Listset {
+  listset_car: list A
+}.
+Arguments listset_car {_} _.
+Arguments Listset {_} _.
 
-Section list_collection.
-Context {A : Type} `{∀ x y : A, Decision (x = y)}.
+Section listset.
+Context {A : Type}.
 
-Global Instance listset_elem_of: ElemOf A (listset A) := λ x l, In x (`l).
-Global Instance listset_empty: Empty (listset A) := []↾@NoDup_nil _.
-Global Instance listset_singleton: Singleton A (listset A) := λ x,
-  [x]↾NoDup_singleton x.
-
-Fixpoint listset_difference_raw (l k : list A) :=
-  match l with
-  | [] => []
-  | x :: l =>
-    if decide_rel In x k
-    then listset_difference_raw l k
-    else x :: listset_difference_raw l k
+Instance listset_elem_of: ElemOf A (listset A) := λ x l,
+  x ∈ listset_car l.
+Instance listset_empty: Empty (listset A) :=
+  Listset [].
+Instance listset_singleton: Singleton A (listset A) := λ x,
+  Listset [x].
+Instance listset_union: Union (listset A) := λ l k,
+  match l, k with
+  | Listset l', Listset k' => Listset (l' ++ k')
   end.
-Lemma listset_difference_raw_in l k x :
-  In x (listset_difference_raw l k) ↔ In x l ∧ ¬In x k.
-Proof.
-  split; induction l; simpl; try case_decide; simpl; intuition congruence.
-Qed.
-Lemma listset_difference_raw_nodup l k :
-  NoDup l → NoDup (listset_difference_raw l k).
-Proof.
-  induction 1; simpl; try case_decide.
-  * constructor.
-  * easy.
-  * constructor. rewrite listset_difference_raw_in; intuition. easy.
-Qed.
-Global Instance listset_difference: Difference (listset A) := λ l k,
-  listset_difference_raw (`l) (`k)↾
-    listset_difference_raw_nodup (`l) (`k) (proj2_sig l).
 
-Definition listset_union_raw (l k : list A) := listset_difference_raw l k ++ k.
-Lemma listset_union_raw_in l k x :
-  In x (listset_union_raw l k) ↔ In x l ∨ In x k.
-Proof.
-  unfold listset_union_raw. rewrite in_app_iff, listset_difference_raw_in.
-  intuition. case (decide (In x k)); intuition.
-Qed.
-Lemma listset_union_raw_nodup l k :
-  NoDup l → NoDup k → NoDup (listset_union_raw l k).
+Global Instance: SimpleCollection A (listset A).
 Proof.
-  intros. apply NoDup_app.
-  * now apply listset_difference_raw_nodup.
-  * easy.
-  * intro. rewrite listset_difference_raw_in. intuition.
+  split.
+  * by apply not_elem_of_nil.
+  * by apply elem_of_list_singleton.
+  * intros [?] [?]. apply elem_of_app.
 Qed.
-Global Instance listset_union: Union (listset A) := λ l k,
-  listset_union_raw (`l) (`k)↾
-    listset_union_raw_nodup (`l) (`k) (proj2_sig l) (proj2_sig k).
 
-Fixpoint listset_intersection_raw (l k : list A) :=
+Context `{∀ x y : A, Decision (x = y)}.
+
+Instance listset_intersection: Intersection (listset A) := λ l k,
+  match l, k with
+  | Listset l', Listset k' => Listset (list_intersection l' k')
+  end.
+Instance listset_difference: Difference (listset A) := λ l k,
+  match l, k with
+  | Listset l', Listset k' => Listset (list_difference l' k')
+  end.
+Instance listset_intersection_with: IntersectionWith A (listset A) := λ f l k,
+  match l, k with
+  | Listset l', Listset k' => Listset (list_intersection_with f l' k')
+  end.
+Instance listset_filter: Filter A (listset A) := λ P _ l,
   match l with
-  | [] => []
-  | x :: l =>
-    if decide_rel In x k
-    then x :: listset_intersection_raw l k
-    else listset_intersection_raw l k
+  | Listset l' => Listset (filter P l')
   end.
-Lemma listset_intersection_raw_in l k x :
-  In x (listset_intersection_raw l k) ↔ In x l ∧ In x k.
-Proof.
-  split; induction l; simpl; try case_decide; simpl; intuition congruence.
-Qed.
-Lemma listset_intersection_raw_nodup l k :
-  NoDup l → NoDup (listset_intersection_raw l k).
+
+Global Instance: Collection A (listset A).
 Proof.
-  induction 1; simpl; try case_decide.
-  * constructor.
-  * constructor. rewrite listset_intersection_raw_in; intuition. easy.
-  * easy.
+  split.
+  * apply _.
+  * intros [?] [?]. apply elem_of_list_intersection.
+  * intros [?] [?]. apply elem_of_list_difference.
+  * intros ? [?] [?]. apply elem_of_list_intersection_with.
 Qed.
-Global Instance listset_intersection: Intersection (listset A) := λ l k,
-  listset_intersection_raw (`l) (`k)↾
-    listset_intersection_raw_nodup (`l) (`k) (proj2_sig l).
 
-Definition listset_add_raw x (l : list A) : list A :=
-  if decide_rel In x l then l else x :: l.
-Lemma listset_add_raw_in x l y : In y (listset_add_raw x l) ↔ y = x ∨ In y l.
-Proof. unfold listset_add_raw. case_decide; firstorder congruence. Qed.
-Lemma listset_add_raw_nodup x l : NoDup l → NoDup (listset_add_raw x l).
+Instance listset_elems: Elements A (listset A) :=
+  remove_dups ∘ listset_car.
+
+Global Instance: FinCollection A (listset A).
 Proof.
-  unfold listset_add_raw. case_decide; try constructor; firstorder.
+  split.
+  * apply _.
+  * intros [?] ??. apply elem_of_list_filter.
+  * symmetry. apply elem_of_remove_dups.
+  * intros. apply remove_dups_nodup.
 Qed.
+End listset.
+
+(** These instances are declared using [Hint Extern] to avoid too
+eager type class search. *)
+Hint Extern 1 (ElemOf _ (listset _)) =>
+  eapply @listset_elem_of : typeclass_instances.
+Hint Extern 1 (Empty (listset _)) =>
+  eapply @listset_empty : typeclass_instances.
+Hint Extern 1 (Singleton _ (listset _)) =>
+  eapply @listset_singleton : typeclass_instances.
+Hint Extern 1 (Union (listset _)) =>
+  eapply @listset_union : typeclass_instances.
+Hint Extern 1 (Intersection (listset _)) =>
+  eapply @listset_intersection : typeclass_instances.
+Hint Extern 1 (IntersectionWith _ (listset _)) =>
+  eapply @listset_intersection_with : typeclass_instances.
+Hint Extern 1 (Difference (listset _)) =>
+  eapply @listset_difference : typeclass_instances.
+Hint Extern 1 (Elements _ (listset _)) =>
+  eapply @listset_elems : typeclass_instances.
+Hint Extern 1 (Filter _ (listset _)) =>
+  eapply @listset_filter : typeclass_instances.
 
-Fixpoint listset_map_raw (f : A → A) (l : list A) :=
+Instance listset_ret: MRet listset := λ A x,
+  {[ x ]}.
+Instance listset_fmap: FMap listset := λ A B f l,
   match l with
-  | [] => []
-  | x :: l => listset_add_raw (f x) (listset_map_raw f l)
+  | Listset l' => Listset (f <$> l')
   end.
-Lemma listset_map_raw_nodup f l : NoDup (listset_map_raw f l).
-Proof. induction l; simpl. constructor. now apply listset_add_raw_nodup. Qed.
-Lemma listset_map_raw_in f l x :
-  In x (listset_map_raw f l) ↔ ∃ y, x = f y ∧ In y l.
-Proof.
-  split.
-  * induction l; simpl; [easy |].
-    rewrite listset_add_raw_in. firstorder.
-  * intros [?[??]]. subst. induction l; simpl in *; [easy |].
-    rewrite listset_add_raw_in. firstorder congruence.
-Qed.
-Global Instance listset_map: Map A (listset A) := λ f l,
-  listset_map_raw f (`l)↾listset_map_raw_nodup f (`l).
+Instance listset_bind: MBind listset := λ A B f l,
+  match l with
+  | Listset l' => Listset (mbind (listset_car ∘ f) l')
+  end.
+Instance listset_join: MJoin listset := λ A, mbind id.
 
-Global Instance: Collection A (listset A).
+Instance: CollectionMonad listset.
 Proof.
   split.
-  * easy.
-  * compute. intuition.
-  * intros. apply listset_union_raw_in.
-  * intros. apply listset_intersection_raw_in.
-  * intros. apply listset_difference_raw_in.
-  * intros. apply listset_map_raw_in.
+  * intros. apply _.
+  * intros ??? [?] ?. apply elem_of_list_bind.
+  * intros. apply elem_of_list_ret.
+  * intros ??? [?]. apply elem_of_list_fmap.
+  * intros ? [?] ?.
+    unfold mjoin, listset_join, elem_of, listset_elem_of.
+    simpl. by rewrite elem_of_list_bind.
 Qed.
-
-Global Instance listset_elems: Elements A (listset A) := @proj1_sig _ _.
-
-Global Instance: FinCollection A (listset A).
-Proof. split. apply _. easy. now intros [??]. Qed.
-End list_collection.
diff --git a/theories/listset_nodup.v b/theories/listset_nodup.v
new file mode 100644
index 0000000000000000000000000000000000000000..a63c969a5f1149fd13471c81502c4760b651b8ab
--- /dev/null
+++ b/theories/listset_nodup.v
@@ -0,0 +1,103 @@
+(* Copyright (c) 2012, Robbert Krebbers. *)
+(* This file is distributed under the terms of the BSD license. *)
+(** This file implements finite as unordered lists without duplicates.
+Although this implementation is slow, it is very useful as decidable equality
+is the only constraint on the carrier set. *)
+Require Export base decidable collections list.
+
+Record listset_nodup A := ListsetNoDup {
+  listset_nodup_car : list A;
+  listset_nodup_prf : NoDup listset_nodup_car
+}.
+Arguments ListsetNoDup {_} _ _.
+Arguments listset_nodup_car {_} _.
+Arguments listset_nodup_prf {_} _.
+
+Section list_collection.
+Context {A : Type} `{∀ x y : A, Decision (x = y)}.
+
+Notation C := (listset_nodup A).
+Notation LS := ListsetNoDup.
+
+Instance listset_nodup_elem_of: ElemOf A C := λ x l,
+  x ∈ listset_nodup_car l.
+Instance listset_nodup_empty: Empty C :=
+  LS [] (@NoDup_nil_2 _).
+Instance listset_nodup_singleton: Singleton A C := λ x,
+  LS [x] (NoDup_singleton x).
+Instance listset_nodup_difference: Difference C := λ l k,
+  LS _ (list_difference_nodup _ (listset_nodup_car k) (listset_nodup_prf l)).
+
+Definition listset_nodup_union_raw (l k : list A) : list A :=
+  list_difference l k ++ k.
+Lemma elem_of_listset_nodup_union_raw l k x :
+  x ∈ listset_nodup_union_raw l k ↔ x ∈ l ∨ x ∈ k.
+Proof.
+  unfold listset_nodup_union_raw.
+  rewrite elem_of_app, elem_of_list_difference.
+  intuition. case (decide (x ∈ k)); intuition.
+Qed.
+Lemma listset_nodup_union_raw_nodup l k :
+  NoDup l → NoDup k → NoDup (listset_nodup_union_raw l k).
+Proof.
+  intros. apply NoDup_app. repeat split.
+  * by apply list_difference_nodup.
+  * intro. rewrite elem_of_list_difference. intuition.
+  * done.
+Qed.
+Instance listset_nodup_union: Union C := λ l k,
+  LS _ (listset_nodup_union_raw_nodup _ _
+     (listset_nodup_prf l) (listset_nodup_prf k)).
+Instance listset_nodup_intersection: Intersection C := λ l k,
+  LS _ (list_intersection_nodup _
+     (listset_nodup_car k) (listset_nodup_prf l)).
+Instance listset_nodup_intersection_with:
+    IntersectionWith A C := λ f l k,
+  LS (remove_dups
+      (list_intersection_with f (listset_nodup_car l) (listset_nodup_car k)))
+    (remove_dups_nodup _).
+Instance listset_nodup_filter: Filter A C :=
+  λ P _ l, LS _ (filter_nodup P _ (listset_nodup_prf l)).
+
+Global Instance: Collection A C.
+Proof.
+  split; [split | | |].
+  * by apply not_elem_of_nil.
+  * by apply elem_of_list_singleton.
+  * intros. apply elem_of_listset_nodup_union_raw.
+  * intros. apply elem_of_list_intersection.
+  * intros. apply elem_of_list_difference.
+  * intros. unfold intersection_with, listset_nodup_intersection_with,
+      elem_of, listset_nodup_elem_of. simpl.
+    rewrite elem_of_remove_dups.
+    by apply elem_of_list_intersection_with.
+Qed.
+
+Global Instance listset_nodup_elems: Elements A C := listset_nodup_car.
+
+Global Instance: FinCollection A C.
+Proof.
+  split.
+  * apply _.
+  * intros. apply elem_of_list_filter.
+  * done.
+  * by intros [??].
+Qed.
+End list_collection.
+
+Hint Extern 1 (ElemOf _ (listset_nodup _)) =>
+  eapply @listset_nodup_elem_of : typeclass_instances.
+Hint Extern 1 (Empty (listset_nodup _)) =>
+  eapply @listset_nodup_empty : typeclass_instances.
+Hint Extern 1 (Singleton _ (listset_nodup _)) =>
+  eapply @listset_nodup_singleton : typeclass_instances.
+Hint Extern 1 (Union (listset_nodup _)) =>
+  eapply @listset_nodup_union : typeclass_instances.
+Hint Extern 1 (Intersection (listset_nodup _)) =>
+  eapply @listset_nodup_intersection : typeclass_instances.
+Hint Extern 1 (Difference (listset_nodup _)) =>
+  eapply @listset_nodup_difference : typeclass_instances.
+Hint Extern 1 (Elements _ (listset_nodup _)) =>
+  eapply @listset_nodup_elems : typeclass_instances.
+Hint Extern 1 (Filter _ (listset_nodup _)) =>
+  eapply @listset_nodup_filter : typeclass_instances.
diff --git a/theories/nmap.v b/theories/nmap.v
index 1ee586ee547b0358f4bd6398bcee35386d67a8e3..312e9ee9b9e52fae96f934b56095504021691c21 100644
--- a/theories/nmap.v
+++ b/theories/nmap.v
@@ -7,56 +7,74 @@ Require Export prelude fin_maps.
 
 Local Open Scope N_scope.
 
-Record Nmap A := { Nmap_0 : option A; Nmap_pos : Pmap A }.
+Record Nmap A := NMap { Nmap_0 : option A; Nmap_pos : Pmap A }.
 Arguments Nmap_0 {_} _.
 Arguments Nmap_pos {_} _.
-Arguments Build_Nmap {_} _ _.
+Arguments NMap {_} _ _.
 
-Global Instance Pmap_dec `{∀ x y : A, Decision (x = y)} :
+Instance Pmap_dec `{∀ x y : A, Decision (x = y)} :
   ∀ x y : Nmap A, Decision (x = y).
 Proof. solve_decision. Defined.
 
-Global Instance Nempty {A} : Empty (Nmap A) := Build_Nmap None ∅.
-Global Instance Nlookup: Lookup N Nmap := λ A i t,
+Instance Nempty {A} : Empty (Nmap A) := NMap None ∅.
+Instance Nlookup {A} : Lookup N A (Nmap A) := λ i t,
   match i with
   | N0 => Nmap_0 t
   | Npos p => Nmap_pos t !! p
   end.
-Global Instance Npartial_alter: PartialAlter N Nmap := λ A f i t,
+Instance Npartial_alter {A} : PartialAlter N A (Nmap A) := λ f i t,
   match i, t with
-  | N0, Build_Nmap o t => Build_Nmap (f o) t
-  | Npos p, Build_Nmap o t => Build_Nmap o (partial_alter f p t)
+  | N0, NMap o t => NMap (f o) t
+  | Npos p, NMap o t => NMap o (partial_alter f p t)
   end.
-Global Instance Ndom: Dom N Nmap := λ C _ _ _ _ t,
+Instance Nto_list {A} : FinMapToList N A (Nmap A) := λ t,
   match t with
-  | Build_Nmap o t => option_case (λ _, {[ 0 ]}) ∅ o ∪ (Pdom_raw Npos (`t))
+  | NMap o t => option_case (λ x, [(0,x)]) [] o ++
+     (fst_map Npos <$> finmap_to_list t)
   end.
-Global Instance Nmerge: Merge Nmap := λ A f t1 t2,
+Instance Nmerge {A} : Merge A (Nmap A) := λ f t1 t2,
   match t1, t2 with
-  | Build_Nmap o1 t1, Build_Nmap o2 t2 => Build_Nmap (f o1 o2) (merge f t1 t2)
+  | NMap o1 t1, NMap o2 t2 => NMap (f o1 o2) (merge f t1 t2)
   end.
-Global Instance Nfmap: FMap Nmap := λ A B f t,
+Instance Nfmap: FMap Nmap := λ A B f t,
   match t with
-  | Build_Nmap o t => Build_Nmap (fmap f o) (fmap f t)
+  | NMap o t => NMap (fmap f o) (fmap f t)
   end.
 
-Global Instance: FinMap N Nmap.
+Instance: FinMap N Nmap.
 Proof.
   split.
   * intros ? [??] [??] H. f_equal.
-    + now apply (H 0).
-    + apply finmap_eq. intros i. now apply (H (Npos i)).
-  * now intros ? [|?].
-  * intros ? f [? t] [|i].
-    + easy.
-    + now apply (lookup_partial_alter f t i).
-  * intros ? f [? t] [|i] [|j]; try intuition congruence.
-    intros. apply (lookup_partial_alter_ne f t i j). congruence.
-  * intros ??? [??] []. easy. apply lookup_fmap.
-  * intros ?? ???????? [o t] n; unfold dom, lookup, Ndom, Nlookup; simpl.
-    rewrite elem_of_union, Plookup_raw_dom.
-    destruct o, n; esolve_elem_of (simplify_is_Some; eauto).
-  * intros ? f ? [o1 t1] [o2 t2] [|?].
-    + easy.
+    + apply (H 0).
+    + apply finmap_eq. intros i. apply (H (Npos i)).
+  * by intros ? [|?].
+  * intros ? f [? t] [|i]; simpl.
+    + done.
+    + apply lookup_partial_alter.
+  * intros ? f [? t] [|i] [|j]; simpl; try intuition congruence.
+    intros. apply lookup_partial_alter_ne. congruence.
+  * intros ??? [??] []; simpl. done. apply lookup_fmap.
+  * intros ? [[x|] t]; unfold finmap_to_list; simpl.
+    + constructor.
+      - rewrite elem_of_list_fmap. by intros [[??] [??]].
+      - rewrite (NoDup_fmap _). apply finmap_to_list_nodup.
+    + rewrite (NoDup_fmap _). apply finmap_to_list_nodup.
+  * intros ? t i x. unfold finmap_to_list. split.
+    + destruct t as [[y|] t]; simpl.
+      - rewrite elem_of_cons, elem_of_list_fmap.
+        intros [? | [[??] [??]]]; simplify_equality; simpl; [done |].
+        by apply elem_of_finmap_to_list.
+      - rewrite elem_of_list_fmap.
+        intros [[??] [??]]; simplify_equality; simpl.
+        by apply elem_of_finmap_to_list.
+    + destruct t as [[y|] t]; simpl.
+      - rewrite elem_of_cons, elem_of_list_fmap.
+        destruct i as [|i]; simpl; [intuition congruence |].
+        intros. right. exists (i, x). by rewrite elem_of_finmap_to_list.
+      - rewrite elem_of_list_fmap.
+        destruct i as [|i]; simpl; [done |].
+        intros. exists (i, x). by rewrite elem_of_finmap_to_list.
+  * intros ? f ? [o1 t1] [o2 t2] [|?]; simpl.
+    + done.
     + apply (merge_spec f t1 t2).
 Qed.
diff --git a/theories/numbers.v b/theories/numbers.v
index dbf29a5b00bb81f12a1cd71395b25836674093a3..c4198fefa7ea175d2bd269771f7837cc8956ea6f 100644
--- a/theories/numbers.v
+++ b/theories/numbers.v
@@ -1,19 +1,69 @@
 (* Copyright (c) 2012, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-(** This file collects some trivial facts on Coq's number data types [nat],
-[N], and [Z], and introduces some useful notations. *)
+(** This file collects some trivial facts on the Coq types [nat] and [N] for
+natural numbers, and the type [Z] for integers. It also declares some useful
+notations. *)
 Require Export PArith NArith ZArith.
-Require Export base decidable fin_collections.
+Require Export base decidable.
+
+Reserved Notation "x ≤ y ≤ z" (at level 70, y at next level).
+Reserved Notation "x ≤ y < z" (at level 70, y at next level).
+Reserved Notation "x < y < z" (at level 70, y at next level).
+Reserved Notation "x < y ≤ z" (at level 70, y at next level).
 
 Infix "≤" := le : nat_scope.
+Notation "x ≤ y ≤ z" := (x ≤ y ∧ y ≤ z)%nat : nat_scope.
+Notation "x ≤ y < z" := (x ≤ y ∧ y < z)%nat : nat_scope.
+Notation "x < y < z" := (x < y ∧ y < z)%nat : nat_scope.
+Notation "x < y ≤ z" := (x < y ∧ y ≤ z)%nat : nat_scope.
+Notation "(≤)" := le (only parsing) : nat_scope.
+Notation "(<)" := lt (only parsing) : nat_scope.
+
+Infix "`div`" := NPeano.div (at level 35) : nat_scope.
+Infix "`mod`" := NPeano.modulo (at level 35) : nat_scope.
 
 Instance nat_eq_dec: ∀ x y : nat, Decision (x = y) := eq_nat_dec.
+Instance nat_le_dec: ∀ x y : nat, Decision (x ≤ y) := le_dec.
+Instance nat_lt_dec: ∀ x y : nat, Decision (x < y) := lt_dec.
+Instance nat_inhabited: Inhabited nat := populate 0%nat.
+
+Lemma lt_n_SS n : n < S (S n).
+Proof. auto with arith. Qed.
+Lemma lt_n_SSS n : n < S (S (S n)).
+Proof. auto with arith. Qed.
+
+Definition sum_list_with {A} (f : A → nat) : list A → nat :=
+  fix go l :=
+  match l with
+  | [] => 0
+  | x :: l => f x + go l
+  end.
+Notation sum_list := (sum_list_with id).
+
 Instance positive_eq_dec: ∀ x y : positive, Decision (x = y) := Pos.eq_dec.
+Instance positive_inhabited: Inhabited positive := populate 1%positive.
+
 Notation "(~0)" := xO (only parsing) : positive_scope.
 Notation "(~1)" := xI (only parsing) : positive_scope.
 
+Instance: Injective (=) (=) xO.
+Proof. by injection 1. Qed.
+Instance: Injective (=) (=) xI.
+Proof. by injection 1. Qed.
+
 Infix "≤" := N.le : N_scope.
+Notation "x ≤ y ≤ z" := (x ≤ y ∧ y ≤ z)%N : N_scope.
+Notation "x ≤ y < z" := (x ≤ y ∧ y < z)%N : N_scope.
+Notation "x < y < z" := (x < y ∧ y < z)%N : N_scope.
+Notation "x < y ≤ z" := (x < y ∧ y ≤ z)%N : N_scope.
 Notation "(≤)" := N.le (only parsing) : N_scope.
+Notation "(<)" := N.lt (only parsing) : N_scope.
+
+Infix "`div`" := N.div (at level 35) : N_scope.
+Infix "`mod`" := N.modulo (at level 35) : N_scope.
+
+Instance: Injective (=) (=) Npos.
+Proof. by injection 1. Qed.
 
 Instance N_eq_dec: ∀ x y : N, Decision (x = y) := N.eq_dec.
 Program Instance N_le_dec (x y : N) : Decision (x ≤ y)%N :=
@@ -22,15 +72,33 @@ Program Instance N_le_dec (x y : N) : Decision (x ≤ y)%N :=
   | _ => left _
   end.
 Next Obligation. congruence. Qed.
+Program Instance N_lt_dec (x y : N) : Decision (x < y)%N :=
+  match Ncompare x y with
+  | Lt => left _
+  | _ => right _
+  end.
+Next Obligation. congruence. Qed.
+Instance N_inhabited: Inhabited N := populate 1%N.
 
 Infix "≤" := Z.le : Z_scope.
+Notation "x ≤ y ≤ z" := (x ≤ y ∧ y ≤ z)%Z : Z_scope.
+Notation "x ≤ y < z" := (x ≤ y ∧ y < z)%Z : Z_scope.
+Notation "x < y < z" := (x < y ∧ y < z)%Z : Z_scope.
+Notation "x < y ≤ z" := (x < y ∧ y ≤ z)%Z : Z_scope.
 Notation "(≤)" := Z.le (only parsing) : Z_scope.
+Notation "(<)" := Z.lt (only parsing) : Z_scope.
+
+Infix "`div`" := Z.div (at level 35) : Z_scope.
+Infix "`mod`" := Z.modulo (at level 35) : Z_scope.
+
 Instance Z_eq_dec: ∀ x y : Z, Decision (x = y) := Z.eq_dec.
 Instance Z_le_dec: ∀ x y : Z, Decision (x ≤ y)%Z := Z_le_dec.
+Instance Z_lt_dec: ∀ x y : Z, Decision (x < y)%Z := Z_lt_dec.
+Instance Z_inhabited: Inhabited Z := populate 1%Z.
 
 (** * Conversions *)
-(** Converts an integer [x] into a natural number by giving [None] in case [x]
-is negative. *)
+(** The function [Z_to_option_N] converts an integer [x] into a natural number
+by giving [None] in case [x] is negative. *)
 Definition Z_to_option_N (x : Z) : option N :=
   match x with
   | Z0 => Some N0
@@ -43,40 +111,66 @@ by yielding one if [P] holds and zero if [P] does not. *)
 Definition Z_decide (P : Prop) {dec : Decision P} : Z :=
   (if dec then 1 else 0)%Z.
 
-(** The function [Z_decide_rel] is the more efficient variant of [Z_decide] for
-decidable binary relations. It yields one if [R x y] and zero if not [R x y]. *)
+(** The function [Z_decide_rel] is the more efficient variant of [Z_decide] when
+used for binary relations. It yields one if [R x y] and zero if not [R x y]. *)
 Definition Z_decide_rel {A B} (R : A → B → Prop)
     {dec : ∀ x y, Decision (R x y)} (x : A) (y : B) : Z :=
   (if dec x y then 1 else 0)%Z.
 
-(** * Fresh binary naturals *)
-(** Given a finite set of binary naturals [N], we generate a fresh element by
-taking the maximum, and adding one to it. *)
-Definition Nmax `{Elements N C} : C → N := collection_fold Nmax 0%N.
-
-Instance Nmax_proper `{FinCollection N C} : Proper ((≡) ==> (=)) Nmax.
+(** Some correspondence lemmas between [nat] and [N] that are not part of the
+standard library. We declare a hint database [natify] to rewrite a goal
+involving [N] into a corresponding variant involving [nat]. *)
+Lemma N_to_nat_lt x y : N.to_nat x < N.to_nat y ↔ (x < y)%N.
+Proof. by rewrite <-N.compare_lt_iff, nat_compare_lt, N2Nat.inj_compare. Qed.
+Lemma N_to_nat_le x y : N.to_nat x ≤ N.to_nat y ↔ (x ≤ y)%N.
+Proof. by rewrite <-N.compare_le_iff, nat_compare_le, N2Nat.inj_compare. Qed.
+Lemma N_to_nat_0 : N.to_nat 0 = 0.
+Proof. done. Qed.
+Lemma N_to_nat_1 : N.to_nat 1 = 1.
+Proof. done. Qed.
+Lemma N_to_nat_div x y : N.to_nat (x `div` y) = N.to_nat x `div` N.to_nat y.
 Proof.
-  apply collection_fold_proper. intros.
-  rewrite !N.max_assoc. f_equal. apply N.max_comm.
+  destruct (decide (y = 0%N)).
+  { subst. by destruct x. }
+  apply NPeano.Nat.div_unique with (N.to_nat (x `mod` y)).
+  { by apply N_to_nat_lt, N.mod_lt. }
+  rewrite (N.div_unique_exact (x * y) y x), N.div_mul by lia.
+  by rewrite <-N2Nat.inj_mul, <-N2Nat.inj_add, <-N.div_mod.
 Qed.
-
-Lemma Nmax_max `{FinCollection N C} X x : x ∈ X → (x ≤ Nmax X)%N.
+(* We have [x `mod` 0 = 0] on [nat], and [x `mod` 0 = x] on [N]. *)
+Lemma N_to_nat_mod x y :
+  y ≠ 0%N →
+  N.to_nat (x `mod` y) = N.to_nat x `mod` N.to_nat y.
 Proof.
-  change ((λ b X, x ∈ X → (x ≤ b)%N) (collection_fold N.max 0%N X) X).
-  apply collection_fold_ind.
-  * solve_proper.
-  * solve_elem_of.
-  * solve_elem_of (eauto using N.le_max_l, N.le_max_r, N.le_trans).
+  intros.
+  apply NPeano.Nat.mod_unique with (N.to_nat (x `div` y)).
+  { by apply N_to_nat_lt, N.mod_lt. }
+  rewrite (N.div_unique_exact (x * y) y x), N.div_mul by lia.
+  by rewrite <-N2Nat.inj_mul, <-N2Nat.inj_add, <-N.div_mod.
 Qed.
 
-Instance Nfresh `{FinCollection N C} : Fresh N C := λ l, (1 + Nmax l)%N.
-Instance Nfresh_spec `{FinCollection N C} : FreshSpec N C.
-Proof.
-  split.
-  * intros. unfold fresh, Nfresh.
-    setoid_replace X with Y; [easy |].
-    now apply elem_of_equiv.
-  * intros X E. assert (1 ≤ 0)%N as []; [| easy].
-    apply N.add_le_mono_r with (Nmax X).
-    now apply Nmax_max.
-Qed.
+Hint Rewrite <-N2Nat.inj_iff : natify.
+Hint Rewrite <-N_to_nat_lt : natify.
+Hint Rewrite <-N_to_nat_le : natify.
+Hint Rewrite Nat2N.id : natify.
+Hint Rewrite N2Nat.inj_add : natify.
+Hint Rewrite N2Nat.inj_mul : natify.
+Hint Rewrite N2Nat.inj_sub : natify.
+Hint Rewrite N2Nat.inj_succ : natify.
+Hint Rewrite N2Nat.inj_pred : natify.
+Hint Rewrite N_to_nat_div : natify.
+Hint Rewrite N_to_nat_0 : natify.
+Hint Rewrite N_to_nat_1 : natify.
+Ltac natify := repeat autorewrite with natify in *.
+
+Hint Extern 100 (Nlt _ _) => natify : natify.
+Hint Extern 100 (Nle _ _) => natify : natify.
+Hint Extern 100 (@eq N _ _) => natify : natify.
+Hint Extern 100 (lt _ _) => natify : natify.
+Hint Extern 100 (le _ _) => natify : natify.
+Hint Extern 100 (@eq nat _ _) => natify : natify.
+
+Instance: ∀ x, PropHolds (0 < x)%N → PropHolds (0 < N.to_nat x).
+Proof. unfold PropHolds. intros. by natify. Qed.
+Instance: ∀ x, PropHolds (0 ≤ x)%N → PropHolds (0 ≤ N.to_nat x).
+Proof. unfold PropHolds. intros. by natify. Qed.
diff --git a/theories/option.v b/theories/option.v
index 09beb718e6e3b4ee55705cf6c0c0f5388af42dbb..455d1a73ea7f084e84f3700a8936ab6d8a0c1434 100644
--- a/theories/option.v
+++ b/theories/option.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects general purpose definitions and theorems on the option
 data type that are not in the Coq standard library. *)
-Require Export base tactics decidable orders.
+Require Export base tactics decidable.
 
 (** * General definitions and theorems *)
 (** Basic properties about equality. *)
@@ -22,9 +22,9 @@ Definition option_case {A B} (f : A → B) (b : B) (x : option A) : B :=
   | Some a => f a
   end.
 
-(** The [maybe] function allows us to get the value out of the option type
-by specifying a default value. *)
-Definition maybe {A} (a : A) (x : option A) : A :=
+(** The [from_option] function allows us to get the value out of the option
+type by specifying a default value. *)
+Definition from_option {A} (a : A) (x : option A) : A :=
   match x with
   | None => a
   | Some b => b
@@ -36,45 +36,61 @@ Lemma option_eq {A} (x y : option A) :
   x = y ↔ ∀ a, x = Some a ↔ y = Some a.
 Proof.
   split.
-  { intros. now subst. }
+  { intros. by subst. }
   intros E. destruct x, y.
-  + now apply E.
-  + symmetry. now apply E.
-  + now apply E.
-  + easy.
+  + by apply E.
+  + symmetry. by apply E.
+  + by apply E.
+  + done.
 Qed.
 
-(** We define [is_Some] as a [sig] instead of a [sigT] as extraction of
-witnesses can be derived (see [is_Some_sigT] below). *)
-Definition is_Some `(x : option A) : Prop := ∃ a, x = Some a.
-Hint Extern 10 (is_Some _) => solve [eexists; eauto].
+Inductive is_Some {A} : option A → Prop :=
+  make_is_Some x : is_Some (Some x).
 
-Ltac simplify_is_Some := repeat intro; repeat
+Lemma make_is_Some_alt `(x : option A) a : x = Some a → is_Some x.
+Proof. intros. by subst. Qed.
+Hint Resolve make_is_Some_alt.
+Lemma is_Some_None {A} : ¬is_Some (@None A).
+Proof. by inversion 1. Qed.
+Hint Resolve is_Some_None.
+
+Lemma is_Some_alt `(x : option A) : is_Some x ↔ ∃ y, x = Some y.
+Proof. split. inversion 1; eauto. intros [??]. by subst. Qed.
+
+Ltac inv_is_Some := repeat
   match goal with
-  | _ => progress simplify_equality
-  | H : is_Some _ |- _ => destruct H as [??]
-  | |- is_Some _ => eauto
+  | H : is_Some _ |- _ => inversion H; clear H; subst
   end.
 
-Lemma Some_is_Some `(a : A) : is_Some (Some a).
-Proof. simplify_is_Some. Qed.
-Lemma None_not_is_Some {A} : ¬is_Some (@None A).
-Proof. simplify_is_Some. Qed.
-
-Definition is_Some_sigT `(x : option A) : is_Some x → { a | x = Some a } :=
+Definition is_Some_proj `{x : option A} : is_Some x → A :=
+  match x with
+  | Some a => λ _, a
+  | None => False_rect _ ∘ is_Some_None
+  end.
+Definition Some_dec `(x : option A) : { a | x = Some a } + { x = None } :=
+  match x return { a | x = Some a } + { x = None } with
+  | Some a => inleft (a ↾ eq_refl _)
+  | None => inright eq_refl
+  end.
+Instance is_Some_dec `(x : option A) : Decision (is_Some x) :=
   match x with
-  | None => False_rect _ ∘ ex_ind None_ne_Some
-  | Some a => λ _, a↾eq_refl
+  | Some x => left (make_is_Some x)
+  | None => right is_Some_None
+  end.
+Instance None_dec `(x : option A) : Decision (x = None) :=
+  match x with
+  | Some x => right (Some_ne_None x)
+  | None => left eq_refl
   end.
-Lemma eq_Some_is_Some `(x : option A) a : x = Some a → is_Some x.
-Proof. simplify_is_Some. Qed.
 
 Lemma eq_None_not_Some `(x : option A) : x = None ↔ ¬is_Some x.
-Proof. destruct x; simpl; firstorder congruence. Qed.
+Proof. split. by destruct 2. destruct x. by intros []. done. 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.
 
 Lemma make_eq_Some {A} (x : option A) a :
   is_Some x → (∀ b, x = Some b → b = a) → x = Some a.
-Proof. intros [??] H. subst. f_equal. auto. Qed.
+Proof. destruct 1. intros. f_equal. auto. Qed.
 
 (** Equality on [option] is decidable. *)
 Instance option_eq_dec `{dec : ∀ x y : A, Decision (x = y)}
@@ -83,17 +99,17 @@ Instance option_eq_dec `{dec : ∀ x y : A, Decision (x = y)}
   | Some a =>
     match y with
     | Some b =>
-      match dec a b with
-      | left H => left (f_equal _ H)
-      | right H => right (H ∘ injective Some _ _)
-      end
+       match dec a b with
+       | left H => left (f_equal _ H)
+       | right H => right (H ∘ injective Some _ _)
+       end
     | None => right (Some_ne_None _)
     end
   | None =>
-    match y with
-    | Some _ => right (None_ne_Some _)
-    | None => left eq_refl
-    end
+     match y with
+     | Some _ => right (None_ne_Some _)
+     | None => left eq_refl
+     end
   end.
 
 (** * Monadic operations *)
@@ -109,96 +125,142 @@ Instance option_join: MJoin option := λ A x,
   | None => None
   end.
 Instance option_fmap: FMap option := @option_map.
+Instance option_guard: MGuard option := λ P dec A x,
+  if dec then x else None.
 
 Lemma option_fmap_is_Some {A B} (f : A → B) (x : option A) :
   is_Some x ↔ is_Some (f <$> x).
-Proof. destruct x; split; intros [??]; subst; compute; eauto; discriminate. Qed.
+Proof. split; inversion 1. done. by destruct x. Qed.
 Lemma option_fmap_is_None {A B} (f : A → B) (x : option A) :
   x = None ↔ f <$> x = None.
-Proof. unfold fmap, option_fmap. destruct x; simpl; split; congruence. Qed.
+Proof. unfold fmap, option_fmap. by destruct x. Qed.
+
+Lemma option_bind_assoc {A B C} (f : A → option B)
+    (g : B → option C) (x : option A) : (x ≫= f) ≫= g = x ≫= (mbind g ∘ f).
+Proof. by destruct x; simpl. Qed.
 
-Ltac simplify_option_bind := repeat
+Tactic Notation "simplify_option_equality" "by" tactic3(tac) := repeat
   match goal with
-  | |- context C [mbind (M:=option) ?f None] =>
-    let X := (context C [ None ]) in change X
-  | H : context C [mbind (M:=option) ?f None] |- _ =>
-    let X := (context C [ None ]) in change X in H
-  | |- context C [mbind (M:=option) ?f (Some ?a)] =>
-    let X := (context C [ f a ]) in
-    let X' := eval simpl in X in change X'
-  | H : context C [mbind (M:=option) ?f (Some ?a)] |- _ =>
-    let X := context C [ f a ] in
-    let X' := eval simpl in X in change X' in H
-  | _ => progress simplify_equality
+  | _ => first [progress simpl in * | progress simplify_equality]
+  | H : context [mbind (M:=option) (A:=?A) ?f ?o] |- _ =>
+    let Hx := fresh in
+    first
+      [ let x := fresh in evar (x:A);
+        let x' := eval unfold x in x in clear x;
+        assert (o = Some x') as Hx by tac
+      | assert (o = None) as Hx by tac ];
+    rewrite Hx in H; clear Hx
+  | H : context [fmap (M:=option) (A:=?A) ?f ?o] |- _ =>
+    let Hx := fresh in
+    first
+      [ let x := fresh in evar (x:A);
+        let x' := eval unfold x in x in clear x;
+        assert (o = Some x') as Hx by tac
+      | assert (o = None) as Hx by tac ];
+    rewrite Hx in H; clear Hx
+  | H : context [ match ?o with _ => _ end ] |- _ =>
+    match type of o with
+    | option ?A =>
+      let Hx := fresh in
+      first
+        [ let x := fresh in evar (x:A);
+          let x' := eval unfold x in x in clear x;
+          assert (o = Some x') as Hx by tac
+        | assert (o = None) as Hx by tac ];
+      rewrite Hx in H; clear Hx
+    end
   | H : mbind (M:=option) ?f ?o = ?x |- _ =>
+    match o with Some _ => fail 1 | None => fail 1 | _ => idtac end;
+    match x with Some _ => idtac | None => idtac | _ => fail 1 end;
     destruct o eqn:?
-  | H : context [ ?o = _ ] |- mbind (M:=option) ?f ?o = ?x =>
-    erewrite H by eauto
+  | H : ?x = mbind (M:=option) ?f ?o |- _ =>
+    match o with Some _ => fail 1 | None => fail 1 | _ => idtac end;
+    match x with Some _ => idtac | None => idtac | _ => fail 1 end;
+    destruct o eqn:?
+  | H : fmap (M:=option) ?f ?o = ?x |- _ =>
+    match o with Some _ => fail 1 | None => fail 1 | _ => idtac end;
+    match x with Some _ => idtac | None => idtac | _ => fail 1 end;
+    destruct o eqn:?
+  | H : ?x = fmap (M:=option) ?f ?o |- _ =>
+    match o with Some _ => fail 1 | None => fail 1 | _ => idtac end;
+    match x with Some _ => idtac | None => idtac | _ => fail 1 end;
+    destruct o eqn:?
+  | |- context [mbind (M:=option) (A:=?A) ?f ?o] =>
+    let Hx := fresh in
+    first
+      [ let x := fresh in evar (x:A);
+        let x' := eval unfold x in x in clear x;
+        assert (o = Some x') as Hx by tac
+      | assert (o = None) as Hx by tac ];
+    rewrite Hx; clear Hx
+  | |- context [fmap (M:=option) (A:=?A) ?f ?o] =>
+    let Hx := fresh in
+    first
+      [ let x := fresh in evar (x:A);
+        let x' := eval unfold x in x in clear x;
+        assert (o = Some x') as Hx by tac
+      | assert (o = None) as Hx by tac ];
+    rewrite Hx; clear Hx
+  | |- context [ match ?o with _ => _ end ] =>
+    match type of o with
+    | option ?A =>
+      let Hx := fresh in
+      first
+        [ let x := fresh in evar (x:A);
+          let x' := eval unfold x in x in clear x;
+          assert (o = Some x') as Hx by tac
+        | assert (o = None) as Hx by tac ];
+      rewrite Hx; clear Hx
+    end
+  | H : context C [@mguard option _ ?P ?dec _ ?x] |- _ =>
+    let X := context C [ if dec then x else None ] in
+    change X in H; destruct dec
+  | |- context C [@mguard option _ ?P ?dec _ ?x] =>
+    let X := context C [ if dec then x else None ] in
+    change X; destruct dec
+  | H1 : ?o = Some ?x, H2 : ?o = Some ?y |- _ =>
+    assert (y = x) by congruence; clear H2
+  | H1 : ?o = Some ?x, H2 : ?o = None |- _ =>
+    congruence
   end.
+Tactic Notation "simplify_option_equality" :=
+  simplify_option_equality by eauto.
+
+Hint Extern 100 => simplify_option_equality : simplify_option_equality.
 
 (** * Union, intersection and difference *)
-Instance option_union: UnionWith option := λ A f x y,
+Instance option_union_with {A} : UnionWith A (option A) := λ f x y,
   match x, y with
-  | Some a, Some b => Some (f a b)
+  | Some a, Some b => f a b
   | Some a, None => Some a
   | None, Some b => Some b
   | None, None => None
   end.
-Instance option_intersection: IntersectionWith option := λ A f x y,
+Instance option_intersection_with {A} :
+    IntersectionWith A (option A) := λ f x y,
   match x, y with
-  | Some a, Some b => Some (f a b)
+  | Some a, Some b => f a b
   | _, _ => None
   end.
-Instance option_difference: DifferenceWith option := λ A f x y,
+Instance option_difference_with {A} :
+    DifferenceWith A (option A) := λ f x y,
   match x, y with
   | Some a, Some b => f a b
   | Some a, None => Some a
   | None, _ => None
   end.
 
-Section option_union_intersection.
-  Context {A} (f : A → A → A).
+Section option_union_intersection_difference.
+  Context {A} (f : A → A → option A).
 
   Global Instance: LeftId (=) None (union_with f).
-  Proof. now intros [?|]. Qed.
+  Proof. by intros [?|]. Qed.
   Global Instance: RightId (=) None (union_with f).
-  Proof. now intros [?|]. Qed.
+  Proof. by intros [?|]. Qed.
   Global Instance: Commutative (=) f → Commutative (=) (union_with f).
-  Proof.
-    intros ? [?|] [?|]; compute; try reflexivity.
-    now rewrite (commutative f).
-  Qed.
-  Global Instance: Associative (=) f → Associative (=) (union_with f).
-  Proof.
-    intros ? [?|] [?|] [?|]; compute; try reflexivity.
-    now rewrite (associative f).
-  Qed.
-  Global Instance: Idempotent (=) f → Idempotent (=) (union_with f).
-  Proof.
-    intros ? [?|]; compute; try reflexivity.
-    now rewrite (idempotent f).
-  Qed.
-
+  Proof. by intros ? [?|] [?|]; compute; rewrite 1?(commutative f). Qed.
   Global Instance: Commutative (=) f → Commutative (=) (intersection_with f).
-  Proof.
-    intros ? [?|] [?|]; compute; try reflexivity.
-    now rewrite (commutative f).
-  Qed.
-  Global Instance: Associative (=) f → Associative (=) (intersection_with f).
-  Proof.
-    intros ? [?|] [?|] [?|]; compute; try reflexivity.
-    now rewrite (associative f).
-  Qed.
-  Global Instance: Idempotent (=) f → Idempotent (=) (intersection_with f).
-  Proof.
-    intros ? [?|]; compute; try reflexivity.
-    now rewrite (idempotent f).
-  Qed.
-End option_union_intersection.
-
-Section option_difference.
-  Context {A} (f : A → A → option A).
-
+  Proof. by intros ? [?|] [?|]; compute; rewrite 1?(commutative f). Qed.
   Global Instance: RightId (=) None (difference_with f).
-  Proof. now intros [?|]. Qed.
-End option_difference.
+  Proof. by intros [?|]. Qed.
+End option_union_intersection_difference.
diff --git a/theories/orders.v b/theories/orders.v
index a763f908cf51a4743b8c16ae50386c59f2ff74ee..89c4554122d4f25206b541d0573c1f960c5d5b92 100644
--- a/theories/orders.v
+++ b/theories/orders.v
@@ -2,7 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects common properties of pre-orders and semi lattices. This
 theory will mainly be used for the theory on collections and finite maps. *)
-Require Export base.
+Require Import SetoidList.
+Require Export base decidable tactics list.
 
 (** * Pre-orders *)
 (** We extend a pre-order to a partial order by defining equality as
@@ -26,8 +27,50 @@ Section preorder.
     * transitivity x1. tauto. transitivity x2; tauto.
     * transitivity y1. tauto. transitivity y2; tauto.
   Qed.
+
+  Global Instance preorder_subset: Subset A := λ X Y, X ⊆ Y ∧ Y ⊈ X.
+  Lemma subset_spec X Y : X ⊂ Y ↔ X ⊆ Y ∧ Y ⊈ X.
+  Proof. done. Qed.
+
+  Lemma subset_subseteq X Y : X ⊂ Y → X ⊆ Y.
+  Proof. by intros [? _]. Qed.
+  Lemma subset_trans_l X Y Z : X ⊂ Y → Y ⊆ Z → X ⊂ Z.
+  Proof.
+    intros [? Hxy] ?. split.
+    * by transitivity Y.
+    * contradict Hxy. by transitivity Z.
+  Qed.
+  Lemma subset_trans_r X Y Z : X ⊆ Y → Y ⊂ Z → X ⊂ Z.
+  Proof.
+    intros ? [? Hyz]. split.
+    * by transitivity Y.
+    * contradict Hyz. by transitivity X.
+  Qed.
+
+  Global Instance: StrictOrder (⊂).
+  Proof.
+    split.
+    * firstorder.
+    * eauto using subset_trans_r, subset_subseteq.
+  Qed.
+  Global Instance: Proper ((≡) ==> (≡) ==> iff) (⊂).
+  Proof. unfold subset, preorder_subset. solve_proper. Qed.
+
+  Context `{∀ X Y : A, Decision (X ⊆ Y)}.
+  Global Instance preorder_equiv_dec_slow (X Y : A) :
+    Decision (X ≡ Y) | 100 := _.
+  Global Instance preorder_subset_dec_slow (X Y : A) :
+    Decision (X ⊂ Y) | 100 := _.
+
+  Lemma subseteq_inv X Y : X ⊆ Y → X ⊂ Y ∨ X ≡ Y.
+  Proof.
+    destruct (decide (Y ⊆ X)).
+    * by right.
+    * by left.
+  Qed.
 End preorder.
 
+Typeclasses Opaque preorder_equiv.
 Hint Extern 0 (@Equivalence _ (≡)) =>
   class_apply preorder_equivalence : typeclass_instances.
 
@@ -47,7 +90,7 @@ Section bounded_join_sl.
   Lemma union_compat x1 x2 y1 y2 : x1 ⊆ x2 → y1 ⊆ y2 → x1 ∪ y1 ⊆ x2 ∪ y2.
   Proof. auto. Qed.
   Lemma union_empty x : x ∪ ∅ ⊆ x.
-  Proof. apply union_least. easy. auto. Qed.
+  Proof. by apply union_least. Qed.
   Lemma union_comm x y : x ∪ y ⊆ y ∪ x.
   Proof. auto. Qed.
   Lemma union_assoc_1 x y z : (x ∪ y) ∪ z ⊆ x ∪ (y ∪ z).
@@ -55,7 +98,7 @@ Section bounded_join_sl.
   Lemma union_assoc_2 x y z : x ∪ (y ∪ z) ⊆ (x ∪ y) ∪ z.
   Proof. auto. Qed.
 
-  Global Instance: Proper ((≡) ==> (≡) ==> (≡)) (∪).
+  Global Instance union_proper: Proper ((≡) ==> (≡) ==> (≡)) (∪).
   Proof.
     unfold equiv, preorder_equiv. split; apply union_compat; simpl in *; tauto.
   Qed.
@@ -79,6 +122,33 @@ Section bounded_join_sl.
 
   Lemma equiv_empty X : X ⊆ ∅ → X ≡ ∅.
   Proof. split; eauto. Qed.
+
+  Global Instance: Proper (eqlistA (≡) ==> (≡)) union_list.
+  Proof.
+    induction 1; simpl.
+    * done.
+    * by apply union_proper.
+  Qed.
+
+  Lemma empty_union X Y : X ∪ Y ≡ ∅ ↔ X ≡ ∅ ∧ Y ≡ ∅.
+  Proof.
+    split.
+    * intros E. split; apply equiv_empty;
+        by transitivity (X ∪ Y); [auto | rewrite E].
+    * intros [E1 E2]. by rewrite E1, E2, (left_id _ _).
+  Qed.
+  Lemma empty_list_union Xs : ⋃ Xs ≡ ∅ ↔ Forall (≡ ∅) Xs.
+  Proof.
+    split.
+    * induction Xs; simpl; rewrite ?empty_union; intuition.
+    * induction 1 as [|?? E1 ? E2]; simpl. done. by apply empty_union.
+  Qed.
+
+  Context `{∀ X Y : A, Decision (X ⊆ Y)}.
+  Lemma non_empty_union X Y : X ∪ Y ≢ ∅ → X ≢ ∅ ∨ Y ≢ ∅.
+  Proof. rewrite empty_union. destruct (decide (X ≡ ∅)); intuition. Qed.
+  Lemma non_empty_list_union Xs : ⋃ Xs ≢ ∅ → Exists (≢ ∅) Xs.
+  Proof. rewrite empty_list_union. apply (not_Forall_Exists _). Qed.
 End bounded_join_sl.
 
 (** * Meet semi lattices *)
@@ -95,7 +165,8 @@ Section meet_sl.
   Proof. intros. transitivity x1; auto. Qed.
   Hint Resolve intersection_compat_l intersection_compat_r.
 
-  Lemma intersection_compat x1 x2 y1 y2 : x1 ⊆ x2 → y1 ⊆ y2 → x1 ∩ y1 ⊆ x2 ∩ y2.
+  Lemma intersection_compat x1 x2 y1 y2 :
+    x1 ⊆ x2 → y1 ⊆ y2 → x1 ∩ y1 ⊆ x2 ∩ y2.
   Proof. auto. Qed.
   Lemma intersection_comm x y : x ∩ y ⊆ y ∩ x.
   Proof. auto. Qed.
@@ -123,3 +194,17 @@ Section meet_sl.
   Lemma subseteq_intersection_2 X Y : X ∩ Y ≡ X → X ⊆ Y.
   Proof. apply subseteq_intersection. Qed.
 End meet_sl.
+
+(** * Lower bounded lattices *)
+Section lower_bounded_lattice.
+  Context `{LowerBoundedLattice A}.
+
+  Global Instance: LeftAbsorb (≡) ∅ (∩).
+  Proof.
+    split.
+    * by apply subseteq_intersection_l.
+    * by apply subseteq_empty.
+  Qed.
+  Global Instance: RightAbsorb (≡) ∅ (∩).
+  Proof. intros ?. by rewrite (commutative _), (left_absorb _ _). Qed.
+End lower_bounded_lattice.
diff --git a/theories/pmap.v b/theories/pmap.v
index b8a35bb77cfc5d82fd4dfdfb51941b3ce7da163d..b7c04de4839d8bc00ff95769746e3d68059a237f 100644
--- a/theories/pmap.v
+++ b/theories/pmap.v
@@ -14,13 +14,6 @@ Local Open Scope positive_scope.
 Local Hint Extern 0 (@eq positive _ _) => congruence.
 Local Hint Extern 0 (¬@eq positive _ _) => congruence.
 
-(** Enable unfolding of the finite map operations in this file. *)
-Local Arguments lookup _ _ _ _ _ !_ /.
-Local Arguments partial_alter _ _ _ _ _ _ !_ /.
-Local Arguments fmap _ _ _ _ _ !_ /.
-Local Arguments merge _ _ _ _ !_ _ /.
-Local Arguments mbind _ _ _ _ _ !_/.
-
 (** * The tree data structure *)
 (** The data type [Pmap_raw] specifies radix-2 search trees. These trees do
 not ensure canonical representations of maps. For example the empty map can
@@ -46,10 +39,9 @@ Local Hint Constructors Pmap_ne.
 Instance Pmap_ne_dec `(t : Pmap_raw A) : Decision (Pmap_ne t).
 Proof.
   red. induction t as [|? IHl [?|] ? IHr].
-  * right. now inversion 1.
-  * now intuition.
-  * destruct IHl, IHr; try solve [left; auto]; right;
-      inversion_clear 1; contradiction.
+  * right. by inversion 1.
+  * intuition.
+  * destruct IHl, IHr; try (by left; auto); right; by inversion_clear 1.
 Qed.
 
 (** The following predicate describes well well formed trees. A tree is well
@@ -66,9 +58,10 @@ Proof.
   red. induction t as [|l IHl [?|] r IHr]; simpl.
   * intuition.
   * destruct IHl, IHr; try solve [left; intuition];
-      right; inversion_clear 1; intuition.
+      right; by inversion_clear 1.
   * destruct IHl, IHr, (decide (Pmap_ne l)), (decide (Pmap_ne r));
-      try solve [left; intuition]; right; inversion_clear 1; intuition.
+      try solve [left; intuition];
+      right; inversion_clear 1; intuition.
 Qed.
 
 (** Now we restrict the data type of trees to those that are well formed. *)
@@ -84,30 +77,30 @@ Global Instance Pmap_dec `{∀ x y : A, Decision (x = y)} (t1 t2 : Pmap A) :
 
 Instance Pempty_raw {A} : Empty (Pmap_raw A) := Pleaf.
 Global Instance Pempty {A} : Empty (Pmap A) :=
-  (∅ : Pmap_raw A)↾bool_decide_pack _ Pmap_wf_leaf.
+  (∅ : Pmap_raw A) ↾ bool_decide_pack _ Pmap_wf_leaf.
 
-Instance Plookup_raw: Lookup positive Pmap_raw :=
-  fix Plookup_raw A (i : positive) (t : Pmap_raw A) {struct t} : option A :=
+Instance Plookup_raw {A} : Lookup positive A (Pmap_raw A) :=
+  fix Plookup_raw (i : positive) (t : Pmap_raw A) {struct t} : option A :=
   match t with
   | Pleaf => None
   | Pnode l o r =>
-    match i with
-    | 1 => o
-    | i~0 => @lookup _ _ Plookup_raw _ i l
-    | i~1 => @lookup _ _ Plookup_raw _ i r
-    end
+     match i with
+     | 1 => o
+     | i~0 => @lookup _ _ _ Plookup_raw i l
+     | i~1 => @lookup _ _ _ Plookup_raw i r
+     end
   end.
-Global Instance Plookup: Lookup positive Pmap := λ A i t, `t !! i.
+Instance Plookup {A} : Lookup positive A (Pmap A) := λ i t, `t !! i.
 
 Lemma Plookup_raw_empty {A} i : (∅ : Pmap_raw A) !! i = None.
-Proof. now destruct i. Qed.
+Proof. by destruct i. Qed.
 
 Lemma Pmap_ne_lookup `(t : Pmap_raw A) : Pmap_ne t → ∃ i x, t !! i = Some x.
 Proof.
   induction 1 as [? x ?| l r ? IHl | l r ? IHr].
-  * intros. now exists 1 x.
-  * destruct IHl as [i [x ?]]. now exists (i~0) x.
-  * destruct IHr as [i [x ?]]. now exists (i~1) x.
+  * intros. by exists 1 x.
+  * destruct IHl as [i [x ?]]. by exists (i~0) x.
+  * destruct IHr as [i [x ?]]. by exists (i~1) x.
 Qed.
 
 Lemma Pmap_wf_eq_get {A} (t1 t2 : Pmap_raw A) :
@@ -116,29 +109,29 @@ Proof.
   intros t1wf. revert t2.
   induction t1wf as [| ? x ? ? IHl ? IHr | l r ? IHl ? IHr Hne1 ].
   * destruct 1 as [| | ???? [?|?]]; intros Hget.
-    + easy.
+    + done.
     + discriminate (Hget 1).
-    + destruct (Pmap_ne_lookup l) as [i [??]]; auto.
+    + destruct (Pmap_ne_lookup l) as [i [??]]; trivial.
       specialize (Hget (i~0)). simpl in *. congruence.
-    + destruct (Pmap_ne_lookup r) as [i [??]]; auto.
+    + destruct (Pmap_ne_lookup r) as [i [??]]; trivial.
       specialize (Hget (i~1)). simpl in *. congruence.
   * destruct 1; intros Hget.
     + discriminate (Hget xH).
     + f_equal.
-      - apply IHl; auto. intros i. now apply (Hget (i~0)).
-      - now apply (Hget 1).
-      - apply IHr; auto. intros i. now apply (Hget (i~1)).
+      - apply IHl; trivial. intros i. apply (Hget (i~0)).
+      - apply (Hget 1).
+      - apply IHr; trivial. intros i. apply (Hget (i~1)).
     + specialize (Hget 1). simpl in *. congruence.
   * destruct 1; intros Hget.
     + destruct Hne1.
-      destruct (Pmap_ne_lookup l) as [i [??]]; auto.
+      destruct (Pmap_ne_lookup l) as [i [??]]; trivial.
       - specialize (Hget (i~0)). simpl in *. congruence.
-      - destruct (Pmap_ne_lookup r) as [i [??]]; auto.
+      - destruct (Pmap_ne_lookup r) as [i [??]]; trivial.
         specialize (Hget (i~1)). simpl in *. congruence.
     + specialize (Hget 1). simpl in *. congruence.
     + f_equal.
-      - apply IHl; auto. intros i. now apply (Hget (i~0)).
-      - apply IHr; auto. intros i. now apply (Hget (i~1)).
+      - apply IHl; trivial. intros i. apply (Hget (i~0)).
+      - apply IHr; trivial. intros i. apply (Hget (i~1)).
 Qed.
 
 Fixpoint Psingleton_raw {A} (i : positive) (x : A) : Pmap_raw A :=
@@ -156,7 +149,7 @@ Proof. induction i; simpl; intuition. Qed.
 Local Hint Resolve Psingleton_raw_wf.
 
 Lemma Plookup_raw_singleton {A} i (x : A) : Psingleton_raw i x !! i = Some x.
-Proof. now induction i. Qed.
+Proof. by induction i. Qed.
 Lemma Plookup_raw_singleton_ne {A} i j (x : A) :
   i ≠ j → Psingleton_raw i x !! j = None.
 Proof. revert j. induction i; intros [?|?|]; simpl; auto. congruence. Qed.
@@ -174,33 +167,32 @@ Local Hint Resolve Pnode_canon_wf.
 
 Lemma Pnode_canon_lookup_xH `(l : Pmap_raw A) o (r : Pmap_raw A) :
   Pnode_canon l o r !! 1 = o.
-Proof. now destruct l,o,r. Qed.
+Proof. by destruct l,o,r. Qed.
 Lemma Pnode_canon_lookup_xO `(l : Pmap_raw A) o (r : Pmap_raw A) i :
   Pnode_canon l o r !! i~0 = l !! i.
-Proof. now destruct l,o,r. Qed.
+Proof. by destruct l,o,r. Qed.
 Lemma Pnode_canon_lookup_xI `(l : Pmap_raw A) o (r : Pmap_raw A) i :
   Pnode_canon l o r !! i~1 = r !! i.
-Proof. now destruct l,o,r. Qed.
+Proof. by destruct l,o,r. Qed.
 Ltac Pnode_canon_rewrite := repeat (
   rewrite Pnode_canon_lookup_xH ||
   rewrite Pnode_canon_lookup_xO ||
   rewrite Pnode_canon_lookup_xI).
 
-Instance Ppartial_alter_raw: PartialAlter positive Pmap_raw :=
-  fix Ppartial_alter_raw A (f : option A → option A) (i : positive)
-    (t : Pmap_raw A) {struct t} : Pmap_raw A :=
+Instance Ppartial_alter_raw {A} : PartialAlter positive A (Pmap_raw A) :=
+  fix go f i t {struct t} : Pmap_raw A :=
   match t with
   | Pleaf =>
-    match f None with
-    | None => Pleaf
-    | Some x => Psingleton_raw i x
-    end
+     match f None with
+     | None => Pleaf
+     | Some x => Psingleton_raw i x
+     end
   | Pnode l o r =>
-    match i with
-    | 1 => Pnode_canon l (f o) r
-    | i~0 => Pnode_canon (@partial_alter _ _ Ppartial_alter_raw _ f i l) o r
-    | i~1 => Pnode_canon l o (@partial_alter _ _ Ppartial_alter_raw _ f i r)
-    end
+     match i with
+     | 1 => Pnode_canon l (f o) r
+     | i~0 => Pnode_canon (@partial_alter _ _ _ go f i l) o r
+     | i~1 => Pnode_canon l o (@partial_alter _ _ _ go f i r)
+     end
   end.
 
 Lemma Ppartial_alter_raw_wf {A} f i (t : Pmap_raw A) :
@@ -212,34 +204,42 @@ Proof.
   * intros [?|?|]; simpl; intuition.
 Qed.
 
-Global Instance Ppartial_alter: PartialAlter positive Pmap := λ A f i t,
+Instance Ppartial_alter {A} : PartialAlter positive A (Pmap A) := λ f i t,
   dexist (partial_alter f i (`t)) (Ppartial_alter_raw_wf f i _ (proj2_dsig t)).
 
 Lemma Plookup_raw_alter {A} f i (t : Pmap_raw A) :
   partial_alter f i t !! i = f (t !! i).
 Proof.
   revert i. induction t.
-  * simpl. case (f None).
-    + intros. now apply Plookup_raw_singleton.
-    + now destruct i.
-  * intros [?|?|]; simpl; Pnode_canon_rewrite; auto.
+  * intros i. change (
+     match f None with
+     | Some x => Psingleton_raw i x
+     | None => Pleaf
+     end !! i = f None). destruct (f None).
+    + intros. apply Plookup_raw_singleton.
+    + by destruct i.
+  * intros [?|?|]; simpl; by Pnode_canon_rewrite.
 Qed.
 Lemma Plookup_raw_alter_ne {A} f i j (t : Pmap_raw A) :
   i ≠ j → partial_alter f i t !! j = t !! j.
 Proof.
   revert i j. induction t as [|l IHl ? r IHr].
-  * simpl. intros. case (f None).
-    + intros. now apply Plookup_raw_singleton_ne.
-    + easy.
+  * intros. change (
+     match f None with
+     | Some x => Psingleton_raw i x
+     | None => Pleaf
+     end !! j = None). destruct (f None).
+    + intros. by apply Plookup_raw_singleton_ne.
+    + done.
   * intros [?|?|] [?|?|]; simpl; Pnode_canon_rewrite; auto; congruence.
 Qed.
 
-Instance Pfmap_raw: FMap Pmap_raw :=
-  fix Pfmap_raw A B f (t : Pmap_raw A) : Pmap_raw B :=
+Instance Pfmap_raw {A B} (f : A → B) : FMapD Pmap_raw f :=
+  fix go (t : Pmap_raw A) : Pmap_raw B :=
+  let _ : FMapD Pmap_raw f := @go in
   match t with
   | Pleaf => Pleaf
-  | Pnode l x r =>
-    Pnode (@fmap _ Pfmap_raw _ _ f l) (fmap f x) (@fmap _ Pfmap_raw _ _ f r)
+  | Pnode l x r => Pnode (f <$> l) (f <$> x) (f <$> r)
   end.
 
 Lemma Pfmap_raw_ne `(f : A → B) (t : Pmap_raw A) :
@@ -248,41 +248,72 @@ Proof. induction 1; simpl; auto. Qed.
 Local Hint Resolve Pfmap_raw_ne.
 Lemma Pfmap_raw_wf `(f : A → B) (t : Pmap_raw A) :
   Pmap_wf t → Pmap_wf (fmap f t).
-Proof. induction 1; simpl; intuition auto. Qed.
+Proof. induction 1; simpl; intuition. Qed.
 
-Global Instance Pfmap: FMap Pmap := λ A B f t,
+Global Instance Pfmap {A B} (f : A → B) : FMapD Pmap f := λ t,
   dexist _ (Pfmap_raw_wf f _ (proj2_dsig t)).
 
 Lemma Plookup_raw_fmap `(f : A → B) (t : Pmap_raw A) i :
   fmap f t !! i = fmap f (t !! i).
-Proof. revert i. induction t. easy. intros [?|?|]; simpl; auto. Qed.
-
-(** The [dom] function is rather inefficient, but since we do not intend to
-use it for computation it does not really matter to us. *)
-Section dom.
-  Context `{Collection B D}.
-
-  Fixpoint Pdom_raw (f : positive → B) `(t : Pmap_raw A) : D :=
-    match t with
-    | Pleaf => ∅
-    | Pnode l o r => option_case (λ _, {[ f 1 ]}) ∅ o
-                       ∪ Pdom_raw (f ∘ (~0)) l ∪ Pdom_raw (f ∘ (~1)) r
-    end.
-
-  Lemma Plookup_raw_dom f `(t : Pmap_raw A) x :
-    x ∈ Pdom_raw f t ↔ ∃ i, x = f i ∧ is_Some (t !! i).
-  Proof.
-    split.
-    * revert f. induction t as [|? IHl [?|] ? IHr]; esolve_elem_of.
-    * intros [i [? Hlookup]]; subst. revert f i Hlookup.
-      induction t as [|? IHl [?|] ? IHr]; intros f [?|?|];
-        solve_elem_of (now apply (IHl (f ∘ (~0)))
-        || now apply (IHr (f ∘ (~1))) || simplify_is_Some).
-  Qed.
-End dom.
-
-Global Instance Pdom : Dom positive Pmap := λ C _ _ _ _ t,
-  Pdom_raw id (`t).
+Proof. revert i. induction t. done. by intros [?|?|]; simpl. Qed.
+
+(** The [finmap_to_list] function is rather inefficient, but since we do not
+use it for computation at this moment, we do not care. *)
+Fixpoint Pto_list_raw {A} (t : Pmap_raw A) : list (positive * A) :=
+  match t with
+  | Pleaf => []
+  | Pnode l o r =>
+     option_case (λ x, [(1,x)]) [] o ++ 
+       (fst_map (~0) <$> Pto_list_raw l) ++ 
+       (fst_map (~1) <$> Pto_list_raw r)
+  end.
+
+Lemma Pto_list_raw_nodup {A} (t : Pmap_raw A) : NoDup (Pto_list_raw t).
+Proof.
+  induction t as [|? IHl [?|] ? IHr]; simpl.
+  * constructor.
+  * rewrite NoDup_cons, NoDup_app, !(NoDup_fmap _).
+    repeat split; trivial.
+    + rewrite elem_of_app, !elem_of_list_fmap.
+      intros [[[??] [??]] | [[??] [??]]]; simpl in *; simplify_equality.
+    + intro. rewrite !elem_of_list_fmap.
+      intros [[??] [??]] [[??] [??]]; simpl in *; simplify_equality.
+  * rewrite NoDup_app, !(NoDup_fmap _).
+    repeat split; trivial.
+    intro. rewrite !elem_of_list_fmap.
+    intros [[??] [??]] [[??] [??]]; simpl in *; simplify_equality.
+Qed.
+
+Lemma Pelem_of_to_list_raw {A} (t : Pmap_raw A) i x :
+  (i,x) ∈ Pto_list_raw t ↔ t !! i = Some x.
+Proof.
+  split.
+  * revert i. induction t as [|? IHl [?|] ? IHr]; intros i; simpl.
+    + by rewrite ?elem_of_nil.
+    + rewrite elem_of_cons, !elem_of_app, !elem_of_list_fmap.
+      intros [? | [[[??] [??]] | [[??] [??]]]]; naive_solver.
+    + rewrite !elem_of_app, !elem_of_list_fmap.
+      intros [[[??] [??]] | [[??] [??]]]; naive_solver.
+  * revert i. induction t as [|? IHl [?|] ? IHr]; simpl.
+    + done.
+    + intros i. rewrite elem_of_cons, elem_of_app.
+      destruct i as [i|i|]; simpl.
+      - right. right. rewrite elem_of_list_fmap.
+        exists (i,x); simpl; auto.
+      - right. left. rewrite elem_of_list_fmap.
+        exists (i,x); simpl; auto.
+      - left. congruence.
+    + intros i. rewrite elem_of_app.
+      destruct i as [i|i|]; simpl.
+      - right. rewrite elem_of_list_fmap.
+        exists (i,x); simpl; auto.
+      - left. rewrite elem_of_list_fmap.
+        exists (i,x); simpl; auto.
+      - done.
+Qed.
+
+Global Instance Pto_list {A} : FinMapToList positive A (Pmap A) :=
+  λ t, Pto_list_raw (`t).
 
 Fixpoint Pmerge_aux `(f : option A → option B) (t : Pmap_raw A) : Pmap_raw B :=
   match t with
@@ -299,25 +330,25 @@ Lemma Pmerge_aux_spec `(f : option A → option B) (Hf : f None = None)
     (t : Pmap_raw A) i :
   Pmerge_aux f t !! i = f (t !! i).
 Proof.
-  revert i. induction t as [| l IHl o r IHr ]; [easy |].
+  revert i. induction t as [| l IHl o r IHr ]; [done |].
   intros [?|?|]; simpl; Pnode_canon_rewrite; auto.
 Qed.
 
-Global Instance Pmerge_raw: Merge Pmap_raw :=
-  fix Pmerge_raw A f (t1 t2 : Pmap_raw A) : Pmap_raw A :=
+Global Instance Pmerge_raw {A} : Merge A (Pmap_raw A) :=
+  fix Pmerge_raw f (t1 t2 : Pmap_raw A) : Pmap_raw A :=
   match t1, t2 with
   | Pleaf, t2 => Pmerge_aux (f None) t2
   | t1, Pleaf => Pmerge_aux (flip f None) t1
   | Pnode l1 o1 r1, Pnode l2 o2 r2 =>
-     Pnode_canon (@merge _ Pmerge_raw _ f l1 l2)
-      (f o1 o2) (@merge _ Pmerge_raw _ f r1 r2)
+     Pnode_canon (@merge _ _ Pmerge_raw f l1 l2)
+      (f o1 o2) (@merge _ _ Pmerge_raw f r1 r2)
   end.
 Local Arguments Pmerge_aux _ _ _ _ : simpl never.
 
 Lemma Pmerge_wf {A} f (t1 t2 : Pmap_raw A) :
   Pmap_wf t1 → Pmap_wf t2 → Pmap_wf (merge f t1 t2).
 Proof. intros t1wf. revert t2. induction t1wf; destruct 1; simpl; auto. Qed.
-Global Instance Pmerge: Merge Pmap := λ A f t1 t2,
+Global Instance Pmerge {A} : Merge A (Pmap A) := λ f t1 t2,
   dexist (merge f (`t1) (`t2))
     (Pmerge_wf _ _ _ (proj2_dsig t1) (proj2_dsig t2)).
 
@@ -325,24 +356,22 @@ Lemma Pmerge_raw_spec {A} f (Hf : f None None = None) (t1 t2 : Pmap_raw A) i :
   merge f t1 t2 !! i = f (t1 !! i) (t2 !! i).
 Proof.
   revert t2 i. induction t1 as [| l1 IHl1 o1 r1 IHr1 ].
-  * simpl. now apply Pmerge_aux_spec.
+  * intros. unfold merge. simpl. by rewrite Pmerge_aux_spec.
   * destruct t2 as [| l2 o2 r2 ].
-    + unfold merge, Pmerge_raw. intros. now rewrite Pmerge_aux_spec.
-    + intros [?|?|]; simpl; Pnode_canon_rewrite; auto.
+    + unfold merge, Pmerge_raw. intros. by rewrite Pmerge_aux_spec.
+    + intros [?|?|]; simpl; by Pnode_canon_rewrite.
 Qed.
 
 Global Instance: FinMap positive Pmap.
 Proof.
   split.
   * intros ? [t1?] [t2?]. intros. apply dsig_eq. simpl.
-    apply Pmap_wf_eq_get; auto; now apply (bool_decide_unpack _).
-  * now destruct i.
-  * intros ?? [??] ?. now apply Plookup_raw_alter.
-  * intros ?? [??] ??. now apply Plookup_raw_alter_ne.
-  * intros ??? [??]. now apply Plookup_raw_fmap.
-  * intros ?????????? [??] i. unfold dom, Pdom.
-    rewrite Plookup_raw_dom. unfold id. split.
-    + intros [? [??]]. now subst.
-    + firstorder.
-  * intros ??? [??] [??] ?. now apply Pmerge_raw_spec.
+    apply Pmap_wf_eq_get; trivial; by apply (bool_decide_unpack _).
+  * by destruct i.
+  * intros ?? [??] ?. by apply Plookup_raw_alter.
+  * intros ?? [??] ??. by apply Plookup_raw_alter_ne.
+  * intros ??? [??]. by apply Plookup_raw_fmap.
+  * intros ? [??]. apply Pto_list_raw_nodup.
+  * intros ? [??]. apply Pelem_of_to_list_raw.
+  * intros ??? [??] [??] ?. by apply Pmerge_raw_spec.
 Qed.
diff --git a/theories/prelude.v b/theories/prelude.v
index 39beaa8d556a2dc5aee8b828094681bb250bbb2a..aa488b804e93688a7f929247d2ae4960f0ab0392 100644
--- a/theories/prelude.v
+++ b/theories/prelude.v
@@ -6,9 +6,11 @@ Require Export
   decidable
   orders
   option
-  list
+  vector
+  numbers
+  ars
   collections
   fin_collections
   listset
-  subset
-  numbers.
+  fresh_numbers
+  list.
diff --git a/theories/subset.v b/theories/subset.v
deleted file mode 100644
index 00a45e9586b9e8f867a3e743db06d33fcd0af5b9..0000000000000000000000000000000000000000
--- a/theories/subset.v
+++ /dev/null
@@ -1,19 +0,0 @@
-(* Copyright (c) 2012, Robbert Krebbers. *)
-(* This file is distributed under the terms of the BSD license. *)
-(** This file implements the operations on non computable subsets [A → Prop]
-over some carrier [A]. *)
-Require Export base.
-
-Definition subset A := A → Prop.
-
-Instance subset_elem_of {A} : ElemOf A (subset A) | 100 := λ x P, P x.
-Instance subset_empty {A} : Empty (subset A) := λ _, False.
-Instance subset_singleton {A} : Singleton A (subset A) := (=).
-Instance subset_union {A} : Union (subset A) := λ P Q x, P x ∨ Q x.
-Instance subset_intersection {A} : Intersection (subset A) := λ P Q x,
-  P x ∧ Q x.
-Instance subset_difference {A} : Difference (subset A) := λ P Q x, P x ∧ ¬Q x.
-Instance subset_map {A} : Map A (subset A) := λ f P x, ∃ y, f y = x ∧ P y.
-
-Instance subset_container: Collection A (subset A) | 100.
-Proof. firstorder try congruence. Qed.
diff --git a/theories/tactics.v b/theories/tactics.v
index ea2b9ce066648b293e5edf80d1a2fe2c66cfc75d..b802ee654d80d1bb4ccb177c717ac2180fa6cba9 100644
--- a/theories/tactics.v
+++ b/theories/tactics.v
@@ -2,20 +2,128 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects some general purpose tactics that are used throughout
 the development. *)
+Require Export Psatz.
 Require Export base.
 
+(** We declare hint databases [lia] and [congruence] containing solely the
+following hints. These hint database are useful in combination with another
+hint database [db] that contain lemmas with premises that should be solved by
+[lia] or [congruence]. One can now just say [auto with db,lia]. *)
+Hint Extern 1000 => lia : lia.
+Hint Extern 1000 => congruence : congruence.
+
+(** The tactic [intuition] expands to [intuition auto with *] by default. This
+is rather efficient when having big hint databases, or expensive [Hint Extern]
+declarations as the above. *)
+Tactic Notation "intuition" := intuition auto.
+
+(** A slightly modified version of Ssreflect's finishing tactic [done]. It
+also performs [reflexivity] and does not compute the goal's [hnf] so as to
+avoid unfolding setoid equalities. Note that this tactic performs much better
+than Coq's [easy] as it does not perform [inversion]. *)
+Ltac done :=
+  trivial; intros; solve
+    [ repeat first
+      [ solve [trivial]
+      | solve [symmetry; trivial]
+      | reflexivity
+      | discriminate
+      | contradiction
+      | split ]
+    | match goal with
+      H : ¬_ |- _ => solve [destruct H; trivial]
+      end ].
+Tactic Notation "by" tactic(tac) :=
+  tac; done.
+
+Ltac case_match :=
+  match goal with
+  | H : context [ match ?x with _ => _ end ] |- _ => destruct x eqn:?
+  | |- context [ match ?x with _ => _ end ] => destruct x eqn:?
+  end.
+
+(** The tactic [clear dependent H1 ... Hn] clears the hypotheses [Hi] and
+their dependencies. *)
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) :=
+  clear dependent H1; clear dependent H2.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) :=
+  clear dependent H1 H2; clear dependent H3.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4) :=
+  clear dependent H1 H2 H3; clear dependent H4.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4)
+  hyp(H5) := clear dependent H1 H2 H3 H4; clear dependent H5.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4) hyp(H5)
+  hyp (H6) := clear dependent H1 H2 H3 H4 H5; clear dependent H6.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4) hyp(H5)
+  hyp (H6) hyp(H7) := clear dependent H1 H2 H3 H4 H5 H6; clear dependent H7.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4) hyp(H5)
+  hyp (H6) hyp(H7) hyp(H8) :=
+  clear dependent H1 H2 H3 H4 H5 H6 H7; clear dependent H8.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4) hyp(H5)
+  hyp (H6) hyp(H7) hyp(H8) hyp(H9) :=
+  clear dependent H1 H2 H3 H4 H5 H6 H7 H8; clear dependent H9.
+Tactic Notation "clear" "dependent" hyp(H1) hyp(H2) hyp(H3) hyp(H4) hyp(H5)
+  hyp (H6) hyp(H7) hyp(H8) hyp(H9) hyp(H10) :=
+  clear dependent H1 H2 H3 H4 H5 H6 H7 H8 H9; clear dependent H10.
+
+(** The tactic [first_of tac1 tac2] calls [tac1] and then calls [tac2] on the
+first subgoal generated by [tac1] *)
+Tactic Notation "first_of" tactic3(tac1) "by" tactic3(tac2) :=
+     (tac1; [ tac2 ])
+  || (tac1; [ tac2 |])
+  || (tac1; [ tac2 | | ])
+  || (tac1; [ tac2 | | | ])
+  || (tac1; [ tac2 | | | | ])
+  || (tac1; [ tac2 | | | | | ])
+  || (tac1; [ tac2 | | | | | |])
+  || (tac1; [ tac2 | | | | | | |])
+  || (tac1; [ tac2 | | | | | | | |])
+  || (tac1; [ tac2 | | | | | | | | |])
+  || (tac1; [ tac2 | | | | | | | | | |])
+  || (tac1; [ tac2 | | | | | | | | | | |])
+  || (tac1; [ tac2 | | | | | | | | | | | |]).
+
+(** The tactic [is_non_dependent H] determines whether the goal's conclusion or
+assumptions depend on [H]. *)
+Tactic Notation "is_non_dependent" constr(H) :=
+  match goal with
+  | _ : context [ H ] |- _ => fail 1
+  | |- context [ H ] => fail 1
+  | _ => idtac
+  end.
+
+(* The tactic [var_eq x y] fails if [x] and [y] are unequal. *)
+Ltac var_eq x1 x2 := match x1 with x2 => idtac | _ => fail 1 end.
+Ltac var_neq x1 x2 := match x1 with x2 => fail 1 | _ => idtac end.
+
+Tactic Notation "repeat_on_hyps" tactic3(tac) :=
+  repeat match goal with H : _ |- _ => progress tac H end.
+
+Ltac block_hyps := repeat_on_hyps (fun H =>
+  match type of H with
+  | block _ => idtac
+  | ?T => change (block T) in H
+  end).
+Ltac unblock_hyps := unfold block in * |-.
+
+(** The tactic [injection' H] is a variant of injection that introduces the
+generated equalities. *)
+Ltac injection' H :=
+  block_goal; injection H; clear H; intros; unblock_goal.
+
 (** The tactic [simplify_equality] repeatedly substitutes, discriminates,
 and injects equalities, and tries to contradict impossible inequalities. *)
 Ltac simplify_equality := repeat
   match goal with
-  | |- _ => progress subst
-  | |- _ = _ => reflexivity
-  | H : _ ≠ _ |- _ => now destruct H
-  | H : _ = _ → False |- _ => now destruct H
+  | H : _ ≠ _ |- _ => by destruct H
+  | H : _ = _ → False |- _ => by destruct H
+  | H : ?x = _ |- _ => subst x
+  | H : _ = ?x |- _ => subst x
   | H : _ = _ |- _ => discriminate H
-  | H : _ = _ |-  ?G =>
-    (* avoid introducing additional equalities *)
-    change (id G); injection H; clear H; intros; unfold id at 1
+  | H : ?f _ = ?f _ |- _ => apply (injective f) in H
+    (* before [injection'] to circumvent bug #2939 in some situations *)
+  | H : _ = _ |- _ => injection' H
+  | H : ?x = ?x |- _ => clear H
   end.
 
 (** Coq's default [remember] tactic does have an option to name the generated
@@ -26,13 +134,12 @@ Tactic Notation "remember" constr(t) "as" "(" ident(x) "," ident(E) ")" :=
   | E' : x = _ |- _ => rename E' into E
   end.
 
-(** Given a list [l], the tactic [map tac l] runs [tac x] for each element [x]
-of the list [l]. It will succeed for the first element [x] of [l] for which 
-[tac x] succeeds. *)
-Tactic Notation "map" tactic(tac) tactic(l) :=
+(** 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. *)
+Tactic Notation "iter" tactic(tac) tactic(l) :=
   let rec go l :=
   match l with
-  | nil => idtac
   | ?x :: ?l => tac x || go l
   end in go l.
 
@@ -97,15 +204,6 @@ Tactic Notation "feed" "destruct" constr(H) :=
 Tactic Notation "feed" "destruct" constr(H) "as" simple_intropattern(IP) :=
   feed (fun p => let H':=fresh in pose proof p as H'; destruct H' as IP) H.
 
-(** The tactic [is_non_dependent H] determines whether the goal's conclusion or
-assumptions depend on [H]. *)
-Tactic Notation "is_non_dependent" constr(H) :=
-  match goal with
-  | _ : context [ H ] |- _ => fail 1
-  | |- context [ H ] => fail 1
-  | _ => idtac
-  end.
-
 (** Coq's [firstorder] tactic fails or loops on rather small goals already. In 
 particular, on those generated by the tactic [unfold_elem_ofs] to solve
 propositions on collections. The [naive_solver] tactic implements an ad-hoc
@@ -134,9 +232,7 @@ Tactic Notation "naive_solver" tactic(tac) :=
   | |- _ => progress simpl in *
   | |- _ => progress simplify_equality
   (**i solve the goal *)
-  | |- _ => eassumption
-  | |- _ => now constructor
-  | |- _ => now symmetry
+  | |- _ => solve [ eassumption | symmetry; eassumption | reflexivity ]
   (**i operations that generate more subgoals *)
   | |- _ ∧ _ => split
   | H : _ ∨ _ |- _ => destruct H
diff --git a/theories/vector.v b/theories/vector.v
new file mode 100644
index 0000000000000000000000000000000000000000..2051432340b439b00300c29546038048ccfb41c7
--- /dev/null
+++ b/theories/vector.v
@@ -0,0 +1,342 @@
+(* Copyright (c) 2012, Robbert Krebbers. *)
+(* This file is distributed under the terms of the BSD license. *)
+(** This file collects general purpose definitions and theorems on vectors
+(lists of fixed length) and the fin type (bounded naturals). It uses the
+definitions from the standard library, but renames or changes their notations,
+so that it becomes more consistent with the naming conventions in this
+development. *)
+Require Import list.
+Open Scope vector_scope.
+
+(** * The fin type *)
+(** The type [fin n] represents natural numbers [i] with [0 ≤ i < n]. We
+define a scope [fin], in which we declare notations for small literals of the
+[fin] type. Whereas the standard library starts counting at [1], we start
+counting at [0]. This way, the embedding [fin_to_nat] preserves [0], and allows
+us to define [fin_to_nat] as a coercion without introducing notational
+ambiguity. *)
+Notation fin := Fin.t.
+Notation FS := Fin.FS.
+
+Delimit Scope fin_scope with fin.
+Arguments Fin.FS _ _%fin.
+
+Notation "0" := Fin.F1 : fin_scope.
+Notation "1" := (FS 0) : fin_scope.
+Notation "2" := (FS 1) : fin_scope.
+Notation "3" := (FS 2) : fin_scope.
+Notation "4" := (FS 3) : fin_scope.
+Notation "5" := (FS 4) : fin_scope.
+Notation "6" := (FS 5) : fin_scope.
+Notation "7" := (FS 6) : fin_scope.
+Notation "8" := (FS 7) : fin_scope.
+Notation "9" := (FS 8) : fin_scope.
+Notation "10" := (FS 9) : fin_scope.
+
+Fixpoint fin_to_nat {n} (i : fin n) : nat :=
+  match i with
+  | 0%fin => 0
+  | FS _ i => S (fin_to_nat i)
+  end.
+Coercion fin_to_nat : fin >-> nat.
+
+Notation fin_rect2 := Fin.rect2.
+Notation FS_inj := Fin.FS_inj.
+
+Instance fin_dec {n} : ∀ i j : fin n, Decision (i = j).
+Proof.
+ refine (fin_rect2
+  (λ n (i j : fin n), { i = j } + { i ≠ j })
+  (λ _, left _)
+  (λ _ _, right _)
+  (λ _ _, right _)
+  (λ _ _ _ H, cast_if H));
+  abstract (f_equal; by auto using FS_inj).
+Defined.
+
+(** The inversion principle [fin_S_inv] is more convenient than its variant
+[Fin.caseS] in the standard library, as we keep the parameter [n] fixed.
+In the tactic [inv_fin i] to perform dependent case analysis on [i], we
+therefore do not have to generalize over the index [n] and all assumptions
+depending on it. Notice that contrary to [dependent destruction], which uses
+the [JMeq_eq] axiom, the tactic [inv_fin] produces axiom free proofs.*)
+Notation fin_0_inv := Fin.case0.
+
+Definition fin_S_inv {n} (P : fin (S n) → Type)
+  (H0 : P 0%fin) (HS : ∀ i, P (FS i)) (i : fin (S n)) : P i.
+Proof.
+ revert P H0 HS. refine (
+  match i with
+  | 0%fin => λ _ H0 _, H0
+  | FS _ i => λ _ _ HS, HS i
+  end).
+Defined.
+
+Ltac inv_fin i :=
+  match type of i with
+  | fin 0 =>
+    revert dependent i;
+    match goal with
+    |- ∀ i, @?P i => apply (fin_0_inv P)
+    end
+  | fin (S ?n) =>
+    revert dependent i;
+    match goal with
+    |- ∀ i, @?P i => apply (fin_S_inv P)
+    end
+  end.
+
+(** * Vectors *)
+(** The type [vec n] represents lists of consisting of exactly [n] elements.
+Whereas the standard library declares exactly the same notations for vectors as
+used for lists, we use slightly different notations so it becomes easier to use
+lists and vectors together. *)
+Notation vec := Vector.t.
+Notation vnil := Vector.nil.
+Arguments vnil {_}.
+Notation vcons := Vector.cons.
+Notation vapp := Vector.append.
+Arguments vcons {_} _ {_} _.
+
+Infix ":::" := vcons (at level 60, right associativity) : vector_scope.
+Notation "[# ] " := vnil : vector_scope.
+Notation "[# x ] " := (vcons x vnil) : vector_scope.
+Notation "[# x ; .. ; y ] " := (vcons x .. (vcons y vnil) ..) : vector_scope.
+
+Infix "+++" := vapp (at level 60, right associativity) : vector_scope.
+
+(** Notice that we cannot define [Vector.nth] as an instance of our [Lookup]
+type class, as it has a dependent type. *)
+Arguments Vector.nth {_ _} !_ !_%fin /.
+Infix "!!!" := Vector.nth (at level 20) : vector_scope.
+
+(** The tactic [vec_double_ind v1 v2] performs double induction on [v1] and [v2]
+provided that they have the same length. *)
+Notation vec_rect2 := Vector.rect2.
+Ltac vec_double_ind v1 v2 :=
+  match type of v1 with
+  | vec _ ?n =>
+    repeat match goal with
+    | H' : context [ n ] |- _ =>
+      var_neq v1 H'; var_neq v2 H'; revert H'
+    end;
+    revert n v1 v2;
+    match goal with
+    | |- ∀ n v1 v2, @?P n v1 v2 => apply (vec_rect2 P)
+    end
+  end.
+
+Notation vcons_inj := VectorSpec.cons_inj.
+Lemma vcons_inj_1 {A n} x y (v w : vec A n) : x ::: v = y ::: w → x = y.
+Proof. apply vcons_inj. Qed.
+Lemma vcons_inj_2 {A n} x y (v w : vec A n) : x ::: v = y ::: w → v = w.
+Proof. apply vcons_inj. Qed.
+
+Instance vec_dec {A} {dec : ∀ x y : A, Decision (x = y)} {n} :
+  ∀ v w : vec A n, Decision (v = w).
+Proof.
+ refine (vec_rect2
+  (λ n (v w : vec A n), { v = w } + { v ≠ w })
+  (left _)
+  (λ _ _ _ H x y, cast_if_and (dec x y) H));
+  f_equal; eauto using vcons_inj_1, vcons_inj_2.
+Defined.
+
+(** Similar to [fin], we provide an inversion principle that keeps the length
+fixed. We define a tactic [inv_vec v] to perform case analysis on [v], using
+this inversion principle. *)
+Notation vec_0_inv := Vector.case0.
+Definition vec_S_inv {A n} (P : vec A (S n) → Type)
+  (Hcons : ∀ x v, P (x ::: v)) v : P v.
+Proof.
+ revert P Hcons. refine (
+  match v with
+  | [#] => tt
+  | x ::: v => λ P Hcons, Hcons x v
+  end).
+Defined.
+
+Ltac inv_vec v :=
+  match type of v with
+  | vec _ 0 =>
+    revert dependent v;
+    match goal with
+    |- ∀ v, @?P v => apply (vec_0_inv P)
+    end
+  | vec _ (S ?n) =>
+    revert dependent v;
+    match goal with
+    |- ∀ v, @?P v => apply (vec_S_inv P)
+    end
+  end.
+
+(** The following tactic performs case analysis on all hypotheses of the shape
+[fin 0], [fin (S n)], [vec A 0] and [vec A (S n)] until no further case
+analyses are possible. *)
+Ltac inv_all_vec_fin :=
+  block_goal;
+  repeat match goal with
+  | v : vec _ _ |- _ => inv_vec v; intros
+  | i : fin _ |- _ => inv_fin i; intros
+  end;
+  unblock_goal.
+
+(** We define a coercion from [vec] to [list] and show that it preserves the
+operations on vectors. We also define a function to go in the other way, but
+do not define it as a coercion, as it would otherwise introduce ambiguity. *)
+Fixpoint vec_to_list {A n} (v : vec A n) : list A :=
+  match v with
+  | [#] => []
+  | x ::: v => x :: vec_to_list v
+  end.
+Coercion vec_to_list : vec >-> list.
+Notation list_to_vec := Vector.of_list.
+
+Lemma vec_to_list_cons {A n} x (v : vec A n) :
+  vec_to_list (x ::: v) = x :: vec_to_list v.
+Proof. done. Qed.
+Lemma vec_to_list_app {A n m} (v : vec A n) (w : vec A m) :
+  vec_to_list (v +++ w) = vec_to_list v ++ vec_to_list w.
+Proof. by induction v; simpl; f_equal. Qed.
+
+Lemma vec_to_list_of_list {A} (l : list A): vec_to_list (list_to_vec l) = l.
+Proof. by induction l; simpl; f_equal. Qed.
+
+Lemma vec_to_list_length {A n} (v : vec A n) : length (vec_to_list v) = n.
+Proof. induction v; simpl; by f_equal. Qed.
+
+Lemma vec_to_list_inj1 {A n m} (v : vec A n) (w : vec A m) :
+  vec_to_list v = vec_to_list w → n = m.
+Proof.
+  revert m w. induction v; intros ? [|???] ?;
+   simpl in *; simplify_equality; f_equal; eauto.
+Qed.
+Lemma vec_to_list_inj2 {A n} (v : vec A n) (w : vec A n) :
+  vec_to_list v = vec_to_list w → v = w.
+Proof.
+  revert w. induction v; intros w; inv_vec w; intros;
+    simpl in *; simplify_equality; f_equal; eauto.
+Qed.
+
+Lemma vlookup_middle {A n m} (v : vec A n) (w : vec A m) x :
+  ∃ i : fin (n + S m), x = (v +++ x ::: w) !!! i.
+Proof.
+  induction v; simpl.
+  * by eexists 0%fin.
+  * destruct IHv as [i ?]. by exists (FS i).
+Qed.
+
+Lemma vec_to_list_lookup_middle {A n} (v : vec A n) (l k : list A) x :
+  vec_to_list v = l ++ x :: k →
+    ∃ i : fin n, l = take i v ∧ x = v !!! i ∧ k = drop (S i) v.
+Proof.
+  intros H.
+  rewrite <-(vec_to_list_of_list l), <-(vec_to_list_of_list k) in H.
+  rewrite <-vec_to_list_cons, <-vec_to_list_app in H.
+  pose proof (vec_to_list_inj1 _ _ H); subst.
+  apply vec_to_list_inj2 in H; subst.
+  induction l. simpl.
+  * eexists 0%fin. simpl. by rewrite vec_to_list_of_list.
+  * destruct IHl as [i ?]. exists (FS i). simpl. intuition congruence.
+Qed.
+
+Lemma vec_to_list_drop_lookup {A n} (v : vec A n) (i : fin n) :
+  drop i v = v !!! i :: drop (S i) v.
+Proof. induction i; inv_vec v; simpl; intros; [done | by rewrite IHi]. Qed.
+
+Lemma vec_to_list_take_drop_lookup {A n} (v : vec A n) (i : fin n) :
+  vec_to_list v = take i v ++ v !!! i :: drop (S i) v.
+Proof.
+  rewrite <-(take_drop i v) at 1. f_equal.
+  apply vec_to_list_drop_lookup.
+Qed. 
+
+Lemma elem_of_vlookup {A n} (v : vec A n) x :
+  x ∈ vec_to_list v ↔ ∃ i, v !!! i = x.
+Proof.
+  split.
+  * induction v; simpl; [by rewrite elem_of_nil |].
+    inversion 1; subst.
+    + by eexists 0%fin.
+    + destruct IHv as [i ?]; trivial. by exists (FS i).
+  * intros [i ?]; subst. induction v; inv_fin i.
+    + by left.
+    + right. apply IHv.
+Qed.
+
+Lemma Forall_vlookup {A} (P : A → Prop) {n} (v : vec A n) :
+  Forall P (vec_to_list v) ↔ ∀ i, P (v !!! i).
+Proof.
+  rewrite Forall_forall.
+  setoid_rewrite elem_of_vlookup. naive_solver.
+Qed.
+Lemma Forall_vlookup_1 {A} (P : A → Prop) {n} (v : vec A n) i :
+  Forall P (vec_to_list v) → P (v !!! i).
+Proof. by rewrite Forall_vlookup. Qed.
+Lemma Forall_vlookup_2 {A} (P : A → Prop) {n} (v : vec A n) :
+  (∀ i, P (v !!! i)) → Forall P (vec_to_list v).
+Proof. by rewrite Forall_vlookup. Qed.
+
+Lemma Exists_vlookup {A} (P : A → Prop) {n} (v : vec A n) :
+  Exists P (vec_to_list v) ↔ ∃ i, P (v !!! i).
+Proof.
+  rewrite Exists_exists.
+  setoid_rewrite elem_of_vlookup. naive_solver.
+Qed.
+
+Lemma Forall2_vlookup {A B} (P : A → B → Prop)
+    {n} (v1 : vec A n) (v2 : vec B n) :
+  Forall2 P (vec_to_list v1) (vec_to_list v2) ↔
+    ∀ i, P (v1 !!! i) (v2 !!! i).
+Proof.
+  split.
+  * vec_double_ind v1 v2.
+    + intros _ i. inv_fin i.
+    + intros n v1 v2 IH a b; simpl. inversion_clear 1.
+      intros i. inv_fin i; simpl; auto.
+  * vec_double_ind v1 v2.
+    + constructor.
+    + intros ??? IH ?? H.
+      constructor. apply (H 0%fin). apply IH, (λ i, H (FS i)).
+Qed.
+
+(** The function [vmap f v] applies a function [f] element wise to [v]. *)
+Notation vmap := Vector.map.
+
+Lemma vlookup_map `(f : A → B) {n} (v : vec A n) i :
+  vmap f v !!! i = f (v !!! i).
+Proof. by apply Vector.nth_map. Qed.
+
+Lemma vec_to_list_map `(f : A → B) {n} (v : vec A n) :
+  vec_to_list (vmap f v) = f <$> vec_to_list v.
+Proof. induction v; simpl. done. by rewrite IHv. Qed.
+
+(** The function [vzip_with f v w] combines the vectors [v] and [w] element
+wise using the function [f]. *)
+Notation vzip_with := Vector.map2.
+
+Lemma vzip_with_lookup `(f : A → B → C) {n} (v1 : vec A n) (v2 : vec B n) i :
+  vzip_with f v1 v2 !!! i = f (v1 !!! i) (v2 !!! i).
+Proof. by apply Vector.nth_map2. Qed.
+
+(** Similar to vlookup, we cannot define [vinsert] as an instance of the
+[Insert] type class, as it has a dependent type. *)
+Fixpoint vinsert {A n} (i : fin n) (x : A) : vec A n → vec A n :=
+  match i with
+  | 0%fin => vec_S_inv _ (λ _ v, x ::: v)
+  | FS _ i => vec_S_inv _ (λ y v, y ::: vinsert i x v)
+  end.
+
+Lemma vec_to_list_insert {A n} i x (v : vec A n) :
+  vec_to_list (vinsert i x v) = insert (fin_to_nat i) x (vec_to_list v).
+Proof. induction v; inv_fin i. done. simpl. intros. by rewrite IHv. Qed.
+
+Lemma vlookup_insert {A n} i x (v : vec A n) : vinsert i x v !!! i = x.
+Proof. by induction i; inv_vec v. Qed.
+
+Lemma vlookup_insert_ne {A n} i j x (v : vec A n) :
+  i ≠ j → vinsert i x v !!! j = v !!! j.
+Proof.
+  induction i; inv_fin j; inv_vec v; simpl; try done.
+  intros. apply IHi. congruence.
+Qed.