Skip to content
Snippets Groups Projects

Compare revisions

Changes are shown as if the source revision was being merged into the target revision. Learn more about comparing revisions.

Source

Select target project
No results found

Target

Select target project
  • iris/stdpp
  • johannes/stdpp
  • proux1/stdpp
  • dosualdo/stdpp
  • benoit/coq-stdpp
  • dfrumin/coq-stdpp
  • haidang/stdpp
  • amintimany/coq-stdpp
  • swasey/coq-stdpp
  • simongregersen/stdpp
  • proux/stdpp
  • janno/coq-stdpp
  • amaurremi/coq-stdpp
  • msammler/stdpp
  • tchajed/stdpp
  • YaZko/stdpp
  • maximedenes/stdpp
  • jakobbotsch/stdpp
  • Blaisorblade/stdpp
  • simonspies/stdpp
  • lepigre/stdpp
  • devilhena/stdpp
  • simonfv/stdpp
  • jihgfee/stdpp
  • snyke7/stdpp
  • Armael/stdpp
  • gmalecha/stdpp
  • olaure01/stdpp
  • sarahzrf/stdpp
  • atrieu/stdpp
  • herbelin/stdpp
  • arthuraa/stdpp
  • lgaeher/stdpp
  • mrhaandi/stdpp
  • mattam82/stdpp
  • Quarkbeast/stdpp
  • aa755/stdpp
  • gmevel/stdpp
  • lstefane/stdpp
  • jung/stdpp
  • vsiles/stdpp
  • dlesbre/stdpp
  • bergwerf/stdpp
  • marijnvanwezel/stdpp
  • ivanbakel/stdpp
  • tperami/stdpp
  • adamAndMath/stdpp
  • Villetaneuse/stdpp
  • sanjit/stdpp
  • yiyunliu/stdpp
  • thomas-lamiaux/stdpp
  • Tragicus/stdpp
  • kbedarka/stdpp
53 results
Show changes
......@@ -14,11 +14,11 @@ Set Default Proof Using "Type*".
Inductive coGset `{Countable A} :=
| FinGSet (X : gset A)
| CoFinGset (X : gset A).
Arguments coGset _ {_ _} : assert.
Global Arguments coGset _ {_ _} : assert.
Instance coGset_eq_dec `{Countable A} : EqDecision (coGset A).
Global Instance coGset_eq_dec `{Countable A} : EqDecision (coGset A).
Proof. solve_decision. Defined.
Instance coGset_countable `{Countable A} : Countable (coGset A).
Global Instance coGset_countable `{Countable A} : Countable (coGset A).
Proof.
apply (inj_countable'
(λ X, match X with FinGSet X => inl X | CoFinGset X => inr X end)
......@@ -80,7 +80,7 @@ Section coGset.
Qed.
End coGset.
Instance coGset_elem_of_dec `{Countable A} : RelDecision (∈@{coGset A}) :=
Global Instance coGset_elem_of_dec `{Countable A} : RelDecision (∈@{coGset A}) :=
λ x X,
match X with
| FinGSet X => decide_rel elem_of x X
......@@ -92,7 +92,7 @@ Section infinite.
Global Instance coGset_leibniz : LeibnizEquiv (coGset A).
Proof.
intros [X|X] [Y|Y]; rewrite elem_of_equiv;
intros [X|X] [Y|Y]; rewrite set_equiv;
unfold elem_of, coGset_elem_of; simpl; intros HXY.
- f_equal. by apply leibniz_equiv.
- by destruct (exist_fresh (X Y)) as [? [? ?%HXY]%not_elem_of_union].
......@@ -138,7 +138,7 @@ End infinite.
Definition coGpick `{Countable A, Infinite A} (X : coGset A) : A :=
fresh (match X with FinGSet _ => | CoFinGset X => X end).
Lemma coGpick_elem_of `{Countable A, Infinite A} X :
Lemma coGpick_elem_of `{Countable A, Infinite A} (X : coGset A) :
¬set_finite X coGpick X X.
Proof.
unfold coGpick.
......@@ -192,15 +192,5 @@ Lemma elem_of_coGset_to_top_set `{Countable A, TopSet A C} X x :
x ∈@{C} coGset_to_top_set X x X.
Proof. destruct X; set_solver. Qed.
(** * Domain of finite maps *)
Instance coGset_dom `{Countable K} {A} : Dom (gmap K A) (coGset K) := λ m,
gset_to_coGset (dom _ m).
Instance coGset_dom_spec `{Countable K} : FinMapDom K (gmap K) (coGset K).
Proof.
split; try apply _. intros B m i. unfold dom, coGset_dom.
by rewrite elem_of_gset_to_coGset, elem_of_dom.
Qed.
Typeclasses Opaque coGset_elem_of coGset_empty coGset_top coGset_singleton.
Typeclasses Opaque coGset_union coGset_intersection coGset_difference.
Typeclasses Opaque coGset_dom.
Global Typeclasses Opaque coGset_elem_of coGset_empty coGset_top coGset_singleton.
Global Typeclasses Opaque coGset_union coGset_intersection coGset_difference.
......@@ -18,7 +18,7 @@ Local Open Scope positive_scope.
Inductive coPset_raw :=
| coPLeaf : bool coPset_raw
| coPNode : bool coPset_raw coPset_raw coPset_raw.
Instance coPset_raw_eq_dec : EqDecision coPset_raw.
Global Instance coPset_raw_eq_dec : EqDecision coPset_raw.
Proof. solve_decision. Defined.
Fixpoint coPset_wf (t : coPset_raw) : bool :=
......@@ -26,9 +26,16 @@ Fixpoint coPset_wf (t : coPset_raw) : bool :=
| coPLeaf _ => true
| coPNode true (coPLeaf true) (coPLeaf true) => false
| coPNode false (coPLeaf false) (coPLeaf false) => false
| coPNode b l r => coPset_wf l && coPset_wf r
| coPNode _ l r => coPset_wf l && coPset_wf r
end.
Arguments coPset_wf !_ / : simpl nomatch, assert.
Global Arguments coPset_wf !_ / : simpl nomatch, assert.
Lemma coPNode_wf b l r :
coPset_wf l coPset_wf r
(l = coPLeaf true r = coPLeaf true b = true False)
(l = coPLeaf false r = coPLeaf false b = false False)
coPset_wf (coPNode b l r).
Proof. destruct b, l as [[]|], r as [[]|]; naive_solver. Qed.
Lemma coPNode_wf_l b l r : coPset_wf (coPNode b l r) coPset_wf l.
Proof. destruct b, l as [[]|],r as [[]|]; simpl; rewrite ?andb_True; tauto. Qed.
......@@ -42,10 +49,10 @@ Definition coPNode' (b : bool) (l r : coPset_raw) : coPset_raw :=
| false, coPLeaf false, coPLeaf false => coPLeaf false
| _, _, _ => coPNode b l r
end.
Arguments coPNode' : simpl never.
Lemma coPNode_wf b l r : coPset_wf l coPset_wf r coPset_wf (coPNode' b l r).
Global Arguments coPNode' : simpl never.
Lemma coPNode'_wf b l r : coPset_wf l coPset_wf r coPset_wf (coPNode' b l r).
Proof. destruct b, l as [[]|], r as [[]|]; simpl; auto. Qed.
Hint Resolve coPNode_wf : core.
Global Hint Resolve coPNode'_wf : core.
Fixpoint coPset_elem_of_raw (p : positive) (t : coPset_raw) {struct t} : bool :=
match t, p with
......@@ -55,7 +62,7 @@ Fixpoint coPset_elem_of_raw (p : positive) (t : coPset_raw) {struct t} : bool :=
| coPNode _ _ r, p~1 => coPset_elem_of_raw p r
end.
Local Notation e_of := coPset_elem_of_raw.
Arguments coPset_elem_of_raw _ !_ / : simpl nomatch, assert.
Global Arguments coPset_elem_of_raw _ !_ / : simpl nomatch, assert.
Lemma coPset_elem_of_node b l r p :
e_of p (coPNode' b l r) = e_of p (coPNode b l r).
Proof. by destruct p, b, l as [[]|], r as [[]|]. Qed.
......@@ -87,7 +94,7 @@ Fixpoint coPset_singleton_raw (p : positive) : coPset_raw :=
| p~0 => coPNode' false (coPset_singleton_raw p) (coPLeaf false)
| p~1 => coPNode' false (coPLeaf false) (coPset_singleton_raw p)
end.
Instance coPset_union_raw : Union coPset_raw :=
Global Instance coPset_union_raw : Union coPset_raw :=
fix go t1 t2 := let _ : Union _ := @go in
match t1, t2 with
| coPLeaf false, coPLeaf false => coPLeaf false
......@@ -98,7 +105,7 @@ Instance coPset_union_raw : Union coPset_raw :=
| coPNode b1 l1 r1, coPNode b2 l2 r2 => coPNode' (b1||b2) (l1 l2) (r1 r2)
end.
Local Arguments union _ _!_ !_ / : assert.
Instance coPset_intersection_raw : Intersection coPset_raw :=
Global Instance coPset_intersection_raw : Intersection coPset_raw :=
fix go t1 t2 := let _ : Intersection _ := @go in
match t1, t2 with
| coPLeaf true, coPLeaf true => coPLeaf true
......@@ -152,22 +159,22 @@ Qed.
(** * Packed together + set operations *)
Definition coPset := { t | coPset_wf t }.
Instance coPset_singleton : Singleton positive coPset := λ p,
Global Instance coPset_singleton : Singleton positive coPset := λ p,
coPset_singleton_raw p coPset_singleton_wf _.
Instance coPset_elem_of : ElemOf positive coPset := λ p X, e_of p (`X).
Instance coPset_empty : Empty coPset := coPLeaf false I.
Instance coPset_top : Top coPset := coPLeaf true I.
Instance coPset_union : Union coPset := λ X Y,
Global Instance coPset_elem_of : ElemOf positive coPset := λ p X, e_of p (`X).
Global Instance coPset_empty : Empty coPset := coPLeaf false I.
Global Instance coPset_top : Top coPset := coPLeaf true I.
Global Instance coPset_union : Union coPset := λ X Y,
let (t1,Ht1) := X in let (t2,Ht2) := Y in
(t1 t2) coPset_union_wf _ _ Ht1 Ht2.
Instance coPset_intersection : Intersection coPset := λ X Y,
Global Instance coPset_intersection : Intersection coPset := λ X Y,
let (t1,Ht1) := X in let (t2,Ht2) := Y in
(t1 t2) coPset_intersection_wf _ _ Ht1 Ht2.
Instance coPset_difference : Difference coPset := λ X Y,
Global Instance coPset_difference : Difference coPset := λ X Y,
let (t1,Ht1) := X in let (t2,Ht2) := Y in
(t1 coPset_opp_raw t2) coPset_intersection_wf _ _ Ht1 (coPset_opp_wf _).
Instance coPset_top_set : TopSet positive coPset.
Global Instance coPset_top_set : TopSet positive coPset.
Proof.
split; [split; [split| |]|].
- by intros ??.
......@@ -184,25 +191,25 @@ Qed.
(** Iris and specifically [solve_ndisj] heavily rely on this hint. *)
Local Definition coPset_top_subseteq := top_subseteq (C:=coPset).
Hint Resolve coPset_top_subseteq : core.
Global Hint Resolve coPset_top_subseteq : core.
Instance coPset_leibniz : LeibnizEquiv coPset.
Global Instance coPset_leibniz : LeibnizEquiv coPset.
Proof.
intros X Y; rewrite elem_of_equiv; intros HXY.
intros X Y; rewrite set_equiv; intros HXY.
apply (sig_eq_pi _), coPset_eq; try apply @proj2_sig.
intros p; apply eq_bool_prop_intro, (HXY p).
Qed.
Instance coPset_elem_of_dec : RelDecision (∈@{coPset}).
Global Instance coPset_elem_of_dec : RelDecision (∈@{coPset}).
Proof. solve_decision. Defined.
Instance coPset_equiv_dec : RelDecision (≡@{coPset}).
Global Instance coPset_equiv_dec : RelDecision (≡@{coPset}).
Proof. refine (λ X Y, cast_if (decide (X = Y))); abstract (by fold_leibniz). Defined.
Instance mapset_disjoint_dec : RelDecision (##@{coPset}).
Global Instance mapset_disjoint_dec : RelDecision (##@{coPset}).
Proof.
refine (λ X Y, cast_if (decide (X Y = )));
abstract (by rewrite disjoint_intersection_L).
Defined.
Instance mapset_subseteq_dec : RelDecision (⊆@{coPset}).
Global Instance mapset_subseteq_dec : RelDecision (⊆@{coPset}).
Proof.
refine (λ X Y, cast_if (decide (X Y = Y))); abstract (by rewrite subseteq_union_L).
Defined.
......@@ -221,7 +228,7 @@ Proof.
unfold set_finite, elem_of at 1, coPset_elem_of; simpl; clear Ht; split.
- induction t as [b|b l IHl r IHr]; simpl.
{ destruct b; simpl; [intros [l Hl]|done].
by apply (is_fresh (list_to_set l : Pset)), elem_of_list_to_set, Hl. }
by apply (infinite_is_fresh l), Hl. }
intros [ll Hll]; rewrite andb_True; split.
+ apply IHl; exists (omap (maybe (~0)) ll); intros i.
rewrite elem_of_list_omap; intros; exists (i~0); auto.
......@@ -232,14 +239,17 @@ Proof.
exists ([1] ++ ((~0) <$> ll) ++ ((~1) <$> rl))%list; intros [i|i|]; simpl;
rewrite elem_of_cons, elem_of_app, !elem_of_list_fmap; naive_solver.
Qed.
Instance coPset_finite_dec (X : coPset) : Decision (set_finite X).
Global Instance coPset_finite_dec (X : coPset) : Decision (set_finite X).
Proof.
refine (cast_if (decide (coPset_finite (`X)))); by rewrite coPset_finite_spec.
Defined.
(** * Pick element from infinite sets *)
(* Implemented using depth-first search, which results in very unbalanced
trees. *)
(* The function [coPpick X] gives an element that is in the set [X], provided
that the set [X] is infinite. Note that [coPpick] function is implemented by
depth-first search, so using it repeatedly to obtain elements [x], and
inserting these elements [x] into the set [X], will give rise to a very
unbalanced tree. *)
Fixpoint coPpick_raw (t : coPset_raw) : option positive :=
match t with
| coPLeaf true | coPNode true _ _ => Some 1
......@@ -269,89 +279,109 @@ Proof.
Qed.
(** * Conversion to psets *)
Fixpoint coPset_to_Pset_raw (t : coPset_raw) : Pmap_raw () :=
Fixpoint coPset_to_Pset_raw (t : coPset_raw) : Pmap () :=
match t with
| coPLeaf _ => PLeaf
| coPNode false l r => PNode' None (coPset_to_Pset_raw l) (coPset_to_Pset_raw r)
| coPNode true l r => PNode (Some ()) (coPset_to_Pset_raw l) (coPset_to_Pset_raw r)
| coPLeaf _ => PEmpty
| coPNode false l r => pmap.PNode (coPset_to_Pset_raw l) None (coPset_to_Pset_raw r)
| coPNode true l r => pmap.PNode (coPset_to_Pset_raw l) (Some ()) (coPset_to_Pset_raw r)
end.
Lemma coPset_to_Pset_wf t : coPset_wf t Pmap_wf (coPset_to_Pset_raw t).
Proof. induction t as [|[]]; simpl; eauto using PNode_wf. Qed.
Definition coPset_to_Pset (X : coPset) : Pset :=
let (t,Ht) := X in Mapset (PMap (coPset_to_Pset_raw t) (coPset_to_Pset_wf _ Ht)).
let (t,Ht) := X in Mapset (coPset_to_Pset_raw t).
Lemma elem_of_coPset_to_Pset X i : set_finite X i coPset_to_Pset X i X.
Proof.
rewrite coPset_finite_spec; destruct X as [t Ht].
change (coPset_finite t coPset_to_Pset_raw t !! i = Some () e_of i t).
clear Ht; revert i; induction t as [[]|[] l IHl r IHr]; intros [i|i|];
simpl; rewrite ?andb_True, ?PNode_lookup; naive_solver.
simpl; rewrite ?andb_True, ?pmap.Pmap_lookup_PNode; naive_solver.
Qed.
(** * Conversion from psets *)
Fixpoint Pset_to_coPset_raw (t : Pmap_raw ()) : coPset_raw :=
match t with
| PLeaf => coPLeaf false
| PNode None l r => coPNode false (Pset_to_coPset_raw l) (Pset_to_coPset_raw r)
| PNode (Some _) l r => coPNode true (Pset_to_coPset_raw l) (Pset_to_coPset_raw r)
end.
Lemma Pset_to_coPset_wf t : Pmap_wf t coPset_wf (Pset_to_coPset_raw t).
Definition Pset_to_coPset_raw_aux (go : Pmap_ne () coPset_raw)
(mt : Pmap ()) : coPset_raw :=
match mt with PNodes t => go t | PEmpty => coPLeaf false end.
Fixpoint Pset_ne_to_coPset_raw (t : Pmap_ne ()) : coPset_raw :=
pmap.Pmap_ne_case t $ λ ml mx mr,
coPNode match mx with Some _ => true | None => false end
(Pset_to_coPset_raw_aux Pset_ne_to_coPset_raw ml)
(Pset_to_coPset_raw_aux Pset_ne_to_coPset_raw mr).
Definition Pset_to_coPset_raw : Pmap () coPset_raw :=
Pset_to_coPset_raw_aux Pset_ne_to_coPset_raw.
Lemma Pset_to_coPset_raw_PNode ml mx mr :
pmap.PNode_valid ml mx mr
Pset_to_coPset_raw (pmap.PNode ml mx mr) =
coPNode match mx with Some _ => true | None => false end
(Pset_to_coPset_raw ml) (Pset_to_coPset_raw mr).
Proof. by destruct ml, mx, mr. Qed.
Lemma Pset_to_coPset_raw_wf t : coPset_wf (Pset_to_coPset_raw t).
Proof.
induction t as [|[] l IHl r IHr]; simpl; rewrite ?andb_True; auto.
- intros [??]; destruct l as [|[]], r as [|[]]; simpl in *; auto.
- destruct l as [|[]], r as [|[]]; simpl in *; rewrite ?andb_true_r;
rewrite ?andb_True; rewrite ?andb_True in IHl, IHr; intuition.
induction t as [|ml mx mr] using pmap.Pmap_ind; [done|].
rewrite Pset_to_coPset_raw_PNode by done.
apply coPNode_wf; [done|done|..];
destruct mx; destruct ml using pmap.Pmap_ind; destruct mr using pmap.Pmap_ind;
rewrite ?Pset_to_coPset_raw_PNode by done; naive_solver.
Qed.
Lemma elem_of_Pset_to_coPset_raw i t : e_of i (Pset_to_coPset_raw t) t !! i = Some ().
Proof. by revert i; induction t as [|[[]|]]; intros []; simpl; auto; split. Qed.
Proof.
revert i. induction t as [|ml mx mr] using pmap.Pmap_ind; [done|].
intros []; rewrite Pset_to_coPset_raw_PNode,
pmap.Pmap_lookup_PNode by done; destruct mx as [[]|]; naive_solver.
Qed.
Lemma Pset_to_coPset_raw_finite t : coPset_finite (Pset_to_coPset_raw t).
Proof. induction t as [|[[]|]]; simpl; rewrite ?andb_True; auto. Qed.
Proof.
induction t as [|ml mx mr] using pmap.Pmap_ind; [done|].
rewrite Pset_to_coPset_raw_PNode by done. destruct mx; naive_solver.
Qed.
Definition Pset_to_coPset (X : Pset) : coPset :=
let 'Mapset (PMap t Ht) := X in Pset_to_coPset_raw t Pset_to_coPset_wf _ Ht.
let 'Mapset t := X in Pset_to_coPset_raw t Pset_to_coPset_raw_wf _.
Lemma elem_of_Pset_to_coPset X i : i Pset_to_coPset X i X.
Proof. destruct X as [[t ?]]; apply elem_of_Pset_to_coPset_raw. Qed.
Proof. destruct X; apply elem_of_Pset_to_coPset_raw. Qed.
Lemma Pset_to_coPset_finite X : set_finite (Pset_to_coPset X).
Proof.
apply coPset_finite_spec; destruct X as [[t ?]]; apply Pset_to_coPset_raw_finite.
Qed.
Proof. apply coPset_finite_spec; destruct X; apply Pset_to_coPset_raw_finite. Qed.
(** * Conversion to and from gsets of positives *)
Lemma coPset_to_gset_wf (m : Pmap ()) : gmap_wf positive m.
Proof. unfold gmap_wf. by rewrite bool_decide_spec. Qed.
Definition coPset_to_gset (X : coPset) : gset positive :=
let 'Mapset m := coPset_to_Pset X in
Mapset (GMap m (coPset_to_gset_wf m)).
Mapset (pmap_to_gmap m).
Definition gset_to_coPset (X : gset positive) : coPset :=
let 'Mapset (GMap (PMap t Ht) _) := X in Pset_to_coPset_raw t Pset_to_coPset_wf _ Ht.
let 'Mapset m := X in
Pset_to_coPset_raw (gmap_to_pmap m) Pset_to_coPset_raw_wf _.
Lemma elem_of_coPset_to_gset X i : set_finite X i coPset_to_gset X i X.
Proof.
intros ?. rewrite <-elem_of_coPset_to_Pset by done.
unfold coPset_to_gset. by destruct (coPset_to_Pset X).
intros ?. rewrite <-elem_of_coPset_to_Pset by done. destruct X as [X ?].
unfold elem_of, gset_elem_of, mapset_elem_of, coPset_to_gset; simpl.
by rewrite lookup_pmap_to_gmap.
Qed.
Lemma elem_of_gset_to_coPset X i : i gset_to_coPset X i X.
Proof. destruct X as [[[t ?]]]; apply elem_of_Pset_to_coPset_raw. Qed.
Proof.
destruct X as [m]. unfold elem_of, coPset_elem_of; simpl.
by rewrite elem_of_Pset_to_coPset_raw, lookup_gmap_to_pmap.
Qed.
Lemma gset_to_coPset_finite X : set_finite (gset_to_coPset X).
Proof.
apply coPset_finite_spec; destruct X as [[[t ?]]]; apply Pset_to_coPset_raw_finite.
apply coPset_finite_spec; destruct X as [[?]]; apply Pset_to_coPset_raw_finite.
Qed.
(** * Domain of finite maps *)
Instance Pmap_dom_coPset {A} : Dom (Pmap A) coPset := λ m, Pset_to_coPset (dom _ m).
Instance Pmap_dom_coPset_spec: FinMapDom positive Pmap coPset.
(** * Infinite sets *)
Lemma coPset_infinite_finite (X : coPset) : set_infinite X ¬set_finite X.
Proof.
split; try apply _; intros A m i; unfold dom, Pmap_dom_coPset.
by rewrite elem_of_Pset_to_coPset, elem_of_dom.
split; [intros ??; by apply (set_not_infinite_finite X)|].
intros Hfin xs. exists (coPpick (X list_to_set xs)).
cut (coPpick (X list_to_set xs) X list_to_set xs); [set_solver|].
apply coPpick_elem_of; intros Hfin'.
apply Hfin, (difference_finite_inv _ (list_to_set xs)), Hfin'.
apply list_to_set_finite.
Qed.
Instance gmap_dom_coPset {A} : Dom (gmap positive A) coPset := λ m,
gset_to_coPset (dom _ m).
Instance gmap_dom_coPset_spec: FinMapDom positive (gmap positive) coPset.
Lemma coPset_finite_infinite (X : coPset) : set_finite X ¬set_infinite X.
Proof. rewrite coPset_infinite_finite. split; [tauto|apply dec_stable]. Qed.
Global Instance coPset_infinite_dec (X : coPset) : Decision (set_infinite X).
Proof.
split; try apply _; intros A m i; unfold dom, gmap_dom_coPset.
by rewrite elem_of_gset_to_coPset, elem_of_dom.
Qed.
refine (cast_if (decide (¬set_finite X))); by rewrite coPset_infinite_finite.
Defined.
(** * Suffix sets *)
Fixpoint coPset_suffixes_raw (p : positive) : coPset_raw :=
......@@ -410,7 +440,7 @@ Proof.
Qed.
Lemma coPset_lr_union X : coPset_l X coPset_r X = X.
Proof.
apply elem_of_equiv_L; intros p; apply eq_bool_prop_elim.
apply set_eq; intros p; apply eq_bool_prop_elim.
destruct X as [t Ht]; simpl; clear Ht; rewrite coPset_elem_of_union.
revert p; induction t as [[]|[]]; intros [?|?|]; simpl;
rewrite ?coPset_elem_of_node; simpl;
......@@ -433,3 +463,10 @@ Proof.
exists (coPset_l X), (coPset_r X); eauto 10 using coPset_lr_union,
coPset_lr_disjoint, coPset_l_finite, coPset_r_finite.
Qed.
Lemma coPset_split_infinite (X : coPset) :
set_infinite X
X1 X2, X = X1 X2 X1 X2 = set_infinite X1 set_infinite X2.
Proof.
setoid_rewrite coPset_infinite_finite.
eapply coPset_split.
Qed.
From Coq.QArith Require Import QArith_base Qcanon.
From stdpp Require Export list numbers list_numbers fin.
From stdpp Require Import well_founded.
From stdpp Require Import options.
Local Open Scope positive.
(** Note that [Countable A] gives rise to [EqDecision A] by checking equality of
the results of [encode]. This instance of [EqDecision A] is very inefficient, so
the native decider is typically preferred for actual computation. To avoid
overlapping instances, we include [EqDecision A] explicitly as a parameter of
[Countable A]. *)
Class Countable A `{EqDecision A} := {
encode : A positive;
decode : positive option A;
decode_encode x : decode (encode x) = Some x
}.
Hint Mode Countable ! - : typeclass_instances.
Arguments encode : simpl never.
Arguments decode : simpl never.
Global Hint Mode Countable ! - : typeclass_instances.
Global Arguments encode : simpl never.
Global Arguments decode : simpl never.
Instance encode_inj `{Countable A} : Inj (=) (=) (encode (A:=A)).
Global Instance encode_inj `{Countable A} : Inj (=) (=) (encode (A:=A)).
Proof.
intros x y Hxy; apply (inj Some).
by rewrite <-(decode_encode x), Hxy, decode_encode.
......@@ -22,7 +28,7 @@ Definition encode_nat `{Countable A} (x : A) : nat :=
pred (Pos.to_nat (encode x)).
Definition decode_nat `{Countable A} (i : nat) : option A :=
decode (Pos.of_nat (S i)).
Instance encode_nat_inj `{Countable A} : Inj (=) (=) (encode_nat (A:=A)).
Global Instance encode_nat_inj `{Countable A} : Inj (=) (=) (encode_nat (A:=A)).
Proof. unfold encode_nat; intros x y Hxy; apply (inj encode); lia. Qed.
Lemma decode_encode_nat `{Countable A} (x : A) : decode_nat (encode_nat x) = Some x.
Proof.
......@@ -35,7 +41,7 @@ Definition encode_Z `{Countable A} (x : A) : Z :=
Zpos (encode x).
Definition decode_Z `{Countable A} (i : Z) : option A :=
match i with Zpos i => decode i | _ => None end.
Instance encode_Z_inj `{Countable A} : Inj (=) (=) (encode_Z (A:=A)).
Global Instance encode_Z_inj `{Countable A} : Inj (=) (=) (encode_Z (A:=A)).
Proof. unfold encode_Z; intros x y Hxy; apply (inj encode); lia. Qed.
Lemma decode_encode_Z `{Countable A} (x : A) : decode_Z (encode_Z x) = Some x.
Proof. apply decode_encode. Qed.
......@@ -55,10 +61,9 @@ Section choice.
{ intros help. by apply (help (encode x)). }
intros i. induction i as [|i IH] using Pos.peano_ind; intros p ??.
{ constructor. intros j. assert (p = encode x) by lia; subst.
inversion 1 as [? Hd|?? Hd]; subst;
rewrite decode_encode in Hd; congruence. }
inv 1 as [? Hd|?? Hd]; rewrite decode_encode in Hd; congruence. }
constructor. intros j.
inversion 1 as [? Hd|? y Hd]; subst; auto with lia.
inv 1 as [? Hd|? y Hd]; auto with lia.
Qed.
Context `{ x, Decision (P x)}.
......@@ -84,6 +89,25 @@ Section choice.
Definition choice (HA : x, P x) : { x | P x } := _choose_correct HA.
End choice.
Section choice_proper.
Context `{Countable A}.
Context (P1 P2 : A Prop) `{ x, Decision (P1 x)} `{ x, Decision (P2 x)}.
Context (Heq : x, P1 x P2 x).
Lemma choose_go_proper {i} (acc1 acc2 : Acc (choose_step _) i) :
choose_go P1 acc1 = choose_go P2 acc2.
Proof using Heq.
induction acc1 as [i a1 IH] using Acc_dep_ind;
destruct acc2 as [acc2]; simpl.
destruct (Some_dec _) as [[x Hx]|]; [|done].
do 2 case_decide; done || exfalso; naive_solver.
Qed.
Lemma choose_proper p1 p2 :
choose P1 p1 = choose P2 p2.
Proof using Heq. apply choose_go_proper. Qed.
End choice_proper.
Lemma surj_cancel `{Countable A} `{EqDecision B}
(f : A B) `{!Surj (=) f} : { g : B A & Cancel (=) f g }.
Proof.
......@@ -97,7 +121,7 @@ Section inj_countable.
Context `{Countable A, EqDecision B}.
Context (f : B A) (g : A option B) (fg : x, g (f x) = Some x).
Program Instance inj_countable : Countable B :=
Program Definition inj_countable : Countable B :=
{| encode y := encode (f y); decode p := x decode p; g x |}.
Next Obligation. intros y; simpl; rewrite decode_encode; eauto. Qed.
End inj_countable.
......@@ -106,29 +130,29 @@ Section inj_countable'.
Context `{Countable A, EqDecision B}.
Context (f : B A) (g : A B) (fg : x, g (f x) = x).
Program Instance inj_countable' : Countable B := inj_countable f (Some g) _.
Program Definition inj_countable' : Countable B := inj_countable f (Some g) _.
Next Obligation. intros x. by f_equal/=. Qed.
End inj_countable'.
(** ** Empty *)
Program Instance Empty_set_countable : Countable Empty_set :=
Global Program Instance Empty_set_countable : Countable Empty_set :=
{| encode u := 1; decode p := None |}.
Next Obligation. by intros []. Qed.
(** ** Unit *)
Program Instance unit_countable : Countable unit :=
Global Program Instance unit_countable : Countable unit :=
{| encode u := 1; decode p := Some () |}.
Next Obligation. by intros []. Qed.
(** ** Bool *)
Program Instance bool_countable : Countable bool := {|
Global Program Instance bool_countable : Countable bool := {|
encode b := if b then 1 else 2;
decode p := Some match p return bool with 1 => true | _ => false end
|}.
Next Obligation. by intros []. Qed.
(** ** Option *)
Program Instance option_countable `{Countable A} : Countable (option A) := {|
Global Program Instance option_countable `{Countable A} : Countable (option A) := {|
encode o := match o with None => 1 | Some x => Pos.succ (encode x) end;
decode p := if decide (p = 1) then Some None else Some <$> decode (Pos.pred p)
|}.
......@@ -138,7 +162,7 @@ Next Obligation.
Qed.
(** ** Sums *)
Program Instance sum_countable `{Countable A} `{Countable B} :
Global Program Instance sum_countable `{Countable A} `{Countable B} :
Countable (A + B)%type := {|
encode xy :=
match xy with inl x => (encode x)~0 | inr y => (encode y)~1 end;
......@@ -211,7 +235,7 @@ Proof.
{ intros p'. by induction p'; simplify_option_eq. }
revert q. by induction p; intros [?|?|]; simplify_option_eq.
Qed.
Program Instance prod_countable `{Countable A} `{Countable B} :
Global Program Instance prod_countable `{Countable A} `{Countable B} :
Countable (A * B)%type := {|
encode xy := prod_encode (encode (xy.1)) (encode (xy.2));
decode p :=
......@@ -238,9 +262,9 @@ Next Obligation.
Qed.
(** ** Numbers *)
Instance pos_countable : Countable positive :=
Global Instance pos_countable : Countable positive :=
{| encode := id; decode := Some; decode_encode x := eq_refl |}.
Program Instance N_countable : Countable N := {|
Global Program Instance N_countable : Countable N := {|
encode x := match x with N0 => 1 | Npos p => Pos.succ p end;
decode p := if decide (p = 1) then Some 0%N else Some (Npos (Pos.pred p))
|}.
......@@ -248,18 +272,18 @@ Next Obligation.
intros [|p]; simpl; [done|].
by rewrite decide_False, Pos.pred_succ by (by destruct p).
Qed.
Program Instance Z_countable : Countable Z := {|
Global Program Instance Z_countable : Countable Z := {|
encode x := match x with Z0 => 1 | Zpos p => p~0 | Zneg p => p~1 end;
decode p := Some match p with 1 => Z0 | p~0 => Zpos p | p~1 => Zneg p end
|}.
Next Obligation. by intros [|p|p]. Qed.
Program Instance nat_countable : Countable nat :=
Global Program Instance nat_countable : Countable nat :=
{| encode x := encode (N.of_nat x); decode p := N.to_nat <$> decode p |}.
Next Obligation.
by intros x; lazy beta; rewrite decode_encode; csimpl; rewrite Nat2N.id.
Qed.
Program Instance Qc_countable : Countable Qc :=
Global Program Instance Qc_countable : Countable Qc :=
inj_countable
(λ p : Qc, let 'Qcmake (x # y) _ := p return _ in (x,y))
(λ q : Z * positive, let '(x,y) := q return _ in Some (Q2Qc (x # y))) _.
......@@ -267,19 +291,19 @@ Next Obligation.
intros [[x y] Hcan]. f_equal. apply Qc_is_canon. simpl. by rewrite Hcan.
Qed.
Program Instance Qp_countable : Countable Qp :=
Global Program Instance Qp_countable : Countable Qp :=
inj_countable
Qp_car
(λ p : Qc, guard (0 < p)%Qc as Hp; Some (mk_Qp p Hp)) _.
Qp_to_Qc
(λ p : Qc, Hp guard (0 < p)%Qc; Some (mk_Qp p Hp)) _.
Next Obligation.
intros [p Hp]. unfold mguard, option_guard; simpl.
case_match; [|done]. f_equal. by apply Qp_eq.
intros [p Hp]. case_guard; simplify_eq/=; [|done].
f_equal. by apply Qp.to_Qc_inj_iff.
Qed.
Program Instance fin_countable n : Countable (fin n) :=
Global Program Instance fin_countable n : Countable (fin n) :=
inj_countable
fin_to_nat
(λ m : nat, guard (m < n)%nat as Hm; Some (nat_to_fin Hm)) _.
(λ m : nat, Hm guard (m < n)%nat; Some (nat_to_fin Hm)) _.
Next Obligation.
intros n i; simplify_option_eq.
- by rewrite nat_to_fin_to_nat.
......@@ -289,13 +313,25 @@ Qed.
(** ** Generic trees *)
Local Close Scope positive.
(** This type can help you construct a [Countable] instance for an arbitrary
(even recursive) inductive datatype. The idea is tht you make [T] something like
[T1 + T2 + ...], covering all the data types that can be contained inside your
type.
- Each non-recursive constructor to a [GenLeaf]. Different constructors must use
different variants of [T] to ensure they remain distinguishable!
- Each recursive constructor to a [GenNode] where the [nat] is a (typically
small) constant representing the constructor itself, and then all the data in
the constructor (recursive or otherwise) is put into child nodes.
This data type is the same as [GenTree.tree] in mathcomp, see
https://github.com/math-comp/math-comp/blob/master/ssreflect/choice.v *)
Inductive gen_tree (T : Type) : Type :=
| GenLeaf : T gen_tree T
| GenNode : nat list (gen_tree T) gen_tree T.
Arguments GenLeaf {_} _ : assert.
Arguments GenNode {_} _ _ : assert.
Global Arguments GenLeaf {_} _ : assert.
Global Arguments GenNode {_} _ _ : assert.
Instance gen_tree_dec `{EqDecision T} : EqDecision (gen_tree T).
Global Instance gen_tree_dec `{EqDecision T} : EqDecision (gen_tree T).
Proof.
refine (
fix go t1 t2 := let _ : EqDecision _ := @go in
......@@ -330,13 +366,22 @@ Proof.
- rewrite <-(assoc_L _). revert k. generalize ([inl (length ts, n)] ++ l).
induction ts as [|t ts'' IH]; intros k ts'''; csimpl; auto.
rewrite reverse_cons, <-!(assoc_L _), FIX; simpl; auto.
- simpl. by rewrite take_app_alt, drop_app_alt, reverse_involutive
by (by rewrite reverse_length).
- simpl. by rewrite take_app_length', drop_app_length', reverse_involutive
by (by rewrite length_reverse).
Qed.
Program Instance gen_tree_countable `{Countable T} : Countable (gen_tree T) :=
Global Program Instance gen_tree_countable `{Countable T} : Countable (gen_tree T) :=
inj_countable gen_tree_to_list (gen_tree_of_list []) _.
Next Obligation.
intros T ?? t.
by rewrite <-(right_id_L [] _ (gen_tree_to_list _)), gen_tree_of_to_list.
Qed.
(** ** Sigma *)
Global Program Instance countable_sig `{Countable A} (P : A Prop)
`{!∀ x, Decision (P x), !∀ x, ProofIrrel (P x)} :
Countable { x : A | P x } :=
inj_countable proj1_sig (λ x, Hx guard (P x); Some (x Hx)) _.
Next Obligation.
intros A ?? P ?? [x Hx]. by erewrite (option_guard_True_pi (P x)).
Qed.
......@@ -12,7 +12,7 @@ Proof. firstorder. Qed.
Lemma Is_true_reflect (b : bool) : reflect b b.
Proof. destruct b; [left; constructor | right; intros []]. Qed.
Instance: Inj (=) () Is_true.
Global Instance: Inj (=) () Is_true.
Proof. intros [] []; simpl; intuition. Qed.
Lemma decide_True {A} `{Decision P} (x y : A) :
......@@ -21,13 +21,13 @@ Proof. destruct (decide P); tauto. Qed.
Lemma decide_False {A} `{Decision P} (x y : A) :
¬P (if decide P then x else y) = y.
Proof. destruct (decide P); tauto. Qed.
Lemma decide_iff {A} P Q `{Decision P, Decision Q} (x y : A) :
Lemma decide_ext {A} P Q `{Decision P, Decision Q} (x y : A) :
(P Q) (if decide P then x else y) = (if decide Q then x else y).
Proof. intros [??]. destruct (decide P), (decide Q); tauto. Qed.
Lemma decide_left `{Decision P, !ProofIrrel P} (HP : P) : decide P = left HP.
Lemma decide_True_pi `{Decision P, !ProofIrrel P} (HP : P) : decide P = left HP.
Proof. destruct (decide P); [|contradiction]. f_equal. apply proof_irrel. Qed.
Lemma decide_right `{Decision P, !ProofIrrel (¬ P)} (HP : ¬ P) : decide P = right HP.
Lemma decide_False_pi `{Decision P, !ProofIrrel (¬ P)} (HP : ¬ P) : decide P = right HP.
Proof. destruct (decide P); [contradiction|]. f_equal. apply proof_irrel. Qed.
(** The tactic [destruct_decide] destructs a sumbool [dec]. If one of the
......@@ -85,6 +85,79 @@ Notation cast_if_or3 S1 S2 S3 := (if S1 then left _ else cast_if_or S2 S3).
Notation cast_if_not_or S1 S2 := (if S1 then cast_if S2 else left _).
Notation cast_if_not S := (if S then right _ else left _).
(** * Instances of [Decision] *)
(** Instances of [Decision] for operators of propositional logic. *)
(** The instances for [True] and [False] have a very high cost. If they are
applied too eagerly, HO-unification could wrongfully instantiate TC instances
with [λ .., True] or [λ .., False].
See https://gitlab.mpi-sws.org/iris/stdpp/-/issues/165 *)
Global Instance True_dec: Decision True | 1000 := left I.
Global Instance False_dec: Decision False | 1000 := right (False_rect False).
Global Instance Is_true_dec b : Decision (Is_true b).
Proof. destruct b; simpl; apply _. Defined.
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.
Global Instance iff_dec `(P_dec : Decision P) `(Q_dec : Decision Q) :
Decision (P Q) := and_dec _ _.
(** Instances of [Decision] for common data types. *)
Global Instance bool_eq_dec : EqDecision bool.
Proof. solve_decision. Defined.
Global Instance unit_eq_dec : EqDecision unit.
Proof. solve_decision. Defined.
Global Instance Empty_set_eq_dec : EqDecision Empty_set.
Proof. solve_decision. Defined.
Global Instance prod_eq_dec `{EqDecision A, EqDecision B} : EqDecision (A * B).
Proof. solve_decision. Defined.
Global Instance sum_eq_dec `{EqDecision A, EqDecision B} : EqDecision (A + B).
Proof. solve_decision. Defined.
Global Instance uncurry_dec `(P_dec : (x : A) (y : B), Decision (P x y)) p :
Decision (uncurry P p) :=
match p as p return Decision (uncurry P p) with
| (x,y) => P_dec x y
end.
Global Instance sig_eq_dec `(P : A Prop) `{ x, ProofIrrel (P x), EqDecision A} :
EqDecision (sig P).
Proof.
refine (λ x y, cast_if (decide (`x = `y))); rewrite sig_eq_pi; trivial.
Defined.
(** Some laws for decidable propositions *)
Lemma not_and_l {P Q : Prop} `{Decision P} : ¬(P Q) ¬P ¬Q.
Proof. destruct (decide P); tauto. Qed.
Lemma not_and_r {P Q : Prop} `{Decision Q} : ¬(P Q) ¬P ¬Q.
Proof. destruct (decide Q); tauto. Qed.
Lemma not_and_l_alt {P Q : Prop} `{Decision P} : ¬(P Q) ¬P (¬Q P).
Proof. destruct (decide P); tauto. Qed.
Lemma not_and_r_alt {P Q : Prop} `{Decision Q} : ¬(P Q) (¬P Q) ¬Q.
Proof. destruct (decide Q); tauto. Qed.
Program Definition inj_eq_dec `{EqDecision A} {B} (f : B A)
`{!Inj (=) (=) f} : EqDecision B := λ x y, cast_if (decide (f x = f y)).
Solve Obligations with firstorder congruence.
(** * Instances of [RelDecision] *)
Definition flip_dec {A} (R : relation A) `{!RelDecision R} :
RelDecision (flip R) := λ x y, decide_rel R y x.
(** We do not declare this as an actual instance since Coq can unify [flip ?R]
with any relation. Coq's standard library is carrying out the same approach for
the [Reflexive], [Transitive], etc, instance of [flip]. *)
Global Hint Extern 3 (RelDecision (flip _)) => apply flip_dec : typeclass_instances.
(** We can convert decidable propositions to booleans. *)
Definition bool_decide (P : Prop) {dec : Decision P} : bool :=
if dec then true else false.
......@@ -99,7 +172,7 @@ Lemma decide_bool_decide P {Hdec: Decision P} {X : Type} (x1 x2 : X):
(if decide P then x1 else x2) = (if bool_decide P then x1 else x2).
Proof. unfold bool_decide, decide. destruct Hdec; reflexivity. Qed.
Tactic Notation "case_bool_decide" "as" ident (Hd) :=
Tactic Notation "case_bool_decide" "as" ident(Hd) :=
match goal with
| H : context [@bool_decide ?P ?dec] |- _ =>
destruct_decide (@bool_decide_reflect P dec) as Hd
......@@ -115,15 +188,15 @@ Lemma bool_decide_unpack (P : Prop) {dec : Decision P} : bool_decide P → P.
Proof. rewrite bool_decide_spec; trivial. Qed.
Lemma bool_decide_pack (P : Prop) {dec : Decision P} : P bool_decide P.
Proof. rewrite bool_decide_spec; trivial. Qed.
Hint Resolve bool_decide_pack : core.
Global Hint Resolve bool_decide_pack : core.
Lemma bool_decide_eq_true (P : Prop) `{Decision P} : bool_decide P = true P.
Proof. case_bool_decide; intuition discriminate. Qed.
Lemma bool_decide_eq_false (P : Prop) `{Decision P} : bool_decide P = false ¬P.
Proof. case_bool_decide; intuition discriminate. Qed.
Lemma bool_decide_iff (P Q : Prop) `{Decision P, Decision Q} :
Lemma bool_decide_ext (P Q : Prop) `{Decision P, Decision Q} :
(P Q) bool_decide P = bool_decide Q.
Proof. repeat case_bool_decide; tauto. Qed.
Proof. apply decide_ext. Qed.
Lemma bool_decide_eq_true_1 P `{!Decision P}: bool_decide P = true P.
Proof. apply bool_decide_eq_true. Qed.
......@@ -135,6 +208,40 @@ Proof. apply bool_decide_eq_false. Qed.
Lemma bool_decide_eq_false_2 P `{!Decision P}: ¬P bool_decide P = false.
Proof. apply bool_decide_eq_false. Qed.
Lemma bool_decide_True : bool_decide True = true.
Proof. reflexivity. Qed.
Lemma bool_decide_False : bool_decide False = false.
Proof. reflexivity. Qed.
Lemma bool_decide_not P `{Decision P} :
bool_decide (¬ P) = negb (bool_decide P).
Proof. repeat case_bool_decide; intuition. Qed.
Lemma bool_decide_or P Q `{Decision P, Decision Q} :
bool_decide (P Q) = bool_decide P || bool_decide Q.
Proof. repeat case_bool_decide; intuition. Qed.
Lemma bool_decide_and P Q `{Decision P, Decision Q} :
bool_decide (P Q) = bool_decide P && bool_decide Q.
Proof. repeat case_bool_decide; intuition. Qed.
Lemma bool_decide_impl P Q `{Decision P, Decision Q} :
bool_decide (P Q) = implb (bool_decide P) (bool_decide Q).
Proof. repeat case_bool_decide; intuition. Qed.
Lemma bool_decide_iff P Q `{Decision P, Decision Q} :
bool_decide (P Q) = eqb (bool_decide P) (bool_decide Q).
Proof. repeat case_bool_decide; intuition. Qed.
(** The tactic [compute_done] solves the following kinds of goals:
- Goals [P] where [Decidable P] can be derived.
- Goals that compute to [True] or [x = x].
The goal must be a ground term for this, i.e., not contain variables (that do
not compute away). The goal is solved by using [vm_compute] and then using a
trivial proof term ([I]/[eq_refl]). *)
Tactic Notation "compute_done" :=
try apply (bool_decide_unpack _);
vm_compute;
first [ exact I | exact eq_refl ].
Tactic Notation "compute_by" tactic(tac) :=
tac; compute_done.
(** Backwards compatibility notations. *)
Notation bool_decide_true := bool_decide_eq_true_2.
Notation bool_decide_false := bool_decide_eq_false_2.
......@@ -157,71 +264,3 @@ Proof. apply (sig_eq_pi _). Qed.
Lemma dexists_proj1 `(P : A Prop) `{ x, Decision (P x)} (x : dsig P) p :
dexist (`x) p = x.
Proof. apply dsig_eq; reflexivity. Qed.
(** * Instances of [Decision] *)
(** Instances of [Decision] for operators of propositional logic. *)
Instance True_dec: Decision True := left I.
Instance False_dec: Decision False := right (False_rect False).
Instance Is_true_dec b : Decision (Is_true b).
Proof. destruct b; simpl; apply _. Defined.
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.
Instance iff_dec `(P_dec : Decision P) `(Q_dec : Decision Q) :
Decision (P Q) := and_dec _ _.
(** Instances of [Decision] for common data types. *)
Instance bool_eq_dec : EqDecision bool.
Proof. solve_decision. Defined.
Instance unit_eq_dec : EqDecision unit.
Proof. solve_decision. Defined.
Instance Empty_set_eq_dec : EqDecision Empty_set.
Proof. solve_decision. Defined.
Instance prod_eq_dec `{EqDecision A, EqDecision B} : EqDecision (A * B).
Proof. solve_decision. Defined.
Instance sum_eq_dec `{EqDecision A, EqDecision B} : EqDecision (A + B).
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.
Instance sig_eq_dec `(P : A Prop) `{ x, ProofIrrel (P x), EqDecision A} :
EqDecision (sig P).
Proof.
refine (λ x y, cast_if (decide (`x = `y))); rewrite sig_eq_pi; trivial.
Defined.
(** Some laws for decidable propositions *)
Lemma not_and_l {P Q : Prop} `{Decision P} : ¬(P Q) ¬P ¬Q.
Proof. destruct (decide P); tauto. Qed.
Lemma not_and_r {P Q : Prop} `{Decision Q} : ¬(P Q) ¬P ¬Q.
Proof. destruct (decide Q); tauto. Qed.
Lemma not_and_l_alt {P Q : Prop} `{Decision P} : ¬(P Q) ¬P (¬Q P).
Proof. destruct (decide P); tauto. Qed.
Lemma not_and_r_alt {P Q : Prop} `{Decision Q} : ¬(P Q) (¬P Q) ¬Q.
Proof. destruct (decide Q); tauto. Qed.
Program Definition inj_eq_dec `{EqDecision A} {B} (f : B A)
`{!Inj (=) (=) f} : EqDecision B := λ x y, cast_if (decide (f x = f y)).
Solve Obligations with firstorder congruence.
(** * Instances of [RelDecision] *)
Definition flip_dec {A} (R : relation A) `{!RelDecision R} :
RelDecision (flip R) := λ x y, decide_rel R y x.
(** We do not declare this as an actual instance since Coq can unify [flip ?R]
with any relation. Coq's standard library is carrying out the same approach for
the [Reflexive], [Transitive], etc, instance of [flip]. *)
Hint Extern 3 (RelDecision (flip _)) => apply flip_dec : typeclass_instances.
(include_subdirs qualified)
(coq.theory
(name stdpp)
(package coq-stdpp))
......@@ -15,15 +15,15 @@ ambiguity. *)
Notation fin := Fin.t.
Notation FS := Fin.FS.
Declare Scope fin_scope.
Delimit Scope fin_scope with fin.
Arguments Fin.FS _ _%fin : assert.
Bind Scope fin_scope with fin.
Global Arguments Fin.FS _ _%fin : assert.
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.
(** Allow any non-negative number literal to be parsed as a [fin]. For example
[42%fin : fin 64], or [42%fin : fin _], or [42%fin : fin (43 + _)]. *)
Number Notation fin Nat.of_num_uint Nat.to_num_uint (via nat
mapping [[Fin.F1] => O, [Fin.FS] => S]) : 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.
......@@ -32,7 +32,7 @@ Coercion fin_to_nat : fin >-> nat.
Notation nat_to_fin := Fin.of_nat_lt.
Notation fin_rect2 := Fin.rect2.
Instance fin_dec {n} : EqDecision (fin n).
Global Instance fin_dec {n} : EqDecision (fin n).
Proof.
refine (fin_rect2
(λ n (i j : fin n), { i = j } + { i j })
......@@ -64,17 +64,19 @@ Ltac inv_fin i :=
| fin ?n =>
match eval hnf in n with
| 0 =>
revert dependent i; match goal with |- i, @?P i => apply (fin_0_inv P) end
generalize dependent i;
match goal with |- i, @?P i => apply (fin_0_inv P) end
| S ?n =>
revert dependent i; match goal with |- i, @?P i => apply (fin_S_inv P) end
generalize dependent i;
match goal with |- i, @?P i => apply (fin_S_inv P) end
end
end.
Instance FS_inj: Inj (=) (=) (@FS n).
Proof. intros n i j. apply Fin.FS_inj. Qed.
Instance fin_to_nat_inj : Inj (=) (=) (@fin_to_nat n).
Global Instance FS_inj {n} : Inj (=) (=) (@FS n).
Proof. intros i j. apply Fin.FS_inj. Qed.
Global Instance fin_to_nat_inj {n} : Inj (=) (=) (@fin_to_nat n).
Proof.
intros n i. induction i; intros j; inv_fin j; intros; f_equal/=; auto with lia.
intros i. induction i; intros j; inv_fin j; intros; f_equal/=; auto with lia.
Qed.
Lemma fin_to_nat_lt {n} (i : fin n) : fin_to_nat i < n.
......@@ -86,26 +88,26 @@ Qed.
Lemma nat_to_fin_to_nat {n} (i : fin n) H : @nat_to_fin (fin_to_nat i) n H = i.
Proof. apply (inj fin_to_nat), fin_to_nat_to_fin. Qed.
Fixpoint fin_plus_inv {n1 n2} : (P : fin (n1 + n2) Type)
Fixpoint fin_add_inv {n1 n2} : (P : fin (n1 + n2) Type)
(H1 : i1 : fin n1, P (Fin.L n2 i1))
(H2 : i2, P (Fin.R n1 i2)) (i : fin (n1 + n2)), P i :=
match n1 with
| 0 => λ P H1 H2 i, H2 i
| S n => λ P H1 H2, fin_S_inv P (H1 0%fin) (fin_plus_inv _ (λ i, H1 (FS i)) H2)
| S n => λ P H1 H2, fin_S_inv P (H1 0%fin) (fin_add_inv _ (λ i, H1 (FS i)) H2)
end.
Lemma fin_plus_inv_L {n1 n2} (P : fin (n1 + n2) Type)
Lemma fin_add_inv_l {n1 n2} (P : fin (n1 + n2) Type)
(H1: i1 : fin n1, P (Fin.L _ i1)) (H2: i2, P (Fin.R _ i2)) (i: fin n1) :
fin_plus_inv P H1 H2 (Fin.L n2 i) = H1 i.
fin_add_inv P H1 H2 (Fin.L n2 i) = H1 i.
Proof.
revert P H1 H2 i.
induction n1 as [|n1 IH]; intros P H1 H2 i; inv_fin i; simpl; auto.
intros i. apply (IH (λ i, P (FS i))).
Qed.
Lemma fin_plus_inv_R {n1 n2} (P : fin (n1 + n2) Type)
Lemma fin_add_inv_r {n1 n2} (P : fin (n1 + n2) Type)
(H1: i1 : fin n1, P (Fin.L _ i1)) (H2: i2, P (Fin.R _ i2)) (i: fin n2) :
fin_plus_inv P H1 H2 (Fin.R n1 i) = H2 i.
fin_add_inv P H1 H2 (Fin.R n1 i) = H2 i.
Proof.
revert P H1 H2 i; induction n1 as [|n1 IH]; intros P H1 H2 i; simpl; auto.
apply (IH (λ i, P (FS i))).
......
......@@ -9,249 +9,455 @@ Set Default Proof Using "Type*".
Class FinMapDom K M D `{ A, Dom (M A) D, FMap M,
A, Lookup K A (M A), A, Empty (M A), A, PartialAlter K A (M A),
OMap M, Merge M, A, FinMapToList K A (M A), EqDecision K,
OMap M, Merge M, A, MapFold K A (M A), EqDecision K,
ElemOf K D, Empty D, Singleton K D,
Union D, Intersection D, Difference D} := {
finmap_dom_map :> FinMap K M;
finmap_dom_set :> Set_ K D;
elem_of_dom {A} (m : M A) i : i dom D m is_Some (m !! i)
finmap_dom_map :: FinMap K M;
finmap_dom_set :: Set_ K D;
elem_of_dom {A} (m : M A) i : i dom m is_Some (m !! i)
}.
Section fin_map_dom.
Context `{FinMapDom K M D}.
Lemma lookup_lookup_total_dom `{!Inhabited A} (m : M A) i :
i dom D m m !! i = Some (m !!! i).
i dom m m !! i = Some (m !!! i).
Proof. rewrite elem_of_dom. apply lookup_lookup_total. Qed.
Lemma dom_filter {A} (P : K * A Prop) `{!∀ x, Decision (P x)} (m : M A) X :
( i, i X x, m !! i = Some x P (i, x))
dom D (filter P m) X.
Proof.
intros HX i. rewrite elem_of_dom, HX.
unfold is_Some. by setoid_rewrite map_filter_lookup_Some.
Qed.
Lemma dom_filter_subseteq {A} (P : K * A Prop) `{!∀ x, Decision (P x)} (m : M A):
dom D (filter P m) dom D m.
Proof.
intros ?. rewrite 2!elem_of_dom.
destruct 1 as [?[Eq _]%map_filter_lookup_Some]. by eexists.
Qed.
Lemma dom_imap_subseteq {A B} (f: K A option B) (m: M A) :
dom D (map_imap f m) dom D m.
dom (map_imap f m) dom m.
Proof.
intros k. rewrite 2!elem_of_dom, map_lookup_imap.
destruct 1 as [?[?[Eq _]]%bind_Some]. by eexists.
Qed.
Lemma dom_imap {A B} (f: K A option B) (m: M A) X :
Lemma dom_imap {A B} (f : K A option B) (m : M A) (X : D) :
( i, i X x, m !! i = Some x is_Some (f i x))
dom D (map_imap f m) X.
dom (map_imap f m) X.
Proof.
intros HX k. rewrite elem_of_dom, HX, map_lookup_imap.
unfold is_Some. setoid_rewrite bind_Some. naive_solver.
Qed.
Lemma elem_of_dom_2 {A} (m : M A) i x : m !! i = Some x i dom D m.
Lemma elem_of_dom_2 {A} (m : M A) i x : m !! i = Some x i dom m.
Proof. rewrite elem_of_dom; eauto. Qed.
Lemma not_elem_of_dom {A} (m : M A) i : i dom D m m !! i = None.
Lemma not_elem_of_dom {A} (m : M A) i : i dom m m !! i = None.
Proof. by rewrite elem_of_dom, eq_None_not_Some. Qed.
Lemma subseteq_dom {A} (m1 m2 : M A) : m1 m2 dom D m1 dom D m2.
Lemma not_elem_of_dom_1 {A} (m : M A) i : i dom m m !! i = None.
Proof. apply not_elem_of_dom. Qed.
Lemma not_elem_of_dom_2 {A} (m : M A) i : m !! i = None i dom m.
Proof. apply not_elem_of_dom. Qed.
Lemma subseteq_dom {A} (m1 m2 : M A) : m1 m2 dom m1 dom m2.
Proof.
rewrite map_subseteq_spec.
intros ??. rewrite !elem_of_dom. inversion 1; eauto.
intros ??. rewrite !elem_of_dom. inv 1; eauto.
Qed.
Lemma subset_dom {A} (m1 m2 : M A) : m1 m2 dom D m1 dom D m2.
Lemma subset_dom {A} (m1 m2 : M A) : m1 m2 dom m1 dom m2.
Proof.
intros [Hss1 Hss2]; split; [by apply subseteq_dom |].
contradict Hss2. rewrite map_subseteq_spec. intros i x Hi.
specialize (Hss2 i). rewrite !elem_of_dom in Hss2.
destruct Hss2; eauto. by simplify_map_eq.
Qed.
Lemma dom_empty {A} : dom D (@empty (M A) _) ∅.
Lemma dom_filter {A} (P : K * A Prop) `{!∀ x, Decision (P x)} (m : M A) (X : D) :
( i, i X x, m !! i = Some x P (i, x))
dom (filter P m) X.
Proof.
intros HX i. rewrite elem_of_dom, HX.
unfold is_Some. by setoid_rewrite map_lookup_filter_Some.
Qed.
Lemma dom_filter_subseteq {A} (P : K * A Prop) `{!∀ x, Decision (P x)} (m : M A):
dom (filter P m) dom m.
Proof. apply subseteq_dom, map_filter_subseteq. Qed.
Lemma filter_dom {A} `{!Elements K D, !FinSet K D}
(P : K Prop) `{!∀ x, Decision (P x)} (m : M A) :
filter P (dom m) dom (filter (λ kv, P kv.1) m).
Proof.
intros i. rewrite elem_of_filter, !elem_of_dom. unfold is_Some.
setoid_rewrite map_lookup_filter_Some. naive_solver.
Qed.
Lemma dom_empty {A} : dom (@empty (M A) _) ∅.
Proof.
intros x. rewrite elem_of_dom, lookup_empty, <-not_eq_None_Some. set_solver.
Qed.
Lemma dom_empty_inv {A} (m : M A) : dom D m m = ∅.
Lemma dom_empty_iff {A} (m : M A) : dom m m = ∅.
Proof.
split; [|intros ->; by rewrite dom_empty].
intros E. apply map_empty. intros. apply not_elem_of_dom.
rewrite E. set_solver.
Qed.
Lemma dom_alter {A} f (m : M A) i : dom D (alter f i m) dom D m.
Lemma dom_empty_inv {A} (m : M A) : dom m m = ∅.
Proof. apply dom_empty_iff. Qed.
Lemma dom_alter {A} f (m : M A) i : dom (alter f i m) dom m.
Proof.
apply elem_of_equiv; intros j; rewrite !elem_of_dom; unfold is_Some.
apply set_equiv; intros j; rewrite !elem_of_dom; unfold is_Some.
destruct (decide (i = j)); simplify_map_eq/=; eauto.
destruct (m !! j); naive_solver.
Qed.
Lemma dom_insert {A} (m : M A) i x : dom D (<[i:=x]>m) {[ i ]} dom D m.
Lemma dom_insert {A} (m : M A) i x : dom (<[i:=x]>m) {[ i ]} dom m.
Proof.
apply elem_of_equiv. intros j. rewrite elem_of_union, !elem_of_dom.
apply set_equiv. intros j. rewrite elem_of_union, !elem_of_dom.
unfold is_Some. setoid_rewrite lookup_insert_Some.
destruct (decide (i = j)); set_solver.
Qed.
Lemma dom_insert_subseteq {A} (m : M A) i x : dom D m dom D (<[i:=x]>m).
Lemma dom_insert_lookup {A} (m : M A) i x :
is_Some (m !! i) dom (<[i:=x]>m) dom m.
Proof.
intros Hindom. assert (i dom m) by by apply elem_of_dom.
rewrite dom_insert. set_solver.
Qed.
Lemma dom_insert_subseteq {A} (m : M A) i x : dom m dom (<[i:=x]>m).
Proof. rewrite (dom_insert _). set_solver. Qed.
Lemma dom_insert_subseteq_compat_l {A} (m : M A) i x X :
X dom D m X dom D (<[i:=x]>m).
Proof. intros. trans (dom D m); eauto using dom_insert_subseteq. Qed.
Lemma dom_singleton {A} (i : K) (x : A) : dom D ({[i := x]} : M A) {[ i ]}.
X dom m X dom (<[i:=x]>m).
Proof. intros. trans (dom m); eauto using dom_insert_subseteq. Qed.
Lemma dom_singleton {A} (i : K) (x : A) : dom ({[i := x]} : M A) {[ i ]}.
Proof. rewrite <-insert_empty, dom_insert, dom_empty; set_solver. Qed.
Lemma dom_delete {A} (m : M A) i : dom D (delete i m) dom D m {[ i ]}.
Lemma dom_delete {A} (m : M A) i : dom (delete i m) dom m {[ i ]}.
Proof.
apply elem_of_equiv. intros j. rewrite elem_of_difference, !elem_of_dom.
apply set_equiv. intros j. rewrite elem_of_difference, !elem_of_dom.
unfold is_Some. setoid_rewrite lookup_delete_Some. set_solver.
Qed.
Lemma delete_partial_alter_dom {A} (m : M A) i f :
i dom D m delete i (partial_alter f i m) = m.
i dom m delete i (partial_alter f i m) = m.
Proof. rewrite not_elem_of_dom. apply delete_partial_alter. Qed.
Lemma delete_insert_dom {A} (m : M A) i x :
i dom D m delete i (<[i:=x]>m) = m.
i dom m delete i (<[i:=x]>m) = m.
Proof. rewrite not_elem_of_dom. apply delete_insert. Qed.
Lemma map_disjoint_dom {A} (m1 m2 : M A) : m1 ## m2 dom D m1 ## dom D m2.
Lemma map_disjoint_dom {A} (m1 m2 : M A) : m1 ## m2 dom m1 ## dom m2.
Proof.
rewrite map_disjoint_spec, elem_of_disjoint.
setoid_rewrite elem_of_dom. unfold is_Some. naive_solver.
Qed.
Lemma map_disjoint_dom_1 {A} (m1 m2 : M A) : m1 ## m2 dom D m1 ## dom D m2.
Lemma map_disjoint_dom_1 {A} (m1 m2 : M A) : m1 ## m2 dom m1 ## dom m2.
Proof. apply map_disjoint_dom. Qed.
Lemma map_disjoint_dom_2 {A} (m1 m2 : M A) : dom D m1 ## dom D m2 m1 ## m2.
Lemma map_disjoint_dom_2 {A} (m1 m2 : M A) : dom m1 ## dom m2 m1 ## m2.
Proof. apply map_disjoint_dom. Qed.
Lemma dom_union {A} (m1 m2 : M A) : dom D (m1 m2) dom D m1 dom D m2.
Lemma dom_union {A} (m1 m2 : M A) : dom (m1 m2) dom m1 dom m2.
Proof.
apply elem_of_equiv. intros i. rewrite elem_of_union, !elem_of_dom.
apply set_equiv. intros i. rewrite elem_of_union, !elem_of_dom.
unfold is_Some. setoid_rewrite lookup_union_Some_raw.
destruct (m1 !! i); naive_solver.
Qed.
Lemma dom_intersection {A} (m1 m2: M A) : dom D (m1 m2) dom D m1 dom D m2.
Lemma dom_intersection {A} (m1 m2: M A) : dom (m1 m2) dom m1 dom m2.
Proof.
apply elem_of_equiv. intros i. rewrite elem_of_intersection, !elem_of_dom.
apply set_equiv. intros i. rewrite elem_of_intersection, !elem_of_dom.
unfold is_Some. setoid_rewrite lookup_intersection_Some. naive_solver.
Qed.
Lemma dom_difference {A} (m1 m2 : M A) : dom D (m1 m2) dom D m1 dom D m2.
Lemma dom_difference {A} (m1 m2 : M A) : dom (m1 m2) dom m1 dom m2.
Proof.
apply elem_of_equiv. intros i. rewrite elem_of_difference, !elem_of_dom.
apply set_equiv. intros i. rewrite elem_of_difference, !elem_of_dom.
unfold is_Some. setoid_rewrite lookup_difference_Some.
destruct (m2 !! i); naive_solver.
Qed.
Lemma dom_fmap {A B} (f : A B) (m : M A) : dom D (f <$> m) dom D m.
Lemma dom_fmap {A B} (f : A B) (m : M A) : dom (f <$> m) dom m.
Proof.
apply elem_of_equiv. intros i.
apply set_equiv. intros i.
rewrite !elem_of_dom, lookup_fmap, <-!not_eq_None_Some.
destruct (m !! i); naive_solver.
Qed.
Lemma dom_finite {A} (m : M A) : set_finite (dom D m).
Lemma dom_finite {A} (m : M A) : set_finite (dom m).
Proof.
induction m using map_ind; rewrite ?dom_empty, ?dom_insert.
- by apply empty_finite.
- apply union_finite; [apply singleton_finite|done].
Qed.
Global Instance dom_proper `{!Equiv A} : Proper ((≡@{M A}) ==> ()) (dom D).
Global Instance dom_proper `{!Equiv A} : Proper ((≡@{M A}) ==> ()) dom.
Proof.
intros m1 m2 EQm. apply elem_of_equiv. intros i.
intros m1 m2 EQm. apply set_equiv. intros i.
rewrite !elem_of_dom, EQm. done.
Qed.
Lemma dom_list_to_map {A} (l : list (K * A)) :
dom (list_to_map l : M A) list_to_set l.*1.
Proof.
induction l as [|?? IH].
- by rewrite dom_empty.
- simpl. by rewrite dom_insert, IH.
Qed.
Lemma map_first_key_dom {A B} (m1 : M A) (m2 : M B) i :
dom m1 dom m2 map_first_key m1 i map_first_key m2 i.
Proof.
intros Hm. apply map_first_key_dom'. intros j.
by rewrite <-!elem_of_dom, Hm.
Qed.
Lemma map_first_key_dom_L {A B} (m1 : M A) (m2 : M B) i :
dom m1 = dom m2 map_first_key m1 i map_first_key m2 i.
Proof. intros Hm. apply map_first_key_dom. by rewrite Hm. Qed.
Lemma map_Forall2_dom {A B} (P : K A B Prop) (m1 : M A) (m2 : M B) :
map_Forall2 P m1 m2 dom m1 dom m2.
Proof.
revert m2. induction m1 as [|i x1 m1 ? IH] using map_ind; intros m2.
{ intros ->%map_Forall2_empty_inv_l. by rewrite !dom_empty. }
intros (x2 & m2' & -> & ? & ? & ?)%map_Forall2_insert_inv_l; last done.
by rewrite !dom_insert, IH by done.
Qed.
(** Alternative definition of [dom] in terms of [map_to_list]. *)
Lemma dom_alt {A} (m : M A) :
dom m list_to_set (map_to_list m).*1.
Proof.
rewrite <-(list_to_map_to_list m) at 1.
rewrite dom_list_to_map.
done.
Qed.
Lemma size_dom `{!Elements K D, !FinSet K D} {A} (m : M A) :
size (dom m) = size m.
Proof.
induction m as [|i x m ? IH] using map_ind.
{ by rewrite dom_empty, map_size_empty, size_empty. }
assert ({[i]} ## dom m).
{ intros j. rewrite elem_of_dom. unfold is_Some. set_solver. }
by rewrite dom_insert, size_union, size_singleton, map_size_insert_None, IH.
Qed.
Lemma dom_subseteq_size {A} (m1 m2 : M A) : dom m2 dom m1 size m2 size m1.
Proof.
revert m1. induction m2 as [|i x m2 ? IH] using map_ind; intros m1 Hdom.
{ rewrite map_size_empty. lia. }
rewrite dom_insert in Hdom.
assert (i dom m2) by (by apply not_elem_of_dom).
assert (i dom m1) as [x' Hx']%elem_of_dom by set_solver.
rewrite <-(insert_delete m1 i x') by done.
rewrite !map_size_insert_None, <-Nat.succ_le_mono by (by rewrite ?lookup_delete).
apply IH. rewrite dom_delete. set_solver.
Qed.
Lemma dom_subset_size {A} (m1 m2 : M A) : dom m2 dom m1 size m2 < size m1.
Proof.
revert m1. induction m2 as [|i x m2 ? IH] using map_ind; intros m1 Hdom.
{ destruct m1 as [|i x m1 ? _] using map_ind.
- rewrite !dom_empty in Hdom. set_solver.
- rewrite map_size_empty, map_size_insert_None by done. lia. }
rewrite dom_insert in Hdom.
assert (i dom m2) by (by apply not_elem_of_dom).
assert (i dom m1) as [x' Hx']%elem_of_dom by set_solver.
rewrite <-(insert_delete m1 i x') by done.
rewrite !map_size_insert_None, <-Nat.succ_lt_mono by (by rewrite ?lookup_delete).
apply IH. rewrite dom_delete. split; [set_solver|].
intros ?. destruct Hdom as [? []].
intros j. destruct (decide (i = j)); set_solver.
Qed.
Lemma subseteq_dom_eq {A} (m1 m2 : M A) :
m1 m2 dom m2 dom m1 m1 = m2.
Proof. intros. apply map_subseteq_size_eq; auto using dom_subseteq_size. Qed.
Lemma dom_singleton_inv {A} (m : M A) i :
dom m {[i]} x, m = {[i := x]}.
Proof.
intros Hdom. assert (is_Some (m !! i)) as [x ?].
{ apply (elem_of_dom (D:=D)); set_solver. }
exists x. apply map_eq; intros j.
destruct (decide (i = j)); simplify_map_eq; [done|].
apply not_elem_of_dom. set_solver.
Qed.
Lemma dom_map_zip_with {A B C} (f : A B C) (ma : M A) (mb : M B) :
dom (map_zip_with f ma mb) dom ma dom mb.
Proof.
rewrite set_equiv. intros x.
rewrite elem_of_intersection, !elem_of_dom, map_lookup_zip_with.
destruct (ma !! x), (mb !! x); rewrite !is_Some_alt; naive_solver.
Qed.
Lemma dom_union_inv `{!RelDecision (∈@{D})} {A} (m : M A) (X1 X2 : D) :
X1 ## X2
dom m X1 X2
m1 m2, m = m1 m2 m1 ## m2 dom m1 X1 dom m2 X2.
Proof.
intros.
exists (filter (λ '(k,x), k X1) m), (filter (λ '(k,x), k X1) m).
assert (filter (λ '(k, _), k X1) m ## filter (λ '(k, _), k X1) m).
{ apply map_disjoint_filter_complement. }
split_and!; [|done| |].
- apply map_eq; intros i. apply option_eq; intros x.
rewrite lookup_union_Some, !map_lookup_filter_Some by done.
destruct (decide (i X1)); naive_solver.
- apply dom_filter; intros i; split; [|naive_solver].
intros. assert (is_Some (m !! i)) as [x ?] by (apply elem_of_dom; set_solver).
naive_solver.
- apply dom_filter; intros i; split.
+ intros. assert (is_Some (m !! i)) as [x ?] by (apply elem_of_dom; set_solver).
naive_solver.
+ intros (x&?&?). apply dec_stable; intros ?.
assert (m !! i = None) by (apply not_elem_of_dom; set_solver).
naive_solver.
Qed.
Lemma dom_kmap `{!Elements K D, !FinSet K D, FinMapDom K2 M2 D2}
{A} (f : K K2) `{!Inj (=) (=) f} (m : M A) :
dom (kmap (M2:=M2) f m) ≡@{D2} set_map f (dom m).
Proof.
apply set_equiv. intros i.
rewrite !elem_of_dom, (lookup_kmap_is_Some _), elem_of_map.
by setoid_rewrite elem_of_dom.
Qed.
Lemma dom_omap_subseteq {A B} (f : A option B) (m : M A) :
dom (omap f m) dom m.
Proof.
intros a. rewrite !elem_of_dom. intros [c Hm].
apply lookup_omap_Some in Hm. naive_solver.
Qed.
Lemma map_compose_dom_subseteq {C} `{FinMap K' M'} (m: M' C) (n : M K') :
dom (m n : M C) ⊆@{D} dom n.
Proof. apply dom_omap_subseteq. Qed.
Lemma map_compose_min_r_dom {C} `{FinMap K' M', !RelDecision (∈@{D})}
(m : M C) (n : M' K) :
m n = m filter (λ '(_,b), b dom m) n.
Proof.
rewrite map_compose_min_r. f_equal.
apply map_filter_ext. intros. by rewrite elem_of_dom.
Qed.
Lemma map_compose_empty_iff_dom_img {C} `{FinMap K' M', !RelDecision (∈@{D})}
(m : M C) (n : M' K) :
m n = dom m ## map_img n.
Proof.
rewrite map_compose_empty_iff, elem_of_disjoint.
setoid_rewrite elem_of_dom. setoid_rewrite eq_None_not_Some.
setoid_rewrite elem_of_map_img. naive_solver.
Qed.
(** If [D] has Leibniz equality, we can show an even stronger result. This is a
common case e.g. when having a [gmap K A] where the key [K] has Leibniz equality
(and thus also [gset K], the usual domain) but the value type [A] does not. *)
Global Instance dom_proper_L `{!Equiv A, !LeibnizEquiv D} :
Proper ((≡@{M A}) ==> (=)) (dom D) | 0.
Proper ((≡@{M A}) ==> (=)) (dom) | 0.
Proof. intros ???. unfold_leibniz. by apply dom_proper. Qed.
Section leibniz.
Context `{!LeibnizEquiv D}.
Lemma dom_filter_L {A} (P : K * A Prop) `{!∀ x, Decision (P x)} (m : M A) X :
( i, i X x, m !! i = Some x P (i, x))
dom D (filter P m) = X.
dom (filter P m) = X.
Proof. unfold_leibniz. apply dom_filter. Qed.
Lemma dom_empty_L {A} : dom D (@empty (M A) _) = ∅.
Lemma filter_dom_L {A} `{!Elements K D, !FinSet K D}
(P : K Prop) `{!∀ x, Decision (P x)} (m : M A) :
filter P (dom m) = dom (filter (λ kv, P kv.1) m).
Proof. unfold_leibniz. apply filter_dom. Qed.
Lemma dom_empty_L {A} : dom (@empty (M A) _) = ∅.
Proof. unfold_leibniz; apply dom_empty. Qed.
Lemma dom_empty_inv_L {A} (m : M A) : dom D m = m = ∅.
Lemma dom_empty_iff_L {A} (m : M A) : dom m = m = ∅.
Proof. unfold_leibniz. apply dom_empty_iff. Qed.
Lemma dom_empty_inv_L {A} (m : M A) : dom m = m = ∅.
Proof. by intros; apply dom_empty_inv; unfold_leibniz. Qed.
Lemma dom_alter_L {A} f (m : M A) i : dom D (alter f i m) = dom D m.
Lemma dom_alter_L {A} f (m : M A) i : dom (alter f i m) = dom m.
Proof. unfold_leibniz; apply dom_alter. Qed.
Lemma dom_insert_L {A} (m : M A) i x : dom D (<[i:=x]>m) = {[ i ]} dom D m.
Lemma dom_insert_L {A} (m : M A) i x : dom (<[i:=x]>m) = {[ i ]} dom m.
Proof. unfold_leibniz; apply dom_insert. Qed.
Lemma dom_singleton_L {A} (i : K) (x : A) : dom D ({[i := x]} : M A) = {[ i ]}.
Lemma dom_insert_lookup_L {A} (m : M A) i x :
is_Some (m !! i) dom (<[i:=x]>m) = dom m.
Proof. unfold_leibniz; apply dom_insert_lookup. Qed.
Lemma dom_singleton_L {A} (i : K) (x : A) : dom ({[i := x]} : M A) = {[ i ]}.
Proof. unfold_leibniz; apply dom_singleton. Qed.
Lemma dom_delete_L {A} (m : M A) i : dom D (delete i m) = dom D m {[ i ]}.
Lemma dom_delete_L {A} (m : M A) i : dom (delete i m) = dom m {[ i ]}.
Proof. unfold_leibniz; apply dom_delete. Qed.
Lemma dom_union_L {A} (m1 m2 : M A) : dom D (m1 m2) = dom D m1 dom D m2.
Lemma dom_union_L {A} (m1 m2 : M A) : dom (m1 m2) = dom m1 dom m2.
Proof. unfold_leibniz; apply dom_union. Qed.
Lemma dom_intersection_L {A} (m1 m2 : M A) :
dom D (m1 m2) = dom D m1 dom D m2.
dom (m1 m2) = dom m1 dom m2.
Proof. unfold_leibniz; apply dom_intersection. Qed.
Lemma dom_difference_L {A} (m1 m2 : M A) : dom D (m1 m2) = dom D m1 dom D m2.
Lemma dom_difference_L {A} (m1 m2 : M A) : dom (m1 m2) = dom m1 dom m2.
Proof. unfold_leibniz; apply dom_difference. Qed.
Lemma dom_fmap_L {A B} (f : A B) (m : M A) : dom D (f <$> m) = dom D m.
Lemma dom_fmap_L {A B} (f : A B) (m : M A) : dom (f <$> m) = dom m.
Proof. unfold_leibniz; apply dom_fmap. Qed.
Lemma map_Forall2_dom_L {A B} (P : K A B Prop) (m1 : M A) (m2 : M B) :
map_Forall2 P m1 m2 dom m1 = dom m2.
Proof. unfold_leibniz. apply map_Forall2_dom. Qed.
Lemma dom_imap_L {A B} (f: K A option B) (m: M A) X :
( i, i X x, m !! i = Some x is_Some (f i x))
dom D (map_imap f m) = X.
dom (map_imap f m) = X.
Proof. unfold_leibniz; apply dom_imap. Qed.
Lemma dom_list_to_map_L {A} (l : list (K * A)) :
dom (list_to_map l : M A) = list_to_set l.*1.
Proof. unfold_leibniz. apply dom_list_to_map. Qed.
Lemma dom_singleton_inv_L {A} (m : M A) i :
dom m = {[i]} x, m = {[i := x]}.
Proof. unfold_leibniz. apply dom_singleton_inv. Qed.
Lemma dom_map_zip_with_L {A B C} (f : A B C) (ma : M A) (mb : M B) :
dom (map_zip_with f ma mb) = dom ma dom mb.
Proof. unfold_leibniz. apply dom_map_zip_with. Qed.
Lemma dom_union_inv_L `{!RelDecision (∈@{D})} {A} (m : M A) (X1 X2 : D) :
X1 ## X2
dom m = X1 X2
m1 m2, m = m1 m2 m1 ## m2 dom m1 = X1 dom m2 = X2.
Proof. unfold_leibniz. apply dom_union_inv. Qed.
End leibniz.
Lemma dom_kmap_L `{!Elements K D, !FinSet K D, FinMapDom K2 M2 D2}
`{!LeibnizEquiv D2} {A} (f : K K2) `{!Inj (=) (=) f} (m : M A) :
dom (kmap (M2:=M2) f m) = set_map f (dom m).
Proof. unfold_leibniz. by apply dom_kmap. Qed.
(** * Set solver instances *)
Global Instance set_unfold_dom_empty {A} i : SetUnfoldElemOf i (dom D (∅:M A)) False.
Global Instance set_unfold_dom_empty {A} i : SetUnfoldElemOf i (dom (∅:M A)) False.
Proof. constructor. by rewrite dom_empty, elem_of_empty. Qed.
Global Instance set_unfold_dom_alter {A} f i j (m : M A) Q :
SetUnfoldElemOf i (dom D m) Q
SetUnfoldElemOf i (dom D (alter f j m)) Q.
Proof. constructor. by rewrite dom_alter, (set_unfold_elem_of _ (dom _ _) _). Qed.
SetUnfoldElemOf i (dom m) Q
SetUnfoldElemOf i (dom (alter f j m)) Q.
Proof. constructor. by rewrite dom_alter, (set_unfold_elem_of _ (dom _) _). Qed.
Global Instance set_unfold_dom_insert {A} i j x (m : M A) Q :
SetUnfoldElemOf i (dom D m) Q
SetUnfoldElemOf i (dom D (<[j:=x]> m)) (i = j Q).
SetUnfoldElemOf i (dom m) Q
SetUnfoldElemOf i (dom (<[j:=x]> m)) (i = j Q).
Proof.
constructor. by rewrite dom_insert, elem_of_union,
(set_unfold_elem_of _ (dom _ _) _), elem_of_singleton.
(set_unfold_elem_of _ (dom _) _), elem_of_singleton.
Qed.
Global Instance set_unfold_dom_delete {A} i j (m : M A) Q :
SetUnfoldElemOf i (dom D m) Q
SetUnfoldElemOf i (dom D (delete j m)) (Q i j).
SetUnfoldElemOf i (dom m) Q
SetUnfoldElemOf i (dom (delete j m)) (Q i j).
Proof.
constructor. by rewrite dom_delete, elem_of_difference,
(set_unfold_elem_of _ (dom _ _) _), elem_of_singleton.
(set_unfold_elem_of _ (dom _) _), elem_of_singleton.
Qed.
Global Instance set_unfold_dom_singleton {A} i j :
SetUnfoldElemOf i (dom D ({[ j := x ]} : M A)) (i = j).
Global Instance set_unfold_dom_singleton {A} i j x :
SetUnfoldElemOf i (dom ({[ j := x ]} : M A)) (i = j).
Proof. constructor. by rewrite dom_singleton, elem_of_singleton. Qed.
Global Instance set_unfold_dom_union {A} i (m1 m2 : M A) Q1 Q2 :
SetUnfoldElemOf i (dom D m1) Q1 SetUnfoldElemOf i (dom D m2) Q2
SetUnfoldElemOf i (dom D (m1 m2)) (Q1 Q2).
SetUnfoldElemOf i (dom m1) Q1 SetUnfoldElemOf i (dom m2) Q2
SetUnfoldElemOf i (dom (m1 m2)) (Q1 Q2).
Proof.
constructor. by rewrite dom_union, elem_of_union,
!(set_unfold_elem_of _ (dom _ _) _).
!(set_unfold_elem_of _ (dom _) _).
Qed.
Global Instance set_unfold_dom_intersection {A} i (m1 m2 : M A) Q1 Q2 :
SetUnfoldElemOf i (dom D m1) Q1 SetUnfoldElemOf i (dom D m2) Q2
SetUnfoldElemOf i (dom D (m1 m2)) (Q1 Q2).
SetUnfoldElemOf i (dom m1) Q1 SetUnfoldElemOf i (dom m2) Q2
SetUnfoldElemOf i (dom (m1 m2)) (Q1 Q2).
Proof.
constructor. by rewrite dom_intersection, elem_of_intersection,
!(set_unfold_elem_of _ (dom _ _) _).
!(set_unfold_elem_of _ (dom _) _).
Qed.
Global Instance set_unfold_dom_difference {A} i (m1 m2 : M A) Q1 Q2 :
SetUnfoldElemOf i (dom D m1) Q1 SetUnfoldElemOf i (dom D m2) Q2
SetUnfoldElemOf i (dom D (m1 m2)) (Q1 ¬Q2).
SetUnfoldElemOf i (dom m1) Q1 SetUnfoldElemOf i (dom m2) Q2
SetUnfoldElemOf i (dom (m1 m2)) (Q1 ¬Q2).
Proof.
constructor. by rewrite dom_difference, elem_of_difference,
!(set_unfold_elem_of _ (dom _ _) _).
!(set_unfold_elem_of _ (dom _) _).
Qed.
Global Instance set_unfold_dom_fmap {A B} (f : A B) i (m : M A) Q :
SetUnfoldElemOf i (dom D m) Q
SetUnfoldElemOf i (dom D (f <$> m)) Q.
Proof. constructor. by rewrite dom_fmap, (set_unfold_elem_of _ (dom _ _) _). Qed.
SetUnfoldElemOf i (dom m) Q
SetUnfoldElemOf i (dom (f <$> m)) Q.
Proof. constructor. by rewrite dom_fmap, (set_unfold_elem_of _ (dom _) _). Qed.
End fin_map_dom.
Lemma dom_seq `{FinMapDom nat M D} {A} start (xs : list A) :
dom D (map_seq start (M:=M A) xs) set_seq start (length xs).
dom (map_seq start (M:=M A) xs) set_seq start (length xs).
Proof.
revert start. induction xs as [|x xs IH]; intros start; simpl.
- by rewrite dom_empty.
- by rewrite dom_insert, IH.
Qed.
Lemma dom_seq_L `{FinMapDom nat M D, !LeibnizEquiv D} {A} start (xs : list A) :
dom D (map_seq (M:=M A) start xs) = set_seq start (length xs).
dom (map_seq (M:=M A) start xs) = set_seq start (length xs).
Proof. unfold_leibniz. apply dom_seq. Qed.
Instance set_unfold_dom_seq `{FinMapDom nat M D} {A} start (xs : list A) :
SetUnfoldElemOf i (dom D (map_seq start (M:=M A) xs)) (start i < start + length xs).
Global Instance set_unfold_dom_seq `{FinMapDom nat M D} {A} start (xs : list A) i :
SetUnfoldElemOf i (dom (map_seq start (M:=M A) xs)) (start i < start + length xs).
Proof. constructor. by rewrite dom_seq, elem_of_set_seq. Qed.
source diff could not be displayed: it is too large. Options to address this: view the blob.
......@@ -9,24 +9,39 @@ From stdpp Require Import options.
Set Default Proof Using "Type*".
(** Operations *)
Instance set_size `{Elements A C} : Size C := length elements.
Global Instance set_size `{Elements A C} : Size C := length elements.
Global Typeclasses Opaque set_size.
Definition set_fold `{Elements A C} {B}
(f : A B B) (b : B) : C B := foldr f b elements.
Global Typeclasses Opaque set_fold.
Instance set_filter
Global Instance set_filter
`{Elements A C, Empty C, Singleton A C, Union C} : Filter A C := λ P _ X,
list_to_set (filter P (elements X)).
Typeclasses Opaque set_filter.
Global Typeclasses Opaque set_filter.
Definition set_map `{Elements A C, Singleton B D, Empty D, Union D}
(f : A B) (X : C) : D :=
list_to_set (f <$> elements X).
Typeclasses Opaque set_map.
Instance set_fresh `{Elements A C, Fresh A (list A)} : Fresh A C :=
Global Typeclasses Opaque set_map.
Global Instance: Params (@set_map) 8 := {}.
Definition set_bind `{Elements A SA, Empty SB, Union SB}
(f : A SB) (X : SA) : SB :=
(f <$> elements X).
Global Typeclasses Opaque set_bind.
Global Instance: Params (@set_bind) 6 := {}.
Definition set_omap `{Elements A C, Singleton B D, Empty D, Union D}
(f : A option B) (X : C) : D :=
list_to_set (omap f (elements X)).
Global Typeclasses Opaque set_omap.
Global Instance: Params (@set_omap) 8 := {}.
Global Instance set_fresh `{Elements A C, Fresh A (list A)} : Fresh A C :=
fresh elements.
Typeclasses Opaque set_filter.
Global Typeclasses Opaque set_fresh.
(** We generalize the [fresh] operation on sets to generate lists of fresh
elements w.r.t. a set [X]. *)
......@@ -36,7 +51,7 @@ Fixpoint fresh_list `{Fresh A C, Union C, Singleton A C}
| 0 => []
| S n => let x := fresh X in x :: fresh_list n ({[ x ]} X)
end.
Instance: Params (@fresh_list) 6 := {}.
Global Instance: Params (@fresh_list) 6 := {}.
(** The following inductive predicate classifies that a list of elements is
in fact fresh w.r.t. a set [X]. *)
......@@ -53,7 +68,7 @@ Implicit Types X Y : C.
Lemma fin_set_finite X : set_finite X.
Proof. by exists (elements X); intros; rewrite elem_of_elements. Qed.
Instance elem_of_dec_slow : RelDecision (∈@{C}) | 100.
Local Instance elem_of_dec_slow : RelDecision (∈@{C}) | 100.
Proof.
refine (λ x X, cast_if (decide_rel () x (elements X)));
by rewrite <-(elem_of_elements _).
......@@ -77,16 +92,14 @@ Proof.
apply elem_of_nil_inv; intros x.
rewrite elem_of_elements, elem_of_empty; tauto.
Qed.
Lemma elements_empty_inv X : elements X = [] X ∅.
Proof.
intros HX; apply elem_of_equiv_empty; intros x.
rewrite <-elem_of_elements, HX, elem_of_nil. tauto.
Qed.
Lemma elements_empty' X : elements X = [] X ∅.
Lemma elements_empty_iff X : elements X = [] X ∅.
Proof.
split; intros HX; [by apply elements_empty_inv|].
by rewrite <-Permutation_nil, HX, elements_empty.
rewrite <-Permutation_nil_r. split; [|intros ->; by rewrite elements_empty].
intros HX. apply elem_of_equiv_empty; intros x.
rewrite <-elem_of_elements, HX. apply not_elem_of_nil.
Qed.
Lemma elements_empty_inv X : elements X = [] X ∅.
Proof. apply elements_empty_iff. Qed.
Lemma elements_union_singleton (X : C) x :
x X elements ({[ x ]} X) x :: elements X.
......@@ -99,7 +112,7 @@ Proof.
Qed.
Lemma elements_singleton x : elements ({[ x ]} : C) = [x].
Proof.
apply Permutation_singleton. by rewrite <-(right_id () {[x]}),
apply Permutation_singleton_r. by rewrite <-(right_id () {[x]}),
elements_union_singleton, elements_empty by set_solver.
Qed.
Lemma elements_disj_union (X Y : C) :
......@@ -116,19 +129,34 @@ Proof.
intros x. rewrite !elem_of_elements; auto.
Qed.
Lemma list_to_set_elements X : list_to_set (elements X) X.
Proof. intros ?. rewrite elem_of_list_to_set. apply elem_of_elements. Qed.
Lemma list_to_set_elements_L `{!LeibnizEquiv C} X : list_to_set (elements X) = X.
Proof. unfold_leibniz. apply list_to_set_elements. Qed.
Lemma elements_list_to_set l :
NoDup l elements (list_to_set (C:=C) l) l.
Proof.
intros Hl. induction Hl.
{ rewrite list_to_set_nil. rewrite elements_empty. done. }
rewrite list_to_set_cons, elements_disj_union by set_solver.
rewrite elements_singleton. apply Permutation_skip. done.
Qed.
(** * The [size] operation *)
Global Instance set_size_proper: Proper (() ==> (=)) (@size C _).
Proof. intros ?? E. apply Permutation_length. by rewrite E. Qed.
Lemma size_empty : size ( : C) = 0.
Proof. unfold size, set_size. simpl. by rewrite elements_empty. Qed.
Lemma size_empty_inv (X : C) : size X = 0 X ∅.
Lemma size_empty_iff (X : C) : size X = 0 X ∅.
Proof.
split; [|intros ->; by rewrite size_empty].
intros; apply equiv_empty; intros x; rewrite <-elem_of_elements.
by rewrite (nil_length_inv (elements X)), ?elem_of_nil.
Qed.
Lemma size_empty_iff (X : C) : size X = 0 X ∅.
Proof. split; [apply size_empty_inv|]. by intros ->; rewrite size_empty. Qed.
Lemma size_empty_inv (X : C) : size X = 0 X ∅.
Proof. apply size_empty_iff. Qed.
Lemma size_non_empty_iff (X : C) : size X 0 X ∅.
Proof. by rewrite size_empty_iff. Qed.
......@@ -159,14 +187,14 @@ Qed.
Lemma size_1_elem_of X : size X = 1 x, X {[ x ]}.
Proof.
intros E. destruct (size_pos_elem_of X) as [x ?]; auto with lia.
exists x. apply elem_of_equiv. split.
exists x. apply set_equiv. split.
- rewrite elem_of_singleton. eauto using size_singleton_inv.
- set_solver.
Qed.
Lemma size_union X Y : X ## Y size (X Y) = size X + size Y.
Proof.
intros. unfold size, set_size. simpl. rewrite <-app_length.
intros. unfold size, set_size. simpl. rewrite <-length_app.
apply Permutation_length, NoDup_Permutation.
- apply NoDup_elements.
- apply NoDup_app; repeat split; try apply NoDup_elements.
......@@ -180,6 +208,27 @@ Proof.
rewrite <-union_difference, (comm ()); set_solver.
Qed.
Lemma size_difference X Y : Y X size (X Y) = size X - size Y.
Proof.
intros. rewrite (union_difference Y X) at 2 by done.
rewrite size_union by set_solver. lia.
Qed.
Lemma size_difference_alt X Y : size (X Y) = size X - size (X Y).
Proof.
intros. rewrite <-size_difference by set_solver.
apply set_size_proper. set_solver.
Qed.
Lemma set_subseteq_size_equiv X1 X2 : X1 X2 size X2 size X1 X1 X2.
Proof.
intros. apply (anti_symm _); [done|].
apply empty_difference_subseteq, size_empty_iff.
rewrite size_difference by done. lia.
Qed.
Lemma set_subseteq_size_eq `{!LeibnizEquiv C} X1 X2 :
X1 X2 size X2 size X1 X1 = X2.
Proof. unfold_leibniz. apply set_subseteq_size_equiv. Qed.
Lemma subseteq_size X Y : X Y size X size Y.
Proof. intros. rewrite (union_difference X Y), size_union_alt by done. lia. Qed.
Lemma subset_size X Y : X Y size X < size Y.
......@@ -190,11 +239,18 @@ Proof.
by apply size_non_empty_iff, non_empty_difference.
Qed.
Lemma size_list_to_set l :
NoDup l size (list_to_set (C:=C) l) = length l.
Proof.
intros Hl. unfold size, set_size. simpl.
rewrite elements_list_to_set; done.
Qed.
(** * Induction principles *)
Lemma set_wf : wf (⊂@{C}).
Lemma set_wf : well_founded (⊂@{C}).
Proof. apply (wf_projected (<) size); auto using subset_size, lt_wf. Qed.
Lemma set_ind (P : C Prop) :
Proper (() ==> iff) P
Proper (() ==> impl) P
P ( x X, x X P X P ({[ x ]} X)) X, P X.
Proof.
intros ? Hemp Hadd. apply well_founded_induction with ().
......@@ -210,7 +266,7 @@ Proof. apply set_ind. by intros ?? ->%leibniz_equiv_iff. Qed.
(** * The [set_fold] operation *)
Lemma set_fold_ind {B} (P : B C Prop) (f : A B B) (b : B) :
Proper ((=) ==> () ==> iff) P
( x, Proper (() ==> impl) (P x))
P b ( x X r, x X P r X P (f x r) ({[ x ]} X))
X, P (set_fold f b X) X.
Proof.
......@@ -229,9 +285,9 @@ Lemma set_fold_ind_L `{!LeibnizEquiv C}
{B} (P : B C Prop) (f : A B B) (b : B) :
P b ( x X r, x X P r X P (f x r) ({[ x ]} X))
X, P (set_fold f b X) X.
Proof. apply set_fold_ind. by intros ?? -> ?? ->%leibniz_equiv. Qed.
Lemma set_fold_proper {B} (R : relation B) `{!Equivalence R}
(f : A B B) (b : B) `{!Proper ((=) ==> R ==> R) f}
Proof. apply set_fold_ind. solve_proper. Qed.
Lemma set_fold_proper {B} (R : relation B) `{!PreOrder R}
(f : A B B) (b : B) `{! a, Proper (R ==> R) (f a)}
(Hf : a1 a2 b, R (f a1 (f a2 b)) (f a2 (f a1 b))) :
Proper (() ==> R) (set_fold f b : C B).
Proof. intros ?? E. apply (foldr_permutation R f b); auto. by rewrite E. Qed.
......@@ -242,18 +298,138 @@ Proof. by unfold set_fold; simpl; rewrite elements_empty. Qed.
Lemma set_fold_singleton {B} (f : A B B) (b : B) (a : A) :
set_fold f b ({[a]} : C) = f a b.
Proof. by unfold set_fold; simpl; rewrite elements_singleton. Qed.
(** The following lemma shows that folding over two sets separately (using the
result of the first fold as input for the second fold) is equivalent to folding
over the union, *if* the function is idempotent for the elements that will be
processed twice ([X ∩ Y]) and does not care about the order in which elements
are processed.
This is a generalization of [set_fold_union] (below) with a.) a relation [R]
instead of equality b.) a function [f : A → B → B] instead of [f : A → A → A],
and c.) premises that ensure the elements are in [X ∪ Y]. *)
Lemma set_fold_union_strong {B} (R : relation B) `{!PreOrder R}
(f : A B B) (b : B) X Y :
( x, Proper (R ==> R) (f x))
( x b',
(** This is morally idempotence for elements of [X ∩ Y] *)
x X Y
(** We cannot write this in the usual direction of idempotence properties
(i.e., [R (f x (f x b'))) (f x b')]) because [R] is not symmetric. *)
R (f x b') (f x (f x b')))
( x1 x2 b',
(** This is morally commutativity + associativity for elements of [X ∪ Y] *)
x1 X Y x2 X Y x1 x2
R (f x1 (f x2 b')) (f x2 (f x1 b')))
R (set_fold f b (X Y)) (set_fold f (set_fold f b X) Y).
Proof.
(** This lengthy proof involves various steps by transitivity of [R].
Roughly, we show that the LHS is related to folding over:
elements (Y ∖ X) ++ elements (X ∩ Y) ++ elements (X ∖ Y)
and the RHS is related to folding over:
elements (Y ∖ X) ++ elements (X ∩ Y) ++ elements (X ∩ Y) ++ elements (Y ∖ X)
These steps are justified by lemma [foldr_permutation]. In the middle we
remove the repeated folding over [elements (X ∩ Y)] using [foldr_idemp_strong].
Most of the proof work concerns the side conditions of [foldr_permutation]
and [foldr_idemp_strong], which require relating results about lists and
sets. *)
intros ?.
assert ( b1 b2 l, R b1 b2 R (foldr f b1 l) (foldr f b2 l)) as Hff.
{ intros b1 b2 l Hb. induction l as [|x l]; simpl; [done|]. by f_equiv. }
intros Hfidemp Hfcomm. unfold set_fold; simpl.
trans (foldr f b (elements (Y X) ++ elements (X Y) ++ elements (X Y))).
{ apply (foldr_permutation R f b).
- intros j1 x1 j2 x2 b' Hj Hj1 Hj2. apply Hfcomm.
+ apply elem_of_list_lookup_2 in Hj1. set_solver.
+ apply elem_of_list_lookup_2 in Hj2. set_solver.
+ intros ->. pose proof (NoDup_elements (X Y)).
by eapply Hj, NoDup_lookup.
- rewrite <-!elements_disj_union by set_solver. f_equiv; intros x.
destruct (decide (x X)), (decide (x Y)); set_solver. }
trans (foldr f (foldr f b (elements (X Y) ++ elements (X Y)))
(elements (Y X) ++ elements (X Y))).
{ rewrite !foldr_app. apply Hff. apply (foldr_idemp_strong (flip R)).
- solve_proper.
- intros j a b' ?%elem_of_list_lookup_2. apply Hfidemp. set_solver.
- intros j1 x1 j2 x2 b' Hj Hj1 Hj2. apply Hfcomm.
+ apply elem_of_list_lookup_2 in Hj2. set_solver.
+ apply elem_of_list_lookup_2 in Hj1. set_solver.
+ intros ->. pose proof (NoDup_elements (X Y)).
by eapply Hj, NoDup_lookup. }
trans (foldr f (foldr f b (elements (X Y) ++ elements (X Y))) (elements Y)).
{ apply (foldr_permutation R f _).
- intros j1 x1 j2 x2 b' Hj Hj1 Hj2. apply Hfcomm.
+ apply elem_of_list_lookup_2 in Hj1. set_solver.
+ apply elem_of_list_lookup_2 in Hj2. set_solver.
+ intros ->. assert (NoDup (elements (Y X) ++ elements (X Y))).
{ rewrite <-elements_disj_union by set_solver. apply NoDup_elements. }
by eapply Hj, NoDup_lookup.
- rewrite <-!elements_disj_union by set_solver. f_equiv; intros x.
destruct (decide (x X)); set_solver. }
apply Hff. apply (foldr_permutation R f _).
- intros j1 x1 j2 x2 b' Hj Hj1 Hj2. apply Hfcomm.
+ apply elem_of_list_lookup_2 in Hj1. set_solver.
+ apply elem_of_list_lookup_2 in Hj2. set_solver.
+ intros ->. assert (NoDup (elements (X Y) ++ elements (X Y))).
{ rewrite <-elements_disj_union by set_solver. apply NoDup_elements. }
by eapply Hj, NoDup_lookup.
- rewrite <-!elements_disj_union by set_solver. f_equiv; intros x.
destruct (decide (x Y)); set_solver.
Qed.
Lemma set_fold_union (f : A A A) (b : A) X Y :
IdemP (=) f
Comm (=) f
Assoc (=) f
set_fold f b (X Y) = set_fold f (set_fold f b X) Y.
Proof.
intros. apply (set_fold_union_strong _ _ _ _ _ _).
- intros x b' _. by rewrite (assoc_L f), (idemp f).
- intros x1 x2 b' _ _ _. by rewrite !(assoc_L f), (comm_L f x1).
Qed.
(** Generalization of [set_fold_disj_union] (below) with a.) a relation [R]
instead of equality b.) a function [f : A → B → B] instead of [f : A → A → A],
and c.) premises that ensure the elements are in [X ∪ Y]. *)
Lemma set_fold_disj_union_strong {B} (R : relation B) `{!PreOrder R}
(f : A B B) (b : B) X Y :
( x, Proper (R ==> R) (f x))
( x1 x2 b',
(** This is morally commutativity + associativity for elements of [X ∪ Y] *)
x1 X Y x2 X Y x1 x2
R (f x1 (f x2 b')) (f x2 (f x1 b')))
X ## Y
R (set_fold f b (X Y)) (set_fold f (set_fold f b X) Y).
Proof. intros. apply set_fold_union_strong; set_solver. Qed.
Lemma set_fold_disj_union (f : A A A) (b : A) X Y :
Comm (=) f
Assoc (=) f
X ## Y
set_fold f b (X Y) = set_fold f (set_fold f b X) Y.
Proof.
intros Hcomm Hassoc Hdisj. unfold set_fold; simpl.
by rewrite elements_disj_union, <- foldr_app, (comm (++)).
intros. apply (set_fold_disj_union_strong _ _ _ _ _ _); [|done].
intros x1 x2 b' _ _ _. by rewrite !(assoc_L f), (comm_L f x1).
Qed.
Lemma set_fold_comm_acc_strong {B} (R : relation B) `{!PreOrder R}
(f : A B B) (g : B B) (b : B) X :
( x, Proper (R ==> R) (f x))
( x y, x X R (f x (g y)) (g (f x y)))
R (set_fold f (g b) X) (g (set_fold f b X)).
Proof.
intros. unfold set_fold; simpl.
apply foldr_comm_acc_strong; [done|solve_proper|set_solver].
Qed.
Lemma set_fold_comm_acc {B} (f : A B B) (g : B B) (b : B) X :
( x y, f x (g y) = g (f x y))
set_fold f (g b) X = g (set_fold f b X).
Proof. intros. apply (set_fold_comm_acc_strong _); [solve_proper|auto]. Qed.
(** * Minimal elements *)
Lemma minimal_exists R `{!Transitive R, x y, Decision (R x y)} (X : C) :
Lemma minimal_exists_elem_of R `{!Transitive R, x y, Decision (R x y)} (X : C) :
X x, x X minimal R x X.
Proof.
pattern X; apply set_ind; clear X.
......@@ -269,51 +445,78 @@ Proof.
exists x; split; [set_solver|].
rewrite HX, (right_id _ ()). apply singleton_minimal.
Qed.
Lemma minimal_exists_L R `{!LeibnizEquiv C, !Transitive R,
Lemma minimal_exists_elem_of_L R `{!LeibnizEquiv C, !Transitive R,
x y, Decision (R x y)} (X : C) :
X x, x X minimal R x X.
Proof. unfold_leibniz. apply (minimal_exists R). Qed.
Proof. unfold_leibniz. apply (minimal_exists_elem_of R). Qed.
Lemma minimal_exists R `{!Transitive R,
x y, Decision (R x y)} `{!Inhabited A} (X : C) :
x, minimal R x X.
Proof.
destruct (set_choose_or_empty X) as [ (y & Ha) | Hne].
- edestruct (minimal_exists_elem_of R X) as (x & Hel & Hmin); first set_solver.
exists x. done.
- exists inhabitant. intros y Hel. set_solver.
Qed.
(** * Filter *)
Lemma elem_of_filter (P : A Prop) `{!∀ x, Decision (P x)} X x :
x filter P X P x x X.
Proof.
unfold filter, set_filter.
by rewrite elem_of_list_to_set, elem_of_list_filter, elem_of_elements.
Qed.
Global Instance set_unfold_filter (P : A Prop) `{!∀ x, Decision (P x)} X Q x :
SetUnfoldElemOf x X Q SetUnfoldElemOf x (filter P X) (P x Q).
Proof.
intros ?; constructor. by rewrite elem_of_filter, (set_unfold_elem_of x X Q).
Qed.
Section filter.
Context (P : A Prop) `{!∀ x, Decision (P x)}.
Lemma elem_of_filter X x : x filter P X P x x X.
Proof.
unfold filter, set_filter.
by rewrite elem_of_list_to_set, elem_of_list_filter, elem_of_elements.
Qed.
Global Instance set_unfold_filter X Q :
SetUnfoldElemOf x X Q SetUnfoldElemOf x (filter P X) (P x Q).
Proof.
intros x ?; constructor. by rewrite elem_of_filter, (set_unfold_elem_of x X Q).
Qed.
Lemma filter_empty : filter P (∅:C) ∅.
Proof. set_solver. Qed.
Lemma filter_union X Y : filter P (X Y) filter P X filter P Y.
Proof. set_solver. Qed.
Lemma filter_singleton x : P x filter P ({[ x ]} : C) {[ x ]}.
Proof. set_solver. Qed.
Lemma filter_singleton_not x : ¬P x filter P ({[ x ]} : C) ∅.
Proof. set_solver. Qed.
Lemma filter_empty_not_elem_of X x : filter P X P x x X.
Proof. set_solver. Qed.
Lemma disjoint_filter X Y : X ## Y filter P X ## filter P Y.
Proof. set_solver. Qed.
Lemma filter_union X Y : filter P (X Y) filter P X filter P Y.
Proof. set_solver. Qed.
Lemma disjoint_filter_complement X : filter P X ## filter (λ x, ¬P x) X.
Proof. set_solver. Qed.
Lemma filter_union_complement X : filter P X filter (λ x, ¬P x) X X.
Proof. intros x. destruct (decide (P x)); set_solver. Qed.
Section leibniz_equiv.
Context `{!LeibnizEquiv C}.
Lemma filter_empty_L : filter P (∅:C) = ∅.
Proof. set_solver. Qed.
Lemma filter_union_L X Y : filter P (X Y) = filter P X filter P Y.
Proof. set_solver. Qed.
Proof. unfold_leibniz. apply filter_empty. Qed.
Lemma filter_singleton_L x : P x filter P ({[ x ]} : C) = {[ x ]}.
Proof. set_solver. Qed.
Proof. unfold_leibniz. apply filter_singleton. Qed.
Lemma filter_singleton_not_L x : ¬P x filter P ({[ x ]} : C) = ∅.
Proof. set_solver. Qed.
Proof. unfold_leibniz. apply filter_singleton_not. Qed.
Lemma filter_empty_not_elem_of_L X x : filter P X = P x x X.
Proof. unfold_leibniz. apply filter_empty_not_elem_of. Qed.
Lemma filter_union_L X Y : filter P (X Y) = filter P X filter P Y.
Proof. unfold_leibniz. apply filter_union. Qed.
Lemma filter_union_complement_L X Y : filter P X filter (λ x, ¬P x) X = X.
Proof. unfold_leibniz. apply filter_union_complement. Qed.
End leibniz_equiv.
End filter.
(** * Map *)
Section map.
Context `{Set_ B D}.
Context `{SemiSet B D}.
Lemma elem_of_map (f : A B) (X : C) y :
y set_map (D:=D) f X x, y = f x x X.
......@@ -321,9 +524,9 @@ Section map.
unfold set_map. rewrite elem_of_list_to_set, elem_of_list_fmap.
by setoid_rewrite elem_of_elements.
Qed.
Global Instance set_unfold_map (f : A B) (X : C) (P : A Prop) :
( y, SetUnfoldElemOf y X (P y))
SetUnfoldElemOf x (set_map (D:=D) f X) ( y, x = f y P y).
Global Instance set_unfold_map (f : A B) (X : C) (P : A Prop) y :
( x, SetUnfoldElemOf x X (P x))
SetUnfoldElemOf y (set_map (D:=D) f X) ( x, y = f x P x).
Proof. constructor. rewrite elem_of_map; naive_solver. Qed.
Global Instance set_map_proper :
......@@ -342,8 +545,160 @@ Section map.
Lemma elem_of_map_2_alt (f : A B) (X : C) (x : A) (y : B) :
x X y = f x y set_map (D:=D) f X.
Proof. set_solver. Qed.
Lemma set_map_empty (f : A B) :
set_map (C:=C) (D:=D) f = ∅.
Proof. unfold set_map. rewrite elements_empty. done. Qed.
Lemma set_map_union (f : A B) (X Y : C) :
set_map (D:=D) f (X Y) set_map (D:=D) f X set_map (D:=D) f Y.
Proof. set_solver. Qed.
(** This cannot be using [=] because [list_to_set_singleton] does not hold for [=]. *)
Lemma set_map_singleton (f : A B) (x : A) :
set_map (C:=C) (D:=D) f {[x]} {[f x]}.
Proof. set_solver. Qed.
Lemma set_map_union_L `{!LeibnizEquiv D} (f : A B) (X Y : C) :
set_map (D:=D) f (X Y) = set_map (D:=D) f X set_map (D:=D) f Y.
Proof. unfold_leibniz. apply set_map_union. Qed.
Lemma set_map_singleton_L `{!LeibnizEquiv D} (f : A B) (x : A) :
set_map (C:=C) (D:=D) f {[x]} = {[f x]}.
Proof. unfold_leibniz. apply set_map_singleton. Qed.
End map.
(** * Bind *)
Section set_bind.
Context `{SemiSet B SB}.
Local Notation set_bind := (set_bind (A:=A) (SA:=C) (SB:=SB)).
Lemma elem_of_set_bind (f : A SB) (X : C) y :
y set_bind f X x, x X y f x.
Proof.
unfold set_bind. rewrite !elem_of_union_list. set_solver.
Qed.
Global Instance set_unfold_set_bind (f : A SB) (X : C)
(y : B) (P : A B Prop) (Q : A Prop) :
( x y, SetUnfoldElemOf y (f x) (P x y))
( x, SetUnfoldElemOf x X (Q x))
SetUnfoldElemOf y (set_bind f X) ( x, Q x P x y).
Proof.
intros HSU1 HSU2. constructor.
rewrite elem_of_set_bind. set_solver.
Qed.
Global Instance set_bind_proper :
Proper (pointwise_relation _ () ==> () ==> ()) set_bind.
Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
Global Instance set_bind_mono :
Proper (pointwise_relation _ () ==> () ==> ()) set_bind.
Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
Lemma set_bind_ext (f g : A SB) (X Y : C) :
( x, x X x Y f x g x) X Y set_bind f X set_bind g Y.
Proof. set_solver. Qed.
Lemma set_bind_singleton f x : set_bind f {[x]} f x.
Proof. set_solver. Qed.
Lemma set_bind_singleton_L `{!LeibnizEquiv SB} f x : set_bind f {[x]} = f x.
Proof. unfold_leibniz. apply set_bind_singleton. Qed.
Lemma set_bind_disj_union f (X Y : C) :
X ## Y set_bind f (X Y) set_bind f X set_bind f Y.
Proof. set_solver. Qed.
Lemma set_bind_disj_union_L `{!LeibnizEquiv SB} f (X Y : C) :
X ## Y set_bind f (X Y) = set_bind f X set_bind f Y.
Proof. unfold_leibniz. apply set_bind_disj_union. Qed.
End set_bind.
(** * OMap *)
Section set_omap.
Context `{SemiSet B D}.
Implicit Types (f : A option B).
Implicit Types (x : A) (y : B).
Notation set_omap := (set_omap (C:=C) (D:=D)).
Lemma elem_of_set_omap f X y : y set_omap f X x, x X f x = Some y.
Proof.
unfold set_omap. rewrite elem_of_list_to_set, elem_of_list_omap.
by setoid_rewrite elem_of_elements.
Qed.
Global Instance set_unfold_omap f X (P : A Prop) y :
( x, SetUnfoldElemOf x X (P x))
SetUnfoldElemOf y (set_omap f X) ( x, Some y = f x P x).
Proof. constructor. rewrite elem_of_set_omap; naive_solver. Qed.
Global Instance set_omap_proper :
Proper (pointwise_relation _ (=) ==> () ==> ()) set_omap.
Proof. intros f g Hfg X Y. set_unfold. setoid_rewrite Hfg. naive_solver. Qed.
Global Instance set_omap_mono :
Proper (pointwise_relation _ (=) ==> () ==> ()) set_omap.
Proof. intros f g Hfg X Y. set_unfold. setoid_rewrite Hfg. naive_solver. Qed.
Lemma elem_of_set_omap_1 f X y : y set_omap f X x, Some y = f x x X.
Proof. set_solver. Qed.
Lemma elem_of_set_omap_2 f X x y : x X f x = Some y y set_omap f X.
Proof. set_solver. Qed.
Lemma set_omap_empty f : set_omap f = ∅.
Proof. unfold set_omap. by rewrite elements_empty. Qed.
Lemma set_omap_empty_iff f X : set_omap f X set_Forall (λ x, f x = None) X.
Proof.
split; set_unfold; unfold set_Forall.
- intros Hi x Hx. destruct (f x) as [y|] eqn:Hy; naive_solver.
- intros Hi y (x & Hf & Hx). specialize (Hi x Hx). by rewrite Hi in Hf.
Qed.
Lemma set_omap_union f X Y : set_omap f (X Y) set_omap f X set_omap f Y.
Proof. set_solver. Qed.
Lemma set_omap_singleton f x :
set_omap f {[ x ]} match f x with Some y => {[ y ]} | None => end.
Proof. set_solver. Qed.
Lemma set_omap_singleton_Some f x y : f x = Some y set_omap f {[ x ]} {[ y ]}.
Proof. intros Hx. by rewrite set_omap_singleton, Hx. Qed.
Lemma set_omap_singleton_None f x : f x = None set_omap f {[ x ]} ∅.
Proof. intros Hx. by rewrite set_omap_singleton, Hx. Qed.
Lemma set_omap_alt f X : set_omap f X set_bind (λ x, option_to_set (f x)) X.
Proof. set_solver. Qed.
Lemma set_map_alt (f : A B) X : set_map f X = set_omap (λ x, Some (f x)) X.
Proof. set_solver. Qed.
Lemma set_omap_filter P `{ x, Decision (P x)} f X :
( x, x X is_Some (f x) P x)
set_omap f (filter P X) set_omap f X.
Proof. set_solver. Qed.
Section leibniz.
Context `{!LeibnizEquiv D}.
Lemma set_omap_union_L f X Y : set_omap f (X Y) = set_omap f X set_omap f Y.
Proof. unfold_leibniz. apply set_omap_union. Qed.
Lemma set_omap_singleton_L f x :
set_omap f {[ x ]} = match f x with Some y => {[ y ]} | None => end.
Proof. unfold_leibniz. apply set_omap_singleton. Qed.
Lemma set_omap_singleton_Some_L f x y :
f x = Some y set_omap f {[ x ]} = {[ y ]}.
Proof. unfold_leibniz. apply set_omap_singleton_Some. Qed.
Lemma set_omap_singleton_None_L f x : f x = None set_omap f {[ x ]} = ∅.
Proof. unfold_leibniz. apply set_omap_singleton_None. Qed.
Lemma set_omap_alt_L f X :
set_omap f X = set_bind (λ x, option_to_set (f x)) X.
Proof. unfold_leibniz. apply set_omap_alt. Qed.
Lemma set_omap_filter_L P `{ x, Decision (P x)} f X :
( x, x X is_Some (f x) P x)
set_omap f (filter P X) = set_omap f X.
Proof. unfold_leibniz. apply set_omap_filter. Qed.
End leibniz.
End set_omap.
(** * Decision procedures *)
Lemma set_Forall_elements P X : set_Forall P X Forall P (elements X).
Proof. rewrite Forall_forall. by setoid_rewrite elem_of_elements. Qed.
......@@ -389,6 +744,14 @@ Proof.
- intros [X Hfin]. exists (elements X). set_solver.
Qed.
Lemma dec_pred_finite_set_alt (P : A Prop) `{!∀ x : A, Decision (P x)} :
pred_finite P ( X : C, x, P x x X).
Proof.
rewrite dec_pred_finite_alt; [|done]. split.
- intros [xs Hfin]. exists (list_to_set xs). set_solver.
- intros [X Hfin]. exists (elements X). set_solver.
Qed.
Lemma pred_infinite_set (P : A Prop) :
pred_infinite P ( X : C, x, P x x X).
Proof.
......@@ -432,7 +795,7 @@ Section infinite.
Forall_fresh X xs Y X Forall_fresh Y xs.
Proof. rewrite !Forall_fresh_alt; set_solver. Qed.
Lemma fresh_list_length n X : length (fresh_list n X) = n.
Lemma length_fresh_list n X : length (fresh_list n X) = n.
Proof. revert X. induction n; simpl; auto. Qed.
Lemma fresh_list_is_fresh n X x : x fresh_list n X x X.
Proof.
......@@ -451,3 +814,11 @@ Section infinite.
Qed.
End infinite.
End fin_set.
Lemma size_set_seq `{FinSet nat C} start len :
size (set_seq (C:=C) start len) = len.
Proof.
rewrite <-list_to_set_seq, size_list_to_set.
2:{ apply NoDup_seq. }
rewrite length_seq. done.
Qed.
......@@ -7,11 +7,11 @@ Class Finite A `{EqDecision A} := {
NoDup_enum : NoDup enum;
elem_of_enum x : x enum
}.
Hint Mode Finite ! - : typeclass_instances.
Arguments enum : clear implicits.
Arguments enum _ {_ _} : assert.
Arguments NoDup_enum : clear implicits.
Arguments NoDup_enum _ {_ _} : assert.
Global Hint Mode Finite ! - : typeclass_instances.
Global Arguments enum : clear implicits.
Global Arguments enum _ {_ _} : assert.
Global Arguments NoDup_enum : clear implicits.
Global Arguments NoDup_enum _ {_ _} : assert.
Definition card A `{Finite A} := length (enum A).
Program Definition finite_countable `{Finite A} : Countable A := {|
......@@ -19,16 +19,15 @@ Program Definition finite_countable `{Finite A} : Countable A := {|
Pos.of_nat $ S $ default 0 $ fst <$> list_find (x =.) (enum A);
decode := λ p, enum A !! pred (Pos.to_nat p)
|}.
Arguments Pos.of_nat : simpl never.
Next Obligation.
intros ?? [xs Hxs HA] x; unfold encode, decode; simpl.
destruct (list_find_elem_of (x =.) xs x) as [[i y] Hi]; auto.
rewrite Nat2Pos.id by done; simpl; rewrite Hi; simpl.
destruct (list_find_Some (x =.) xs i y); naive_solver.
Qed.
Hint Immediate finite_countable : typeclass_instances.
Global Hint Immediate finite_countable : typeclass_instances.
Definition find `{Finite A} P `{ x, Decision (P x)} : option A :=
Definition find `{Finite A} (P : A Prop) `{ x, Decision (P x)} : option A :=
list_find P (enum A) ≫= decode_nat fst.
Lemma encode_lt_card `{finA: Finite A} (x : A) : encode_nat x < card A.
......@@ -51,7 +50,7 @@ Proof.
split; [done|]; rewrite Hj; simpl.
apply list_find_Some in Hj as (?&->&?). eauto using NoDup_lookup.
Qed.
Lemma find_Some `{finA: Finite A} P `{ x, Decision (P x)} (x : A) :
Lemma find_Some `{finA: Finite A} (P : A Prop) `{ x, Decision (P x)} (x : A) :
find P = Some x P x.
Proof.
destruct finA as [xs Hxs HA]; unfold find, decode_nat, decode; simpl.
......@@ -59,13 +58,13 @@ Proof.
rewrite !Nat2Pos.id in Hx by done.
destruct (list_find_Some P xs i y); naive_solver.
Qed.
Lemma find_is_Some `{finA: Finite A} P `{ x, Decision (P x)} (x : A) :
Lemma find_is_Some `{finA: Finite A} (P : A Prop) `{ x, Decision (P x)} (x : A) :
P x y, find P = Some y P y.
Proof.
destruct finA as [xs Hxs HA]; unfold find, decode; simpl.
intros Hx. destruct (list_find_elem_of P xs x) as [[i y] Hi]; auto.
rewrite Hi; simpl. rewrite !Nat2Pos.id by done. simpl.
apply list_find_Some in Hi; naive_solver.
rewrite Hi; unfold decode_nat, decode; simpl. rewrite !Nat2Pos.id by done.
simpl. apply list_find_Some in Hi; naive_solver.
Qed.
Definition encode_fin `{Finite A} (x : A) : fin (card A) :=
......@@ -119,9 +118,9 @@ Qed.
Lemma finite_inj_Permutation `{Finite A} `{Finite B} (f : A B)
`{!Inj (=) (=) f} : card A = card B f <$> enum A enum B.
Proof.
intros. apply submseteq_Permutation_length_eq.
- by rewrite fmap_length.
intros. apply submseteq_length_Permutation.
- by apply finite_inj_submseteq.
- rewrite length_fmap. by apply Nat.eq_le_incl.
Qed.
Lemma finite_inj_surj `{Finite A} `{Finite B} (f : A B)
`{!Inj (=) (=) f} : card A = card B Surj (=) f.
......@@ -147,7 +146,7 @@ Proof.
{ exists (card_0_inv B HA). intros y. apply (card_0_inv _ HA y). }
destruct (finite_surj A B) as (g&?); auto with lia.
destruct (surj_cancel g) as (f&?). exists f. apply cancel_inj.
- intros [f ?]. unfold card. rewrite <-(fmap_length f).
- intros [f ?]. unfold card. rewrite <-(length_fmap f).
by apply submseteq_length, (finite_inj_submseteq f).
Qed.
Lemma finite_bijective A `{Finite A} B `{Finite B} :
......@@ -204,7 +203,7 @@ Section enc_finite.
Context (to_nat_c : x, to_nat x < c).
Context (to_of_nat : i, i < c to_nat (of_nat i) = i).
Program Instance enc_finite : Finite A := {| enum := of_nat <$> seq 0 c |}.
Local Program Instance enc_finite : Finite A := {| enum := of_nat <$> seq 0 c |}.
Next Obligation.
apply NoDup_alt. intros i j x. rewrite !list_lookup_fmap. intros Hi Hj.
destruct (seq _ _ !! i) as [i'|] eqn:Hi',
......@@ -217,21 +216,41 @@ Section enc_finite.
split; auto. by apply elem_of_list_lookup_2 with (to_nat x), lookup_seq.
Qed.
Lemma enc_finite_card : card A = c.
Proof. unfold card. simpl. by rewrite fmap_length, seq_length. Qed.
Proof. unfold card. simpl. by rewrite length_fmap, length_seq. Qed.
End enc_finite.
(** If we have a surjection [f : A → B] and [A] is finite, then [B] is finite
too. The surjection [f] could map multiple [x : A] on the same [B], so we
need to remove duplicates in [enum]. If [f] is injective, we do not need to do that,
leading to a potentially faster implementation of [enum], see [bijective_finite]
below. *)
Section surjective_finite.
Context `{Finite A, EqDecision B} (f : A B).
Context `{!Surj (=) f}.
Program Definition surjective_finite: Finite B :=
{| enum := remove_dups (f <$> enum A) |}.
Next Obligation. apply NoDup_remove_dups. Qed.
Next Obligation.
intros y. rewrite elem_of_remove_dups, elem_of_list_fmap.
destruct (surj f y). eauto using elem_of_enum.
Qed.
End surjective_finite.
Section bijective_finite.
Context `{Finite A, EqDecision B} (f : A B) (g : B A).
Context `{!Inj (=) (=) f, !Cancel (=) f g}.
Context `{Finite A, EqDecision B} (f : A B).
Context `{!Inj (=) (=) f, !Surj (=) f}.
Program Instance bijective_finite: Finite B := {| enum := f <$> enum A |}.
Next Obligation. apply (NoDup_fmap_2 _), NoDup_enum. Qed.
Program Definition bijective_finite : Finite B :=
{| enum := f <$> enum A |}.
Next Obligation. apply (NoDup_fmap f), NoDup_enum. Qed.
Next Obligation.
intros y. rewrite elem_of_list_fmap. eauto using elem_of_enum.
intros b. rewrite elem_of_list_fmap. destruct (surj f b).
eauto using elem_of_enum.
Qed.
End bijective_finite.
Program Instance option_finite `{Finite A} : Finite (option A) :=
Global Program Instance option_finite `{Finite A} : Finite (option A) :=
{| enum := None :: (Some <$> enum A) |}.
Next Obligation.
constructor.
......@@ -243,21 +262,21 @@ Next Obligation.
apply elem_of_list_fmap. eauto using elem_of_enum.
Qed.
Lemma option_cardinality `{Finite A} : card (option A) = S (card A).
Proof. unfold card. simpl. by rewrite fmap_length. Qed.
Proof. unfold card. simpl. by rewrite length_fmap. Qed.
Program Instance Empty_set_finite : Finite Empty_set := {| enum := [] |}.
Global Program Instance Empty_set_finite : Finite Empty_set := {| enum := [] |}.
Next Obligation. by apply NoDup_nil. Qed.
Next Obligation. intros []. Qed.
Lemma Empty_set_card : card Empty_set = 0.
Proof. done. Qed.
Program Instance unit_finite : Finite () := {| enum := [tt] |}.
Global Program Instance unit_finite : Finite () := {| enum := [tt] |}.
Next Obligation. apply NoDup_singleton. Qed.
Next Obligation. intros []. by apply elem_of_list_singleton. Qed.
Lemma unit_card : card unit = 1.
Proof. done. Qed.
Program Instance bool_finite : Finite bool := {| enum := [true; false] |}.
Global Program Instance bool_finite : Finite bool := {| enum := [true; false] |}.
Next Obligation.
constructor; [ by rewrite elem_of_list_singleton | apply NoDup_singleton ].
Qed.
......@@ -265,7 +284,7 @@ Next Obligation. intros [|]; [ left | right; left ]. Qed.
Lemma bool_card : card bool = 2.
Proof. done. Qed.
Program Instance sum_finite `{Finite A, Finite B} : Finite (A + B)%type :=
Global Program Instance sum_finite `{Finite A, Finite B} : Finite (A + B)%type :=
{| enum := (inl <$> enum A) ++ (inr <$> enum B) |}.
Next Obligation.
intros. apply NoDup_app; split_and?.
......@@ -278,79 +297,68 @@ Next Obligation.
[left|right]; (eexists; split; [done|apply elem_of_enum]).
Qed.
Lemma sum_card `{Finite A, Finite B} : card (A + B) = card A + card B.
Proof. unfold card. simpl. by rewrite app_length, !fmap_length. Qed.
Proof. unfold card. simpl. by rewrite length_app, !length_fmap. Qed.
Program Instance prod_finite `{Finite A, Finite B} : Finite (A * B)%type :=
{| enum := foldr (λ x, (pair x <$> enum B ++.)) [] (enum A) |}.
Global Program Instance prod_finite `{Finite A, Finite B} : Finite (A * B)%type :=
{| enum := a enum A; (a,.) <$> enum B |}.
Next Obligation.
intros A ?????. induction (NoDup_enum A) as [|x xs Hx Hxs IH]; simpl.
{ constructor. }
apply NoDup_app; split_and?.
- by apply (NoDup_fmap_2 _), NoDup_enum.
- intros [? y]. rewrite elem_of_list_fmap. intros (?&?&?); simplify_eq.
clear IH. induction Hxs as [|x' xs ?? IH]; simpl.
{ rewrite elem_of_nil. tauto. }
rewrite elem_of_app, elem_of_list_fmap.
intros [(?&?&?)|?]; simplify_eq.
+ destruct Hx. by left.
+ destruct IH; [ | by auto ]. by intro; destruct Hx; right.
- done.
intros A ?????. apply NoDup_bind.
- intros a1 a2 [a b] ?? (?&?&_)%elem_of_list_fmap (?&?&_)%elem_of_list_fmap.
naive_solver.
- intros a ?. rewrite (NoDup_fmap _). apply NoDup_enum.
- apply NoDup_enum.
Qed.
Next Obligation.
intros ?????? [x y]. induction (elem_of_enum x); simpl.
- rewrite elem_of_app, !elem_of_list_fmap. eauto using elem_of_enum.
- rewrite elem_of_app; eauto.
intros ?????? [a b]. apply elem_of_list_bind.
exists a. eauto using elem_of_enum, elem_of_list_fmap_1.
Qed.
Lemma prod_card `{Finite A} `{Finite B} : card (A * B) = card A * card B.
Proof.
unfold card; simpl. induction (enum A); simpl; auto.
rewrite app_length, fmap_length. auto.
rewrite length_app, length_fmap. auto.
Qed.
Definition list_enum {A} (l : list A) : n, list { l : list A | length l = n } :=
fix go n :=
Fixpoint vec_enum {A} (l : list A) (n : nat) : list (vec A n) :=
match n with
| 0 => [[]eq_refl]
| S n => foldr (λ x, (sig_map (x ::.) (λ _ H, f_equal S H) <$> (go n) ++.)) [] l
| 0 => [[#]]
| S m => a l; vcons a <$> vec_enum l m
end.
Program Instance list_finite `{Finite A} n : Finite { l : list A | length l = n } :=
{| enum := list_enum (enum A) n |}.
Global Program Instance vec_finite `{Finite A} n : Finite (vec A n) :=
{| enum := vec_enum (enum A) n |}.
Next Obligation.
intros A ?? n. induction n as [|n IH]; simpl; [apply NoDup_singleton |].
revert IH. generalize (list_enum (enum A) n). intros l Hl.
induction (NoDup_enum A) as [|x xs Hx Hxs IH]; simpl; auto; [constructor |].
apply NoDup_app; split_and?.
- by apply (NoDup_fmap_2 _).
- intros [k1 Hk1]. clear Hxs IH. rewrite elem_of_list_fmap.
intros ([k2 Hk2]&?&?) Hxk2; simplify_eq/=. destruct Hx. revert Hxk2.
induction xs as [|x' xs IH]; simpl in *; [by rewrite elem_of_nil |].
rewrite elem_of_app, elem_of_list_fmap, elem_of_cons.
intros [([??]&?&?)|?]; simplify_eq/=; auto.
- apply IH.
intros A ?? n. induction n as [|n IH]; csimpl; [apply NoDup_singleton|].
apply NoDup_bind.
- intros x1 x2 y ?? (?&?&_)%elem_of_list_fmap (?&?&_)%elem_of_list_fmap.
congruence.
- intros x ?. rewrite NoDup_fmap by (intros ?; apply vcons_inj_2). done.
- apply NoDup_enum.
Qed.
Next Obligation.
intros A ?? n [l Hl]. revert l Hl.
induction n as [|n IH]; intros [|x l] Hl; simpl; simplify_eq.
{ apply elem_of_list_singleton. by apply (sig_eq_pi _). }
revert IH. generalize (list_enum (enum A) n). intros k Hk.
induction (elem_of_enum x) as [x xs|x xs]; simpl in *.
- rewrite elem_of_app, elem_of_list_fmap. left. injection Hl. intros Hl'.
eexists (lHl'). split; [|done]. by apply (sig_eq_pi _).
- rewrite elem_of_app. eauto.
intros A ?? n v. induction v as [|x n v IH]; csimpl; [apply elem_of_list_here|].
apply elem_of_list_bind. eauto using elem_of_enum, elem_of_list_fmap_1.
Qed.
Lemma list_card `{Finite A} n : card { l : list A | length l = n } = card A ^ n.
Lemma vec_card `{Finite A} n : card (vec A n) = card A ^ n.
Proof.
unfold card; simpl. induction n as [|n IH]; simpl; auto.
rewrite <-IH. clear IH. generalize (list_enum (enum A) n).
induction (enum A) as [|x xs IH]; intros l; simpl; auto.
by rewrite app_length, fmap_length, IH.
unfold card; simpl. induction n as [|n IH]; simpl; [done|].
rewrite <-IH. clear IH. generalize (vec_enum (enum A) n).
induction (enum A) as [|x xs IH]; intros l; csimpl; auto.
by rewrite length_app, length_fmap, IH.
Qed.
Global Instance list_finite `{Finite A} n : Finite { l : list A | length l = n }.
Proof.
refine (bijective_finite (λ v : vec A n, vec_to_list v length_vec_to_list _)).
- abstract (by intros v1 v2 [= ?%vec_to_list_inj2]).
- abstract (intros [l <-]; exists (list_to_vec l);
apply (sig_eq_pi _), vec_to_list_to_vec).
Defined.
Lemma list_card `{Finite A} n : card { l : list A | length l = n } = card A ^ n.
Proof. unfold card; simpl. rewrite length_fmap. apply vec_card. Qed.
Fixpoint fin_enum (n : nat) : list (fin n) :=
match n with 0 => [] | S n => 0%fin :: (FS <$> fin_enum n) end.
Program Instance fin_finite n : Finite (fin n) := {| enum := fin_enum n |}.
Global Program Instance fin_finite n : Finite (fin n) := {| enum := fin_enum n |}.
Next Obligation.
intros n. induction n; simpl; constructor.
- rewrite elem_of_list_fmap. by intros (?&?&?).
......@@ -361,7 +369,18 @@ Next Obligation.
rewrite elem_of_cons, ?elem_of_list_fmap; eauto.
Qed.
Lemma fin_card n : card (fin n) = n.
Proof. unfold card; simpl. induction n; simpl; rewrite ?fmap_length; auto. Qed.
Proof. unfold card; simpl. induction n; simpl; rewrite ?length_fmap; auto. Qed.
(* shouldn’t be an instance (cycle with [sig_finite]): *)
Lemma finite_sig_dec `{!EqDecision A} (P : A Prop) `{Finite (sig P)} x :
Decision (P x).
Proof.
assert {xs : list A | x, P x x xs} as [xs ?].
{ clear x. exists (proj1_sig <$> enum _). intros x. split; intros Hx.
- apply elem_of_list_fmap_1_alt with (x Hx); [apply elem_of_enum|]; done.
- apply elem_of_list_fmap in Hx as [[x' Hx'] [-> _]]; done. }
destruct (decide (x xs)); [left | right]; naive_solver.
Qed. (* <- could be Defined but this lemma will probably not be used for computing *)
Section sig_finite.
Context {A} (P : A Prop) `{ x, Decision (P x)}.
......@@ -396,5 +415,48 @@ Section sig_finite.
split; [by destruct p | apply elem_of_enum].
Qed.
Lemma sig_card : card (sig P) = length (filter P (enum A)).
Proof. by rewrite <-list_filter_sig_filter, fmap_length. Qed.
Proof. by rewrite <-list_filter_sig_filter, length_fmap. Qed.
End sig_finite.
Lemma finite_pigeonhole `{Finite A} `{Finite B} (f : A B) :
card B < card A x1 x2, x1 x2 f x1 = f x2.
Proof.
intros. apply dec_stable; intros Heq.
cut (Inj eq eq f); [intros ?%inj_card; lia|].
intros x1 x2 ?. apply dec_stable. naive_solver.
Qed.
Lemma nat_pigeonhole (f : nat nat) (n1 n2 : nat) :
n2 < n1
( i, i < n1 f i < n2)
i1 i2, i1 < i2 < n1 f i1 = f i2.
Proof.
intros Hn Hf. pose (f' (i : fin n1) := nat_to_fin (Hf _ (fin_to_nat_lt i))).
destruct (finite_pigeonhole f') as (i1&i2&Hi&Hf'); [by rewrite !fin_card|].
apply (not_inj (f:=fin_to_nat)) in Hi. apply (f_equal fin_to_nat) in Hf'.
unfold f' in Hf'. rewrite !fin_to_nat_to_fin in Hf'.
pose proof (fin_to_nat_lt i1); pose proof (fin_to_nat_lt i2).
destruct (decide (i1 < i2)); [exists i1, i2|exists i2, i1]; lia.
Qed.
Lemma list_pigeonhole {A} (l1 l2 : list A) :
l1 l2
length l2 < length l1
i1 i2 x, i1 < i2 l1 !! i1 = Some x l1 !! i2 = Some x.
Proof.
intros Hl Hlen.
assert ( i : fin (length l1), (j : fin (length l2)) x,
l1 !! (fin_to_nat i) = Some x
l2 !! (fin_to_nat j) = Some x) as [f Hf]%fin_choice.
{ intros i. destruct (lookup_lt_is_Some_2 l1 i)
as [x Hix]; [apply fin_to_nat_lt|].
assert (x l2) as [j Hjx]%elem_of_list_lookup_1
by (by eapply Hl, elem_of_list_lookup_2).
exists (nat_to_fin (lookup_lt_Some _ _ _ Hjx)), x.
by rewrite fin_to_nat_to_fin. }
destruct (finite_pigeonhole f) as (i1&i2&Hi&Hf'); [by rewrite !fin_card|].
destruct (Hf i1) as (x1&?&?), (Hf i2) as (x2&?&?).
assert (x1 = x2) as -> by congruence.
apply (not_inj (f:=fin_to_nat)) in Hi. apply (f_equal fin_to_nat) in Hf'.
destruct (decide (i1 < i2)); [exists i1, i2|exists i2, i1]; eauto with lia.
Qed.
File moved
(** This files implements an efficient implementation of finite maps whose keys
range over Coq's data type of any countable type [K]. The data structure is
similar to [Pmap], which in turn is based on the "canonical" binary tries
representation by Appel and Leroy, https://hal.inria.fr/hal-03372247. It thus
has the same good properties:
- It guarantees logarithmic-time [lookup] and [partial_alter], and linear-time
[merge]. It has a low constant factor for computation in Coq compared to other
versions (see the Appel and Leroy paper for benchmarks).
- It satisfies extensional equality [(∀ i, m1 !! i = m2 !! i) → m1 = m2].
- It can be used in nested recursive definitions, e.g.,
[Inductive test := Test : gmap test → test]. This is possible because we do
_not_ use a Sigma type to ensure canonical representations (a Sigma type would
break Coq's strict positivity check).
Compared to [Pmap], we not only need to make sure the trie representation is
canonical, we also need to make sure that all positions (of type positive) are
valid encodings of [K]. That is, for each position [q] in the trie, we have
[encode <$> decode q = Some q].
Instead of formalizing this condition using a Sigma type (which would break
the strict positivity check in nested recursive definitions), we make
[gmap_dep_ne A P] dependent on a predicate [P : positive → Prop] that describes
the subset of valid positions, and instantiate it with [gmap_key K].
The predicate [P : positive → Prop] is considered irrelevant by extraction, so
after extraction, the resulting data structure is identical to [Pmap]. *)
From stdpp Require Export countable infinite fin_maps fin_map_dom.
From stdpp Require Import mapset pmap.
From stdpp Require Import options.
Local Open Scope positive_scope.
Local Notation "P ~ 0" := (λ p, P p~0) : function_scope.
Local Notation "P ~ 1" := (λ p, P p~1) : function_scope.
Implicit Type P : positive Prop.
(** * The tree data structure *)
Inductive gmap_dep_ne (A : Type) (P : positive Prop) :=
| GNode001 : gmap_dep_ne A P~1 gmap_dep_ne A P
| GNode010 : P 1 A gmap_dep_ne A P
| GNode011 : P 1 A gmap_dep_ne A P~1 gmap_dep_ne A P
| GNode100 : gmap_dep_ne A P~0 gmap_dep_ne A P
| GNode101 : gmap_dep_ne A P~0 gmap_dep_ne A P~1 gmap_dep_ne A P
| GNode110 : gmap_dep_ne A P~0 P 1 A gmap_dep_ne A P
| GNode111 : gmap_dep_ne A P~0 P 1 A gmap_dep_ne A P~1 gmap_dep_ne A P.
Global Arguments GNode001 {A P} _ : assert.
Global Arguments GNode010 {A P} _ _ : assert.
Global Arguments GNode011 {A P} _ _ _ : assert.
Global Arguments GNode100 {A P} _ : assert.
Global Arguments GNode101 {A P} _ _ : assert.
Global Arguments GNode110 {A P} _ _ _ : assert.
Global Arguments GNode111 {A P} _ _ _ _ : assert.
(** Using [Variant] we suppress the generation of the induction scheme. We use
the induction scheme [gmap_ind] in terms of the smart constructors to reduce the
number of cases, similar to Appel and Leroy. *)
Variant gmap_dep (A : Type) (P : positive Prop) :=
| GEmpty : gmap_dep A P
| GNodes : gmap_dep_ne A P gmap_dep A P.
Global Arguments GEmpty {A P}.
Global Arguments GNodes {A P} _.
Record gmap_key K `{Countable K} (q : positive) :=
GMapKey { _ : encode (A:=K) <$> decode q = Some q }.
Add Printing Constructor gmap_key.
Global Arguments GMapKey {_ _ _ _} _.
Lemma gmap_key_encode `{Countable K} (k : K) : gmap_key K (encode k).
Proof. constructor. by rewrite decode_encode. Qed.
Global Instance gmap_key_pi `{Countable K} q : ProofIrrel (gmap_key K q).
Proof. intros [?] [?]. f_equal. apply (proof_irrel _). Qed.
Record gmap K `{Countable K} A := GMap { gmap_car : gmap_dep A (gmap_key K) }.
Add Printing Constructor gmap.
Global Arguments GMap {_ _ _ _} _.
Global Arguments gmap_car {_ _ _ _} _.
Global Instance gmap_dep_ne_eq_dec {A P} :
EqDecision A ( i, ProofIrrel (P i)) EqDecision (gmap_dep_ne A P).
Proof.
intros ? Hirr t1 t2. revert P t1 t2 Hirr.
refine (fix go {P} (t1 t2 : gmap_dep_ne A P) {Hirr : _} : Decision (t1 = t2) :=
match t1, t2 with
| GNode001 r1, GNode001 r2 => cast_if (go r1 r2)
| GNode010 _ x1, GNode010 _ x2 => cast_if (decide (x1 = x2))
| GNode011 _ x1 r1, GNode011 _ x2 r2 =>
cast_if_and (decide (x1 = x2)) (go r1 r2)
| GNode100 l1, GNode100 l2 => cast_if (go l1 l2)
| GNode101 l1 r1, GNode101 l2 r2 => cast_if_and (go l1 l2) (go r1 r2)
| GNode110 l1 _ x1, GNode110 l2 _ x2 =>
cast_if_and (go l1 l2) (decide (x1 = x2))
| GNode111 l1 _ x1 r1, GNode111 l2 _ x2 r2 =>
cast_if_and3 (go l1 l2) (decide (x1 = x2)) (go r1 r2)
| _, _ => right _
end);
clear go; abstract first [congruence|f_equal; done || apply Hirr|idtac].
Defined.
Global Instance gmap_dep_eq_dec {A P} :
( i, ProofIrrel (P i)) EqDecision A EqDecision (gmap_dep A P).
Proof. intros. solve_decision. Defined.
Global Instance gmap_eq_dec `{Countable K} {A} :
EqDecision A EqDecision (gmap K A).
Proof. intros. solve_decision. Defined.
(** The smart constructor [GNode] and eliminator [gmap_dep_ne_case] are used to
reduce the number of cases, similar to Appel and Leroy. *)
Local Definition GNode {A P}
(ml : gmap_dep A P~0)
(mx : option (P 1 * A)) (mr : gmap_dep A P~1) : gmap_dep A P :=
match ml, mx, mr with
| GEmpty, None, GEmpty => GEmpty
| GEmpty, None, GNodes r => GNodes (GNode001 r)
| GEmpty, Some (p,x), GEmpty => GNodes (GNode010 p x)
| GEmpty, Some (p,x), GNodes r => GNodes (GNode011 p x r)
| GNodes l, None, GEmpty => GNodes (GNode100 l)
| GNodes l, None, GNodes r => GNodes (GNode101 l r)
| GNodes l, Some (p,x), GEmpty => GNodes (GNode110 l p x)
| GNodes l, Some (p,x), GNodes r => GNodes (GNode111 l p x r)
end.
Local Definition gmap_dep_ne_case {A P B} (t : gmap_dep_ne A P)
(f : gmap_dep A P~0 option (P 1 * A) gmap_dep A P~1 B) : B :=
match t with
| GNode001 r => f GEmpty None (GNodes r)
| GNode010 p x => f GEmpty (Some (p,x)) GEmpty
| GNode011 p x r => f GEmpty (Some (p,x)) (GNodes r)
| GNode100 l => f (GNodes l) None GEmpty
| GNode101 l r => f (GNodes l) None (GNodes r)
| GNode110 l p x => f (GNodes l) (Some (p,x)) GEmpty
| GNode111 l p x r => f (GNodes l) (Some (p,x)) (GNodes r)
end.
(** Operations *)
Local Definition gmap_dep_ne_lookup {A} : {P}, positive gmap_dep_ne A P option A :=
fix go {P} i t {struct t} :=
match t, i with
| (GNode010 _ x | GNode011 _ x _ | GNode110 _ _ x | GNode111 _ _ x _), 1 => Some x
| (GNode100 l | GNode110 l _ _ | GNode101 l _ | GNode111 l _ _ _), i~0 => go i l
| (GNode001 r | GNode011 _ _ r | GNode101 _ r | GNode111 _ _ _ r), i~1 => go i r
| _, _ => None
end.
Local Definition gmap_dep_lookup {A P}
(i : positive) (mt : gmap_dep A P) : option A :=
match mt with GEmpty => None | GNodes t => gmap_dep_ne_lookup i t end.
Global Instance gmap_lookup `{Countable K} {A} :
Lookup K A (gmap K A) := λ k mt,
gmap_dep_lookup (encode k) (gmap_car mt).
Global Instance gmap_empty `{Countable K} {A} : Empty (gmap K A) := GMap GEmpty.
(** Block reduction, even on concrete [gmap]s.
Marking [gmap_empty] as [simpl never] would not be enough, because of
https://github.com/coq/coq/issues/2972 and
https://github.com/coq/coq/issues/2986.
And marking [gmap] consumers as [simpl never] does not work either, see:
https://gitlab.mpi-sws.org/iris/stdpp/-/merge_requests/171#note_53216 *)
Global Opaque gmap_empty.
Local Fixpoint gmap_dep_ne_singleton {A P} (i : positive) :
P i A gmap_dep_ne A P :=
match i with
| 1 => GNode010
| i~0 => λ p x, GNode100 (gmap_dep_ne_singleton i p x)
| i~1 => λ p x, GNode001 (gmap_dep_ne_singleton i p x)
end.
Local Definition gmap_partial_alter_aux {A P}
(go : i, P i gmap_dep_ne A P gmap_dep A P)
(f : option A option A) (i : positive) (p : P i)
(mt : gmap_dep A P) : gmap_dep A P :=
match mt with
| GEmpty =>
match f None with
| None => GEmpty | Some x => GNodes (gmap_dep_ne_singleton i p x)
end
| GNodes t => go i p t
end.
Local Definition gmap_dep_ne_partial_alter {A} (f : option A option A) :
{P} (i : positive), P i gmap_dep_ne A P gmap_dep A P :=
Eval lazy -[gmap_dep_ne_singleton] in
fix go {P} i p t {struct t} :=
gmap_dep_ne_case t $ λ ml mx mr,
match i with
| 1 => λ p, GNode ml ((p,.) <$> f (snd <$> mx)) mr
| i~0 => λ p, GNode (gmap_partial_alter_aux go f i p ml) mx mr
| i~1 => λ p, GNode ml mx (gmap_partial_alter_aux go f i p mr)
end p.
Local Definition gmap_dep_partial_alter {A P}
(f : option A option A) : i : positive, P i gmap_dep A P gmap_dep A P :=
gmap_partial_alter_aux (gmap_dep_ne_partial_alter f) f.
Global Instance gmap_partial_alter `{Countable K} {A} :
PartialAlter K A (gmap K A) := λ f k '(GMap mt),
GMap $ gmap_dep_partial_alter f (encode k) (gmap_key_encode k) mt.
Local Definition gmap_dep_ne_fmap {A B} (f : A B) :
{P}, gmap_dep_ne A P gmap_dep_ne B P :=
fix go {P} t :=
match t with
| GNode001 r => GNode001 (go r)
| GNode010 p x => GNode010 p (f x)
| GNode011 p x r => GNode011 p (f x) (go r)
| GNode100 l => GNode100 (go l)
| GNode101 l r => GNode101 (go l) (go r)
| GNode110 l p x => GNode110 (go l) p (f x)
| GNode111 l p x r => GNode111 (go l) p (f x) (go r)
end.
Local Definition gmap_dep_fmap {A B P} (f : A B)
(mt : gmap_dep A P) : gmap_dep B P :=
match mt with GEmpty => GEmpty | GNodes t => GNodes (gmap_dep_ne_fmap f t) end.
Global Instance gmap_fmap `{Countable K} : FMap (gmap K) := λ {A B} f '(GMap mt),
GMap $ gmap_dep_fmap f mt.
Local Definition gmap_dep_omap_aux {A B P}
(go : gmap_dep_ne A P gmap_dep B P) (tm : gmap_dep A P) : gmap_dep B P :=
match tm with GEmpty => GEmpty | GNodes t' => go t' end.
Local Definition gmap_dep_ne_omap {A B} (f : A option B) :
{P}, gmap_dep_ne A P gmap_dep B P :=
fix go {P} t :=
gmap_dep_ne_case t $ λ ml mx mr,
GNode (gmap_dep_omap_aux go ml) ('(p,x) mx; (p,.) <$> f x)
(gmap_dep_omap_aux go mr).
Local Definition gmap_dep_omap {A B P} (f : A option B) :
gmap_dep A P gmap_dep B P := gmap_dep_omap_aux (gmap_dep_ne_omap f).
Global Instance gmap_omap `{Countable K} : OMap (gmap K) := λ {A B} f '(GMap mt),
GMap $ gmap_dep_omap f mt.
Local Definition gmap_merge_aux {A B C P}
(go : gmap_dep_ne A P gmap_dep_ne B P gmap_dep C P)
(f : option A option B option C)
(mt1 : gmap_dep A P) (mt2 : gmap_dep B P) : gmap_dep C P :=
match mt1, mt2 with
| GEmpty, GEmpty => GEmpty
| GNodes t1', GEmpty => gmap_dep_ne_omap (λ x, f (Some x) None) t1'
| GEmpty, GNodes t2' => gmap_dep_ne_omap (λ x, f None (Some x)) t2'
| GNodes t1', GNodes t2' => go t1' t2'
end.
Local Definition diag_None' {A B C} {P : Prop}
(f : option A option B option C)
(mx : option (P * A)) (my : option (P * B)) : option (P * C) :=
match mx, my with
| None, None => None
| Some (p,x), None => (p,.) <$> f (Some x) None
| None, Some (p,y) => (p,.) <$> f None (Some y)
| Some (p,x), Some (_,y) => (p,.) <$> f (Some x) (Some y)
end.
Local Definition gmap_dep_ne_merge {A B C} (f : option A option B option C) :
{P}, gmap_dep_ne A P gmap_dep_ne B P gmap_dep C P :=
fix go {P} t1 t2 {struct t1} :=
gmap_dep_ne_case t1 $ λ ml1 mx1 mr1,
gmap_dep_ne_case t2 $ λ ml2 mx2 mr2,
GNode (gmap_merge_aux go f ml1 ml2) (diag_None' f mx1 mx2)
(gmap_merge_aux go f mr1 mr2).
Local Definition gmap_dep_merge {A B C P} (f : option A option B option C) :
gmap_dep A P gmap_dep B P gmap_dep C P :=
gmap_merge_aux (gmap_dep_ne_merge f) f.
Global Instance gmap_merge `{Countable K} : Merge (gmap K) :=
λ {A B C} f '(GMap mt1) '(GMap mt2), GMap $ gmap_dep_merge f mt1 mt2.
Local Definition gmap_fold_aux {A B P}
(go : positive B gmap_dep_ne A P B)
(i : positive) (y : B) (mt : gmap_dep A P) : B :=
match mt with GEmpty => y | GNodes t => go i y t end.
Local Definition gmap_dep_ne_fold {A B} (f : positive A B B) :
{P}, positive B gmap_dep_ne A P B :=
fix go {P} i y t :=
gmap_dep_ne_case t $ λ ml mx mr,
gmap_fold_aux go i~1
(gmap_fold_aux go i~0
match mx with None => y | Some (p,x) => f (Pos.reverse i) x y end ml) mr.
Local Definition gmap_dep_fold {A B P} (f : positive A B B) :
positive B gmap_dep A P B :=
gmap_fold_aux (gmap_dep_ne_fold f).
Global Instance gmap_fold `{Countable K} {A} :
MapFold K A (gmap K A) := λ {B} f y '(GMap mt),
gmap_dep_fold (λ i x, match decode i with Some k => f k x | None => id end) 1 y mt.
(** Proofs *)
Local Definition GNode_valid {A P}
(ml : gmap_dep A P~0) (mx : option (P 1 * A)) (mr : gmap_dep A P~1) :=
match ml, mx, mr with GEmpty, None, GEmpty => False | _, _, _ => True end.
Local Lemma gmap_dep_ind A (Q : P, gmap_dep A P Prop) :
( P, Q P GEmpty)
( P ml mx mr, GNode_valid ml mx mr Q _ ml Q _ mr Q P (GNode ml mx mr))
P mt, Q P mt.
Proof.
intros Hemp Hnode P [|t]; [done|]. induction t.
- by apply (Hnode _ GEmpty None (GNodes _)).
- by apply (Hnode _ GEmpty (Some (_,_)) GEmpty).
- by apply (Hnode _ GEmpty (Some (_,_)) (GNodes _)).
- by apply (Hnode _ (GNodes _) None GEmpty).
- by apply (Hnode _ (GNodes _) None (GNodes _)).
- by apply (Hnode _ (GNodes _) (Some (_,_)) GEmpty).
- by apply (Hnode _ (GNodes _) (Some (_,_)) (GNodes _)).
Qed.
Local Lemma gmap_dep_lookup_GNode {A P} (ml : gmap_dep A P~0) mr mx i :
gmap_dep_lookup i (GNode ml mx mr) =
match i with
| 1 => snd <$> mx | i~0 => gmap_dep_lookup i ml | i~1 => gmap_dep_lookup i mr
end.
Proof. by destruct ml, mx as [[]|], mr, i. Qed.
Local Lemma gmap_dep_ne_lookup_not_None {A P} (t : gmap_dep_ne A P) :
i, P i gmap_dep_ne_lookup i t None.
Proof.
induction t; repeat select ( _, _) (fun H => destruct H);
try first [by eexists 1|by eexists _~0|by eexists _~1].
Qed.
Local Lemma gmap_dep_eq_empty {A P} (mt : gmap_dep A P) :
( i, P i gmap_dep_lookup i mt = None) mt = GEmpty.
Proof.
intros Hlookup. destruct mt as [|t]; [done|].
destruct (gmap_dep_ne_lookup_not_None t); naive_solver.
Qed.
Local Lemma gmap_dep_eq {A P} (mt1 mt2 : gmap_dep A P) :
( i, ProofIrrel (P i))
( i, P i gmap_dep_lookup i mt1 = gmap_dep_lookup i mt2) mt1 = mt2.
Proof.
revert mt2. induction mt1 as [|P ml1 mx1 mr1 _ IHl IHr] using gmap_dep_ind;
intros mt2 ? Hlookup;
destruct mt2 as [|? ml2 mx2 mr2 _ _ _] using gmap_dep_ind.
- done.
- symmetry. apply gmap_dep_eq_empty. naive_solver.
- apply gmap_dep_eq_empty. naive_solver.
- f_equal.
+ apply (IHl _ _). intros i. generalize (Hlookup (i~0)).
by rewrite !gmap_dep_lookup_GNode.
+ generalize (Hlookup 1). rewrite !gmap_dep_lookup_GNode.
destruct mx1 as [[]|], mx2 as [[]|]; intros; simplify_eq/=;
repeat f_equal; try apply proof_irrel; naive_solver.
+ apply (IHr _ _). intros i. generalize (Hlookup (i~1)).
by rewrite !gmap_dep_lookup_GNode.
Qed.
Local Lemma gmap_dep_ne_lookup_singleton {A P} i (p : P i) (x : A) :
gmap_dep_ne_lookup i (gmap_dep_ne_singleton i p x) = Some x.
Proof. revert P p. induction i; by simpl. Qed.
Local Lemma gmap_dep_ne_lookup_singleton_ne {A P} i j (p : P i) (x : A) :
i j gmap_dep_ne_lookup j (gmap_dep_ne_singleton i p x) = None.
Proof. revert P j p. induction i; intros ? [?|?|]; naive_solver. Qed.
Local Lemma gmap_dep_partial_alter_GNode {A P} (f : option A option A)
i (p : P i) (ml : gmap_dep A P~0) mx mr :
GNode_valid ml mx mr
gmap_dep_partial_alter f i p (GNode ml mx mr) =
match i with
| 1 => λ p, GNode ml ((p,.) <$> f (snd <$> mx)) mr
| i~0 => λ p, GNode (gmap_dep_partial_alter f i p ml) mx mr
| i~1 => λ p, GNode ml mx (gmap_dep_partial_alter f i p mr)
end p.
Proof. by destruct ml, mx as [[]|], mr. Qed.
Local Lemma gmap_dep_lookup_partial_alter {A P} (f : option A option A)
(mt : gmap_dep A P) i (p : P i) :
gmap_dep_lookup i (gmap_dep_partial_alter f i p mt) = f (gmap_dep_lookup i mt).
Proof.
revert i p. induction mt using gmap_dep_ind.
{ intros i p; simpl. destruct (f None); simpl; [|done].
by rewrite gmap_dep_ne_lookup_singleton. }
intros [] ?;
rewrite gmap_dep_partial_alter_GNode, !gmap_dep_lookup_GNode by done;
done || by destruct (f _).
Qed.
Local Lemma gmap_dep_lookup_partial_alter_ne {A P} (f : option A option A)
(mt : gmap_dep A P) i (p : P i) j :
i j
gmap_dep_lookup j (gmap_dep_partial_alter f i p mt) = gmap_dep_lookup j mt.
Proof.
revert i p j; induction mt using gmap_dep_ind.
{ intros i p j ?; simpl. destruct (f None); simpl; [|done].
by rewrite gmap_dep_ne_lookup_singleton_ne. }
intros [] ? [] ?;
rewrite gmap_dep_partial_alter_GNode, !gmap_dep_lookup_GNode by done;
auto with lia.
Qed.
Local Lemma gmap_dep_lookup_fmap {A B P} (f : A B) (mt : gmap_dep A P) i :
gmap_dep_lookup i (gmap_dep_fmap f mt) = f <$> gmap_dep_lookup i mt.
Proof.
destruct mt as [|t]; simpl; [done|].
revert i. induction t; intros []; by simpl.
Qed.
Local Lemma gmap_dep_omap_GNode {A B P} (f : A option B)
(ml : gmap_dep A P~0) mx mr :
GNode_valid ml mx mr
gmap_dep_omap f (GNode ml mx mr) =
GNode (gmap_dep_omap f ml) ('(p,x) mx; (p,.) <$> f x) (gmap_dep_omap f mr).
Proof. by destruct ml, mx as [[]|], mr. Qed.
Local Lemma gmap_dep_lookup_omap {A B P} (f : A option B) (mt : gmap_dep A P) i :
gmap_dep_lookup i (gmap_dep_omap f mt) = gmap_dep_lookup i mt ≫= f.
Proof.
revert i. induction mt using gmap_dep_ind; [done|].
intros [];
rewrite gmap_dep_omap_GNode, !gmap_dep_lookup_GNode by done; [done..|].
destruct select (option _) as [[]|]; simpl; by try destruct (f _).
Qed.
Section gmap_merge.
Context {A B C} (f : option A option B option C).
Local Lemma gmap_dep_merge_GNode_GEmpty {P} (ml : gmap_dep A P~0) mx mr :
GNode_valid ml mx mr
gmap_dep_merge f (GNode ml mx mr) GEmpty =
GNode (gmap_dep_omap (λ x, f (Some x) None) ml) (diag_None' f mx None)
(gmap_dep_omap (λ x, f (Some x) None) mr).
Proof. by destruct ml, mx as [[]|], mr. Qed.
Local Lemma gmap_dep_merge_GEmpty_GNode {P} (ml : gmap_dep B P~0) mx mr :
GNode_valid ml mx mr
gmap_dep_merge f GEmpty (GNode ml mx mr) =
GNode (gmap_dep_omap (λ x, f None (Some x)) ml) (diag_None' f None mx)
(gmap_dep_omap (λ x, f None (Some x)) mr).
Proof. by destruct ml, mx as [[]|], mr. Qed.
Local Lemma gmap_dep_merge_GNode_GNode {P}
(ml1 : gmap_dep A P~0) ml2 mx1 mx2 mr1 mr2 :
GNode_valid ml1 mx1 mr1 GNode_valid ml2 mx2 mr2
gmap_dep_merge f (GNode ml1 mx1 mr1) (GNode ml2 mx2 mr2) =
GNode (gmap_dep_merge f ml1 ml2) (diag_None' f mx1 mx2)
(gmap_dep_merge f mr1 mr2).
Proof. by destruct ml1, mx1 as [[]|], mr1, ml2, mx2 as [[]|], mr2. Qed.
Local Lemma gmap_dep_lookup_merge {P} (mt1 : gmap_dep A P) (mt2 : gmap_dep B P) i :
gmap_dep_lookup i (gmap_dep_merge f mt1 mt2) =
diag_None f (gmap_dep_lookup i mt1) (gmap_dep_lookup i mt2).
Proof.
revert mt2 i; induction mt1 using gmap_dep_ind; intros mt2 i.
{ induction mt2 using gmap_dep_ind; [done|].
rewrite gmap_dep_merge_GEmpty_GNode, gmap_dep_lookup_GNode by done.
destruct i as [i|i|];
rewrite ?gmap_dep_lookup_omap, gmap_dep_lookup_GNode; simpl;
[by destruct (gmap_dep_lookup i _)..|].
destruct select (option _) as [[]|]; simpl; by try destruct (f _). }
destruct mt2 using gmap_dep_ind.
{ rewrite gmap_dep_merge_GNode_GEmpty, gmap_dep_lookup_GNode by done.
destruct i as [i|i|];
rewrite ?gmap_dep_lookup_omap, gmap_dep_lookup_GNode; simpl;
[by destruct (gmap_dep_lookup i _)..|].
destruct select (option _) as [[]|]; simpl; by try destruct (f _). }
rewrite gmap_dep_merge_GNode_GNode by done.
destruct i; rewrite ?gmap_dep_lookup_GNode; [done..|].
repeat destruct select (option _) as [[]|]; simpl; by try destruct (f _).
Qed.
End gmap_merge.
Local Lemma gmap_dep_fold_GNode {A B} (f : positive A B B)
{P} i y (ml : gmap_dep A P~0) mx mr :
gmap_dep_fold f i y (GNode ml mx mr) = gmap_dep_fold f i~1
(gmap_dep_fold f i~0
match mx with None => y | Some (_,x) => f (Pos.reverse i) x y end ml) mr.
Proof. by destruct ml, mx as [[]|], mr. Qed.
Local Lemma gmap_dep_fold_ind {A} {P} (Q : gmap_dep A P Prop) :
Q GEmpty
( i p x mt,
gmap_dep_lookup i mt = None
( j A' B (f : positive A' B B) (g : A A') b x',
gmap_dep_fold f j b
(gmap_dep_partial_alter (λ _, Some x') i p (gmap_dep_fmap g mt))
= f (Pos.reverse_go i j) x' (gmap_dep_fold f j b (gmap_dep_fmap g mt)))
Q mt Q (gmap_dep_partial_alter (λ _, Some x) i p mt))
mt, Q mt.
Proof.
intros Hemp Hinsert mt. revert Q Hemp Hinsert.
induction mt as [|P ml mx mr ? IHl IHr] using gmap_dep_ind;
intros Q Hemp Hinsert; [done|].
apply (IHr (λ mt, Q (GNode ml mx mt))).
{ apply (IHl (λ mt, Q (GNode mt mx GEmpty))).
{ destruct mx as [[p x]|]; [|done].
replace (GNode GEmpty (Some (p,x)) GEmpty) with
(gmap_dep_partial_alter (λ _, Some x) 1 p GEmpty) by done.
by apply Hinsert. }
intros i p x mt r ? Hfold.
replace (GNode (gmap_dep_partial_alter (λ _, Some x) i p mt) mx GEmpty)
with (gmap_dep_partial_alter (λ _, Some x) (i~0) p (GNode mt mx GEmpty))
by (by destruct mt, mx as [[]|]).
apply Hinsert.
- by rewrite gmap_dep_lookup_GNode.
- intros j A' B f g b x'.
replace (gmap_dep_partial_alter (λ _, Some x') (i~0) p
(gmap_dep_fmap g (GNode mt mx GEmpty)))
with (GNode (gmap_dep_partial_alter (λ _, Some x') i p (gmap_dep_fmap g mt))
(prod_map id g <$> mx) GEmpty)
by (by destruct mt, mx as [[]|]).
replace (gmap_dep_fmap g (GNode mt mx GEmpty))
with (GNode (gmap_dep_fmap g mt) (prod_map id g <$> mx) GEmpty)
by (by destruct mt, mx as [[]|]).
rewrite !gmap_dep_fold_GNode; simpl; auto.
- done. }
intros i p x mt r ? Hfold.
replace (GNode ml mx (gmap_dep_partial_alter (λ _, Some x) i p mt))
with (gmap_dep_partial_alter (λ _, Some x) (i~1) p (GNode ml mx mt))
by (by destruct ml, mx as [[]|], mt).
apply Hinsert.
- by rewrite gmap_dep_lookup_GNode.
- intros j A' B f g b x'.
replace (gmap_dep_partial_alter (λ _, Some x') (i~1) p
(gmap_dep_fmap g (GNode ml mx mt)))
with (GNode (gmap_dep_fmap g ml) (prod_map id g <$> mx)
(gmap_dep_partial_alter (λ _, Some x') i p (gmap_dep_fmap g mt)))
by (by destruct ml, mx as [[]|], mt).
replace (gmap_dep_fmap g (GNode ml mx mt))
with (GNode (gmap_dep_fmap g ml) (prod_map id g <$> mx) (gmap_dep_fmap g mt))
by (by destruct ml, mx as [[]|], mt).
rewrite !gmap_dep_fold_GNode; simpl; auto.
- done.
Qed.
(** Instance of the finite map type class *)
Global Instance gmap_finmap `{Countable K} : FinMap K (gmap K).
Proof.
split.
- intros A [mt1] [mt2] Hlookup. f_equal. apply (gmap_dep_eq _ _ _).
intros i [Hk]. destruct (decode i) as [k|]; simplify_eq/=. apply Hlookup.
- done.
- intros A f [mt] i. apply gmap_dep_lookup_partial_alter.
- intros A f [mt] i j ?. apply gmap_dep_lookup_partial_alter_ne. naive_solver.
- intros A b f [mt] i. apply gmap_dep_lookup_fmap.
- intros A B f [mt] i. apply gmap_dep_lookup_omap.
- intros A B C f [mt1] [mt2] i. apply gmap_dep_lookup_merge.
- done.
- intros A P Hemp Hins [mt].
apply (gmap_dep_fold_ind (λ mt, P (GMap mt))); clear mt; [done|].
intros i [Hk] x mt ? Hfold. destruct (fmap_Some_1 _ _ _ Hk) as (k&Hk'&->).
assert (GMapKey Hk = gmap_key_encode k) as Hkk by (apply proof_irrel).
rewrite Hkk in Hfold |- *. clear Hk Hkk.
apply (Hins k x (GMap mt)); [done|]. intros A' B f g b x'.
trans ((match decode (encode k) with Some k => f k x' | None => id end)
(map_fold f b (g <$> GMap mt))); [apply (Hfold 1)|].
by rewrite Hk'.
Qed.
Global Program Instance gmap_countable
`{Countable K, Countable A} : Countable (gmap K A) := {
encode m := encode (map_to_list m : list (K * A));
decode p := list_to_map <$> decode p
}.
Next Obligation.
intros K ?? A ?? m; simpl. rewrite decode_encode; simpl.
by rewrite list_to_map_to_list.
Qed.
(** Conversion to/from [Pmap] *)
Local Definition gmap_dep_ne_to_pmap_ne {A} : {P}, gmap_dep_ne A P Pmap_ne A :=
fix go {P} t :=
match t with
| GNode001 r => PNode001 (go r)
| GNode010 _ x => PNode010 x
| GNode011 _ x r => PNode011 x (go r)
| GNode100 l => PNode100 (go l)
| GNode101 l r => PNode101 (go l) (go r)
| GNode110 l _ x => PNode110 (go l) x
| GNode111 l _ x r => PNode111 (go l) x (go r)
end.
Local Definition gmap_dep_to_pmap {A P} (mt : gmap_dep A P) : Pmap A :=
match mt with
| GEmpty => PEmpty
| GNodes t => PNodes (gmap_dep_ne_to_pmap_ne t)
end.
Definition gmap_to_pmap {A} (m : gmap positive A) : Pmap A :=
let '(GMap mt) := m in gmap_dep_to_pmap mt.
Local Lemma lookup_gmap_dep_ne_to_pmap_ne {A P} (t : gmap_dep_ne A P) i :
gmap_dep_ne_to_pmap_ne t !! i = gmap_dep_ne_lookup i t.
Proof. revert i; induction t; intros []; by simpl. Qed.
Lemma lookup_gmap_to_pmap {A} (m : gmap positive A) i :
gmap_to_pmap m !! i = m !! i.
Proof. destruct m as [[|t]]; [done|]. apply lookup_gmap_dep_ne_to_pmap_ne. Qed.
Local Definition pmap_ne_to_gmap_dep_ne {A} :
{P}, ( i, P i) Pmap_ne A gmap_dep_ne A P :=
fix go {P} (p : i, P i) t :=
match t with
| PNode001 r => GNode001 (go p~1 r)
| PNode010 x => GNode010 (p 1) x
| PNode011 x r => GNode011 (p 1) x (go p~1 r)
| PNode100 l => GNode100 (go p~0 l)
| PNode101 l r => GNode101 (go p~0 l) (go p~1 r)
| PNode110 l x => GNode110 (go p~0 l) (p 1) x
| PNode111 l x r => GNode111 (go p~0 l) (p 1) x (go p~1 r)
end%function.
Local Definition pmap_to_gmap_dep {A P}
(p : i, P i) (mt : Pmap A) : gmap_dep A P :=
match mt with
| PEmpty => GEmpty
| PNodes t => GNodes (pmap_ne_to_gmap_dep_ne p t)
end.
Definition pmap_to_gmap {A} (m : Pmap A) : gmap positive A :=
GMap $ pmap_to_gmap_dep gmap_key_encode m.
Local Lemma lookup_pmap_ne_to_gmap_dep_ne {A P} (p : i, P i) (t : Pmap_ne A) i :
gmap_dep_ne_lookup i (pmap_ne_to_gmap_dep_ne p t) = t !! i.
Proof. revert P i p; induction t; intros ? [] ?; by simpl. Qed.
Lemma lookup_pmap_to_gmap {A} (m : Pmap A) i : pmap_to_gmap m !! i = m !! i.
Proof. destruct m as [|t]; [done|]. apply lookup_pmap_ne_to_gmap_dep_ne. Qed.
(** * Curry and uncurry *)
Definition gmap_uncurry `{Countable K1, Countable K2} {A} :
gmap K1 (gmap K2 A) gmap (K1 * K2) A :=
map_fold (λ i1 m' macc,
map_fold (λ i2 x, <[(i1,i2):=x]>) macc m') ∅.
Definition gmap_curry `{Countable K1, Countable K2} {A} :
gmap (K1 * K2) A gmap K1 (gmap K2 A) :=
map_fold (λ '(i1, i2) x,
partial_alter (Some <[i2:=x]> default ) i1) ∅.
Section curry_uncurry.
Context `{Countable K1, Countable K2} {A : Type}.
Lemma lookup_gmap_uncurry (m : gmap K1 (gmap K2 A)) i j :
gmap_uncurry m !! (i,j) = m !! i ≫= (.!! j).
Proof.
apply (map_fold_weak_ind (λ mr m, mr !! (i,j) = m !! i ≫= (.!! j))).
{ by rewrite !lookup_empty. }
clear m; intros i' m2 m m12 Hi' IH.
apply (map_fold_weak_ind (λ m2r m2, m2r !! (i,j) = <[i':=m2]> m !! i ≫= (.!! j))).
{ rewrite IH. destruct (decide (i' = i)) as [->|].
- rewrite lookup_insert, Hi'; simpl; by rewrite lookup_empty.
- by rewrite lookup_insert_ne by done. }
intros j' y m2' m12' Hj' IH'. destruct (decide (i = i')) as [->|].
- rewrite lookup_insert; simpl. destruct (decide (j = j')) as [->|].
+ by rewrite !lookup_insert.
+ by rewrite !lookup_insert_ne, IH', lookup_insert by congruence.
- by rewrite !lookup_insert_ne, IH', lookup_insert_ne by congruence.
Qed.
Lemma lookup_gmap_curry (m : gmap (K1 * K2) A) i j :
gmap_curry m !! i ≫= (.!! j) = m !! (i, j).
Proof.
apply (map_fold_weak_ind (λ mr m, mr !! i ≫= (.!! j) = m !! (i, j))).
{ by rewrite !lookup_empty. }
clear m; intros [i' j'] x m12 mr Hij' IH.
destruct (decide (i = i')) as [->|].
- rewrite lookup_partial_alter. destruct (decide (j = j')) as [->|].
+ destruct (mr !! i'); simpl; by rewrite !lookup_insert.
+ destruct (mr !! i'); simpl; by rewrite !lookup_insert_ne by congruence.
- by rewrite lookup_partial_alter_ne, lookup_insert_ne by congruence.
Qed.
Lemma lookup_gmap_curry_None (m : gmap (K1 * K2) A) i :
gmap_curry m !! i = None ( j, m !! (i, j) = None).
Proof.
apply (map_fold_weak_ind
(λ mr m, mr !! i = None ( j, m !! (i, j) = None))); [done|].
clear m; intros [i' j'] x m12 mr Hij' IH.
destruct (decide (i = i')) as [->|].
- split; [by rewrite lookup_partial_alter|].
intros Hi. specialize (Hi j'). by rewrite lookup_insert in Hi.
- rewrite lookup_partial_alter_ne, IH; [|done]. apply forall_proper.
intros j. rewrite lookup_insert_ne; [done|congruence].
Qed.
Lemma gmap_uncurry_curry (m : gmap (K1 * K2) A) :
gmap_uncurry (gmap_curry m) = m.
Proof.
apply map_eq; intros [i j]. by rewrite lookup_gmap_uncurry, lookup_gmap_curry.
Qed.
Lemma gmap_curry_non_empty (m : gmap (K1 * K2) A) i x :
gmap_curry m !! i = Some x x ∅.
Proof.
intros Hm ->. eapply eq_None_not_Some; [|by eexists].
eapply lookup_gmap_curry_None; intros j.
by rewrite <-lookup_gmap_curry, Hm.
Qed.
Lemma gmap_curry_uncurry_non_empty (m : gmap K1 (gmap K2 A)) :
( i x, m !! i = Some x x )
gmap_curry (gmap_uncurry m) = m.
Proof.
intros Hne. apply map_eq; intros i. destruct (m !! i) as [m2|] eqn:Hm.
- destruct (gmap_curry (gmap_uncurry m) !! i) as [m2'|] eqn:Hcurry.
+ f_equal. apply map_eq. intros j.
trans (gmap_curry (gmap_uncurry m) !! i ≫= (.!! j)).
{ by rewrite Hcurry. }
by rewrite lookup_gmap_curry, lookup_gmap_uncurry, Hm.
+ rewrite lookup_gmap_curry_None in Hcurry.
exfalso; apply (Hne i m2), map_eq; [done|intros j].
by rewrite lookup_empty, <-(Hcurry j), lookup_gmap_uncurry, Hm.
- apply lookup_gmap_curry_None; intros j. by rewrite lookup_gmap_uncurry, Hm.
Qed.
End curry_uncurry.
(** * Finite sets *)
Definition gset K `{Countable K} := mapset (gmap K).
Section gset.
Context `{Countable K}.
(* Lift instances of operational TCs from [mapset] and mark them [simpl never]. *)
Global Instance gset_elem_of: ElemOf K (gset K) := _.
Global Instance gset_empty : Empty (gset K) := _.
Global Instance gset_singleton : Singleton K (gset K) := _.
Global Instance gset_union: Union (gset K) := _.
Global Instance gset_intersection: Intersection (gset K) := _.
Global Instance gset_difference: Difference (gset K) := _.
Global Instance gset_elements: Elements K (gset K) := _.
Global Instance gset_eq_dec : EqDecision (gset K) := _.
Global Instance gset_countable : Countable (gset K) := _.
Global Instance gset_equiv_dec : RelDecision (≡@{gset K}) | 1 := _.
Global Instance gset_elem_of_dec : RelDecision (∈@{gset K}) | 1 := _.
Global Instance gset_disjoint_dec : RelDecision (##@{gset K}) := _.
Global Instance gset_subseteq_dec : RelDecision (⊆@{gset K}) := _.
(** We put in an eta expansion to avoid [injection] from unfolding equalities
like [dom (gset _) m1 = dom (gset _) m2]. *)
Global Instance gset_dom {A} : Dom (gmap K A) (gset K) := λ m,
let '(GMap mt) := m in mapset_dom (GMap mt).
Global Arguments gset_elem_of : simpl never.
Global Arguments gset_empty : simpl never.
Global Arguments gset_singleton : simpl never.
Global Arguments gset_union : simpl never.
Global Arguments gset_intersection : simpl never.
Global Arguments gset_difference : simpl never.
Global Arguments gset_elements : simpl never.
Global Arguments gset_eq_dec : simpl never.
Global Arguments gset_countable : simpl never.
Global Arguments gset_equiv_dec : simpl never.
Global Arguments gset_elem_of_dec : simpl never.
Global Arguments gset_disjoint_dec : simpl never.
Global Arguments gset_subseteq_dec : simpl never.
Global Arguments gset_dom : simpl never.
(* Lift instances of other TCs. *)
Global Instance gset_leibniz : LeibnizEquiv (gset K) := _.
Global Instance gset_semi_set : SemiSet K (gset K) | 1 := _.
Global Instance gset_set : Set_ K (gset K) | 1 := _.
Global Instance gset_fin_set : FinSet K (gset K) := _.
Global Instance gset_dom_spec : FinMapDom K (gmap K) (gset K).
Proof.
pose proof (mapset_dom_spec (M:=gmap K)) as [?? Hdom]; split; auto.
intros A m. specialize (Hdom A m). by destruct m.
Qed.
(** If you are looking for a lemma showing that [gset] is extensional, see
[sets.set_eq]. *)
(** The function [gset_to_gmap x X] converts a set [X] to a map with domain
[X] where each key has value [x]. Compared to the generic conversion
[set_to_map], the function [gset_to_gmap] has [O(n)] instead of [O(n log n)]
complexity and has an easier and better developed theory. *)
Definition gset_to_gmap {A} (x : A) (X : gset K) : gmap K A :=
(λ _, x) <$> mapset_car X.
Lemma lookup_gset_to_gmap {A} (x : A) (X : gset K) i :
gset_to_gmap x X !! i = (guard (i X);; Some x).
Proof.
destruct X as [X].
unfold gset_to_gmap, gset_elem_of, elem_of, mapset_elem_of; simpl.
rewrite lookup_fmap.
case_guard; destruct (X !! i) as [[]|]; naive_solver.
Qed.
Lemma lookup_gset_to_gmap_Some {A} (x : A) (X : gset K) i y :
gset_to_gmap x X !! i = Some y i X x = y.
Proof. rewrite lookup_gset_to_gmap. simplify_option_eq; naive_solver. Qed.
Lemma lookup_gset_to_gmap_None {A} (x : A) (X : gset K) i :
gset_to_gmap x X !! i = None i X.
Proof. rewrite lookup_gset_to_gmap. simplify_option_eq; naive_solver. Qed.
Lemma gset_to_gmap_empty {A} (x : A) : gset_to_gmap x = ∅.
Proof. apply fmap_empty. Qed.
Lemma gset_to_gmap_union_singleton {A} (x : A) i Y :
gset_to_gmap x ({[ i ]} Y) = <[i:=x]>(gset_to_gmap x Y).
Proof.
apply map_eq; intros j; apply option_eq; intros y.
rewrite lookup_insert_Some, !lookup_gset_to_gmap_Some, elem_of_union,
elem_of_singleton; destruct (decide (i = j)); intuition.
Qed.
Lemma gset_to_gmap_singleton {A} (x : A) i :
gset_to_gmap x {[ i ]} = {[i:=x]}.
Proof.
rewrite <-(right_id_L () {[ i ]}), gset_to_gmap_union_singleton.
by rewrite gset_to_gmap_empty.
Qed.
Lemma gset_to_gmap_difference_singleton {A} (x : A) i Y :
gset_to_gmap x (Y {[i]}) = delete i (gset_to_gmap x Y).
Proof.
apply map_eq; intros j; apply option_eq; intros y.
rewrite lookup_delete_Some, !lookup_gset_to_gmap_Some, elem_of_difference,
elem_of_singleton; destruct (decide (i = j)); intuition.
Qed.
Lemma fmap_gset_to_gmap {A B} (f : A B) (X : gset K) (x : A) :
f <$> gset_to_gmap x X = gset_to_gmap (f x) X.
Proof.
apply map_eq; intros j. rewrite lookup_fmap, !lookup_gset_to_gmap.
by simplify_option_eq.
Qed.
Lemma gset_to_gmap_dom {A B} (m : gmap K A) (y : B) :
gset_to_gmap y (dom m) = const y <$> m.
Proof.
apply map_eq; intros j. rewrite lookup_fmap, lookup_gset_to_gmap.
destruct (m !! j) as [x|] eqn:?.
- by rewrite option_guard_True by (rewrite elem_of_dom; eauto).
- by rewrite option_guard_False by (rewrite not_elem_of_dom; eauto).
Qed.
Lemma dom_gset_to_gmap {A} (X : gset K) (x : A) :
dom (gset_to_gmap x X) = X.
Proof.
induction X as [| y X not_in IH] using set_ind_L.
- rewrite gset_to_gmap_empty, dom_empty_L; done.
- rewrite gset_to_gmap_union_singleton, dom_insert_L, IH; done.
Qed.
Lemma gset_to_gmap_set_to_map {A} (X : gset K) (x : A) :
gset_to_gmap x X = set_to_map (.,x) X.
Proof.
apply map_eq; intros k. apply option_eq; intros y.
rewrite lookup_gset_to_gmap_Some, lookup_set_to_map; naive_solver.
Qed.
Lemma map_to_list_gset_to_gmap {A} (X : gset K) (x : A) :
map_to_list (gset_to_gmap x X) (., x) <$> elements X.
Proof.
induction X as [| y X not_in IH] using set_ind_L.
- rewrite gset_to_gmap_empty, elements_empty, map_to_list_empty. done.
- rewrite gset_to_gmap_union_singleton, elements_union_singleton by done.
rewrite map_to_list_insert.
2:{ rewrite lookup_gset_to_gmap_None. done. }
rewrite IH. done.
Qed.
End gset.
Section gset_cprod.
Context `{Countable A, Countable B}.
Global Instance gset_cprod : CProd (gset A) (gset B) (gset (A * B)) :=
λ X Y, set_bind (λ e1, set_map (e1,.) Y) X.
Lemma elem_of_gset_cprod (X : gset A) (Y : gset B) x :
x cprod X Y x.1 X x.2 Y.
Proof. unfold cprod, gset_cprod. destruct x. set_solver. Qed.
Global Instance set_unfold_gset_cprod (X : gset A) (Y : gset B) x (P : Prop) Q :
SetUnfoldElemOf x.1 X P SetUnfoldElemOf x.2 Y Q
SetUnfoldElemOf x (cprod X Y) (P Q).
Proof using.
intros ??; constructor.
by rewrite elem_of_gset_cprod, (set_unfold_elem_of x.1 X P),
(set_unfold_elem_of x.2 Y Q).
Qed.
End gset_cprod.
Global Typeclasses Opaque gset.
From stdpp Require Export countable.
From stdpp Require Import gmap.
From stdpp Require ssreflect. (* don't import yet, but we'll later do that to use ssreflect rewrite *)
From stdpp Require Import options.
Record gmultiset A `{Countable A} := GMultiSet { gmultiset_car : gmap A nat }.
Arguments GMultiSet {_ _ _} _ : assert.
Arguments gmultiset_car {_ _ _} _ : assert.
(** Multisets [gmultiset A] are represented as maps from [A] to natural numbers,
which represent the multiplicity. To ensure we have canonical representations,
the multiplicity is a [positive]. Therefore, [gmultiset_car !! x = None] means
[x] has multiplicity [0] and [gmultiset_car !! x = Some 1] means [x] has
multiplicity 1. *)
Instance gmultiset_eq_dec `{Countable A} : EqDecision (gmultiset A).
Record gmultiset A `{Countable A} := GMultiSet { gmultiset_car : gmap A positive }.
Global Arguments GMultiSet {_ _ _} _ : assert.
Global Arguments gmultiset_car {_ _ _} _ : assert.
Global Instance gmultiset_eq_dec `{Countable A} : EqDecision (gmultiset A).
Proof. solve_decision. Defined.
Program Instance gmultiset_countable `{Countable A} :
Global Program Instance gmultiset_countable `{Countable A} :
Countable (gmultiset A) := {|
encode X := encode (gmultiset_car X); decode p := GMultiSet <$> decode p
|}.
......@@ -19,7 +26,7 @@ Section definitions.
Context `{Countable A}.
Definition multiplicity (x : A) (X : gmultiset A) : nat :=
match gmultiset_car X !! x with Some n => S n | None => 0 end.
match gmultiset_car X !! x with Some n => Pos.to_nat n | None => 0 end.
Global Instance gmultiset_elem_of : ElemOf A (gmultiset A) := λ x X,
0 < multiplicity x X.
Global Instance gmultiset_subseteq : SubsetEq (gmultiset A) := λ X Y, x,
......@@ -28,35 +35,45 @@ Section definitions.
multiplicity x X = multiplicity x Y.
Global Instance gmultiset_elements : Elements A (gmultiset A) := λ X,
let (X) := X in '(x,n) map_to_list X; replicate (S n) x.
let (X) := X in '(x,n) map_to_list X; replicate (Pos.to_nat n) x.
Global Instance gmultiset_size : Size (gmultiset A) := length elements.
Global Instance gmultiset_empty : Empty (gmultiset A) := GMultiSet ∅.
Global Instance gmultiset_singleton : Singleton A (gmultiset A) := λ x,
GMultiSet {[ x := 0 ]}.
Global Instance gmultiset_singleton : SingletonMS A (gmultiset A) := λ x,
GMultiSet {[ x := 1%positive ]}.
Global Instance gmultiset_union : Union (gmultiset A) := λ X Y,
let (X) := X in let (Y) := Y in
GMultiSet $ union_with (λ x y, Some (x `max` y)) X Y.
GMultiSet $ union_with (λ x y, Some (x `max` y)%positive) X Y.
Global Instance gmultiset_intersection : Intersection (gmultiset A) := λ X Y,
let (X) := X in let (Y) := Y in
GMultiSet $ intersection_with (λ x y, Some (x `min` y)) X Y.
GMultiSet $ intersection_with (λ x y, Some (x `min` y)%positive) X Y.
(** Often called the "sum" *)
Global Instance gmultiset_disj_union : DisjUnion (gmultiset A) := λ X Y,
let (X) := X in let (Y) := Y in
GMultiSet $ union_with (λ x y, Some (S (x + y))) X Y.
GMultiSet $ union_with (λ x y, Some (x + y)%positive) X Y.
Global Instance gmultiset_difference : Difference (gmultiset A) := λ X Y,
let (X) := X in let (Y) := Y in
GMultiSet $ difference_with (λ x y,
let z := x - y in guard (0 < z); Some (pred z)) X Y.
guard (y < x)%positive;; Some (x - y)%positive) X Y.
Global Instance gmultiset_scalar_mul : ScalarMul nat (gmultiset A) := λ n X,
let (X) := X in GMultiSet $
match n with 0 => | _ => fmap (λ m, m * Pos.of_nat n)%positive X end.
Global Instance gmultiset_dom : Dom (gmultiset A) (gset A) := λ X,
let (X) := X in dom _ X.
let (X) := X in dom X.
Definition gmultiset_map `{Countable B} (f : A B)
(X : gmultiset A) : gmultiset B :=
GMultiSet $ map_fold
(λ x n, partial_alter (Some from_option (Pos.add n) n) (f x))
(gmultiset_car X).
End definitions.
Typeclasses Opaque gmultiset_elem_of gmultiset_subseteq.
Typeclasses Opaque gmultiset_elements gmultiset_size gmultiset_empty.
Typeclasses Opaque gmultiset_singleton gmultiset_union gmultiset_difference.
Typeclasses Opaque gmultiset_dom.
Global Typeclasses Opaque gmultiset_elem_of gmultiset_subseteq.
Global Typeclasses Opaque gmultiset_elements gmultiset_size gmultiset_empty.
Global Typeclasses Opaque gmultiset_singleton gmultiset_union gmultiset_difference.
Global Typeclasses Opaque gmultiset_scalar_mul gmultiset_dom gmultiset_map.
Section basic_lemmas.
Context `{Countable A}.
......@@ -68,7 +85,7 @@ Section basic_lemmas.
split; [by intros ->|intros HXY].
destruct X as [X], Y as [Y]; f_equal; apply map_eq; intros x.
specialize (HXY x); unfold multiplicity in *; simpl in *.
repeat case_match; naive_solver.
repeat case_match; naive_solver lia.
Qed.
Global Instance gmultiset_leibniz : LeibnizEquiv (gmultiset A).
Proof. intros X Y. by rewrite gmultiset_eq. Qed.
......@@ -78,12 +95,12 @@ Section basic_lemmas.
(* Multiplicity *)
Lemma multiplicity_empty x : multiplicity x = 0.
Proof. done. Qed.
Lemma multiplicity_singleton x : multiplicity x {[ x ]} = 1.
Lemma multiplicity_singleton x : multiplicity x {[+ x +]} = 1.
Proof. unfold multiplicity; simpl. by rewrite lookup_singleton. Qed.
Lemma multiplicity_singleton_ne x y : x y multiplicity x {[ y ]} = 0.
Lemma multiplicity_singleton_ne x y : x y multiplicity x {[+ y +]} = 0.
Proof. intros. unfold multiplicity; simpl. by rewrite lookup_singleton_ne. Qed.
Lemma multiplicity_singleton' x y :
multiplicity x {[ y ]} = if decide (x = y) then 1 else 0.
multiplicity x {[+ y +]} = if decide (x = y) then 1 else 0.
Proof.
destruct (decide _) as [->|].
- by rewrite multiplicity_singleton.
......@@ -114,32 +131,40 @@ Section basic_lemmas.
rewrite lookup_difference_with.
destruct (X !! _), (Y !! _); simplify_option_eq; lia.
Qed.
Lemma multiplicity_scalar_mul n X x :
multiplicity x (n *: X) = n * multiplicity x X.
Proof.
destruct X as [X]; unfold multiplicity; simpl. destruct n as [|n]; [done|].
rewrite lookup_fmap. destruct (X !! _); simpl; lia.
Qed.
(* Set *)
Lemma elem_of_multiplicity x X : x X 0 < multiplicity x X.
Proof. done. Qed.
Global Instance gmultiset_simple_set : SemiSet A (gmultiset A).
Lemma gmultiset_elem_of_empty x : x ∈@{gmultiset A} False.
Proof. rewrite elem_of_multiplicity, multiplicity_empty. lia. Qed.
Lemma gmultiset_elem_of_singleton x y : x ∈@{gmultiset A} {[+ y +]} x = y.
Proof.
split.
- intros x. rewrite elem_of_multiplicity, multiplicity_empty. lia.
- intros x y.
rewrite elem_of_multiplicity, multiplicity_singleton'.
destruct (decide (x = y)); intuition lia.
- intros X Y x. rewrite !elem_of_multiplicity, multiplicity_union. lia.
rewrite elem_of_multiplicity, multiplicity_singleton'.
case_decide; naive_solver lia.
Qed.
Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}).
Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined.
Lemma gmultiset_elem_of_union X Y x : x X Y x X x Y.
Proof. rewrite !elem_of_multiplicity, multiplicity_union. lia. Qed.
Lemma gmultiset_elem_of_disj_union X Y x : x X Y x X x Y.
Proof. rewrite !elem_of_multiplicity, multiplicity_disj_union. lia. Qed.
Lemma gmultiset_elem_of_intersection X Y x : x X Y x X x Y.
Proof. rewrite !elem_of_multiplicity, multiplicity_intersection. lia. Qed.
Lemma gmultiset_elem_of_scalar_mul n X x : x n *: X n 0 x X.
Proof. rewrite !elem_of_multiplicity, multiplicity_scalar_mul. lia. Qed.
Global Instance set_unfold_gmultiset_disj_union x X Y P Q :
SetUnfoldElemOf x X P SetUnfoldElemOf x Y Q
SetUnfoldElemOf x (X Y) (P Q).
Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}).
Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined.
Lemma gmultiset_elem_of_dom x X : x dom X x X.
Proof.
intros ??; constructor. rewrite gmultiset_elem_of_disj_union.
by rewrite <-(set_unfold_elem_of x X P), <-(set_unfold_elem_of x Y Q).
unfold dom, gmultiset_dom, elem_of at 2, gmultiset_elem_of, multiplicity.
destruct X as [X]; simpl; rewrite elem_of_dom, <-not_eq_None_Some.
destruct (X !! x); naive_solver lia.
Qed.
End basic_lemmas.
......@@ -147,30 +172,47 @@ End basic_lemmas.
(** We define a tactic [multiset_solver] that solves goals involving multisets.
The strategy of this tactic is as follows:
1. Unfold all equalities ([=]), equivalences ([≡]), and inclusions ([⊆]) using
the laws of [multiplicity] for the multiset operations. Note that strict
inclusion ([⊂]) is not supported.
2. Use [naive_solver] to decompose the goal into smaller subgoals.
3. Instantiate all universally quantified hypotheses in the subgoals generated
by [naive_solver] to obtain goals that can be solved using [lia].
1. Turn all equalities ([=]), equivalences ([≡]), inclusions ([⊆] and [⊂]),
and set membership relations ([∈]) into arithmetic (in)equalities
involving [multiplicity]. The multiplicities of [∅], [∪], [∩], [⊎] and [∖]
are turned into [0], [max], [min], [+], and [-], respectively.
2. Decompose the goal into smaller subgoals through intuitionistic reasoning.
3. Instantiate universally quantified hypotheses in hypotheses to obtain a
goal that can be solved using [lia].
4. Simplify multiplicities of singletons [{[ x ]}].
Step (1) is implemented using a type class [MultisetUnfold] that hooks into
the [SetUnfold] mechanism of [set_solver]. Since [SetUnfold] already propagates
through logical connectives, we obtain the same behavior for our multiset
solver. Note that no [MultisetUnfold] instance is defined for the (non-trivial)
singleton [{[ y ]}] since the singleton receives special treatment in step (3).
Step (1) and (2) are implemented using the [set_solver] tactic, which internally
calls [naive_solver] for step (2). Step (1) is implemented by extending the
[SetUnfold] mechanism with a class [MultisetUnfold].
Step (3) is achieved using the tactic [multiset_instantiate], which instantiates
universally quantified hypotheses [H : ∀ x : A, P x] in two ways:
Step (3) is implemented using the tactic [multiset_instantiate], which
instantiates universally quantified hypotheses [H : ∀ x : A, P x] in two ways:
- If [P] contains a multiset singleton [{[ y ]}] it adds the hypothesis [H y].
- If the goal or some hypothesis contains [multiplicity y X] it adds the
hypothesis [H y].
*)
- If [P] contains a multiset singleton [{[ y ]}] it adds the hypothesis [H y].
This is needed, for example, to prove [¬ ({[ x ]} ⊆ ∅)], which is turned
into hypothesis [H : ∀ y, multiplicity y {[ x ]} ≤ 0] and goal [False]. The
only way to make progress is to instantiate [H] with the singleton appearing
in [H], so variable [x].
Step (4) is implemented using the tactic [multiset_simplify_singletons], which
simplifies occurrences of [multiplicity x {[ y ]}] as follows:
- First, we try to turn these occurencess into [1] or [0] if either [x = y] or
[x ≠ y] can be proved using [done], respectively.
- Second, we try to turn these occurrences into a fresh [z ≤ 1] if [y] does not
occur elsewhere in the hypotheses or goal.
- Finally, we make a case distinction between [x = y] or [x ≠ y]. This step is
done last so as to avoid needless exponential blow-ups.
The tests [test_big_X] in [tests/multiset_solver.v] show the second step reduces
the running time significantly (from >10 seconds to <1 second). *)
Class MultisetUnfold `{Countable A} (x : A) (X : gmultiset A) (n : nat) :=
{ multiset_unfold : multiplicity x X = n }.
Arguments multiset_unfold {_ _ _} _ _ _ {_} : assert.
Hint Mode MultisetUnfold + + + - + - : typeclass_instances.
Global Arguments multiset_unfold {_ _ _} _ _ _ {_} : assert.
Global Hint Mode MultisetUnfold + + + - + - : typeclass_instances.
Section multiset_unfold.
Context `{Countable A}.
......@@ -182,8 +224,8 @@ Section multiset_unfold.
Proof. done. Qed.
Global Instance multiset_unfold_empty x : MultisetUnfold x 0.
Proof. constructor. by rewrite multiplicity_empty. Qed.
Global Instance multiset_unfold_singleton x y :
MultisetUnfold x {[ x ]} 1.
Global Instance multiset_unfold_singleton x :
MultisetUnfold x {[+ x +]} 1.
Proof. constructor. by rewrite multiplicity_singleton. Qed.
Global Instance multiset_unfold_union x X Y n m :
MultisetUnfold x X n MultisetUnfold x Y m
......@@ -201,59 +243,131 @@ Section multiset_unfold.
MultisetUnfold x X n MultisetUnfold x Y m
MultisetUnfold x (X Y) (n - m).
Proof. intros [HX] [HY]; constructor. by rewrite multiplicity_difference, HX, HY. Qed.
Global Instance multiset_unfold_scalar_mul x m X n :
MultisetUnfold x X n
MultisetUnfold x (m *: X) (m * n).
Proof. intros [HX]; constructor. by rewrite multiplicity_scalar_mul, HX. Qed.
Global Instance set_unfold_multiset_equiv X Y f g :
( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
SetUnfold (X Y) ( x, f x = g x).
SetUnfold (X Y) ( x, f x = g x) | 0.
Proof.
constructor. apply forall_proper; intros x.
by rewrite (multiset_unfold x X (f x)), (multiset_unfold x Y (g x)).
Qed.
Global Instance set_unfold_multiset_eq X Y f g :
( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
SetUnfold (X = Y) ( x, f x = g x).
SetUnfold (X = Y) ( x, f x = g x) | 0.
Proof. constructor. unfold_leibniz. by apply set_unfold_multiset_equiv. Qed.
Global Instance set_unfold_multiset_subseteq X Y f g :
( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
SetUnfold (X Y) ( x, f x g x).
SetUnfold (X Y) ( x, f x g x) | 0.
Proof.
constructor. apply forall_proper; intros x.
by rewrite (multiset_unfold x X (f x)), (multiset_unfold x Y (g x)).
Qed.
Global Instance set_unfold_multiset_subset X Y f g :
( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
SetUnfold (X Y) (( x, f x g x) (¬∀ x, g x f x)) | 0.
Proof.
constructor. unfold strict. f_equiv.
- by apply set_unfold_multiset_subseteq.
- f_equiv. by apply set_unfold_multiset_subseteq.
Qed.
Global Instance set_unfold_multiset_elem_of X x n :
MultisetUnfold x X n SetUnfoldElemOf x X (0 < n) | 100.
Proof. constructor. by rewrite <-(multiset_unfold x X n). Qed.
Global Instance set_unfold_gmultiset_empty x :
SetUnfoldElemOf x ( : gmultiset A) False.
Proof. constructor. apply gmultiset_elem_of_empty. Qed.
Global Instance set_unfold_gmultiset_singleton x y :
SetUnfoldElemOf x ({[+ y +]} : gmultiset A) (x = y).
Proof. constructor; apply gmultiset_elem_of_singleton. Qed.
Global Instance set_unfold_gmultiset_union x X Y P Q :
SetUnfoldElemOf x X P SetUnfoldElemOf x Y Q
SetUnfoldElemOf x (X Y) (P Q).
Proof.
intros ??; constructor. by rewrite gmultiset_elem_of_union,
(set_unfold_elem_of x X P), (set_unfold_elem_of x Y Q).
Qed.
Global Instance set_unfold_gmultiset_disj_union x X Y P Q :
SetUnfoldElemOf x X P SetUnfoldElemOf x Y Q
SetUnfoldElemOf x (X Y) (P Q).
Proof.
intros ??; constructor. rewrite gmultiset_elem_of_disj_union.
by rewrite <-(set_unfold_elem_of x X P), <-(set_unfold_elem_of x Y Q).
Qed.
Global Instance set_unfold_gmultiset_intersection x X Y P Q :
SetUnfoldElemOf x X P SetUnfoldElemOf x Y Q
SetUnfoldElemOf x (X Y) (P Q).
Proof.
intros ??; constructor. rewrite gmultiset_elem_of_intersection.
by rewrite (set_unfold_elem_of x X P), (set_unfold_elem_of x Y Q).
Qed.
Global Instance set_unfold_gmultiset_dom x X :
SetUnfoldElemOf x (dom X) (x X).
Proof. constructor. apply gmultiset_elem_of_dom. Qed.
End multiset_unfold.
(** Step 3: instantiate hypotheses *)
(** For these tactics we want to use ssreflect rewrite. ssreflect matching
interacts better with canonical structures (see
<https://gitlab.mpi-sws.org/iris/stdpp/-/issues/195>). *)
Module Export tactics.
Import ssreflect.
Ltac multiset_instantiate :=
(* Step 3.1: instantiate hypotheses *)
repeat match goal with
| H : ( x : ?A, @?P x) |- _ =>
let e := fresh in evar (e:A);
let e' := eval unfold e in e in clear e;
lazymatch constr:(P e') with
| context [ {[ ?y ]} ] =>
(* Use [unless] to avoid creating a new hypothesis [H y : P y] if [P y]
already exists. *)
unify y e'; unless (P y) by assumption; pose proof (H y)
let e := mk_evar A in
lazymatch constr:(P e) with
| context [ {[+ ?y +]} ] => unify y e; learn_hyp (H y)
end
| H : ( x : ?A, @?P x), _ : context [multiplicity ?y _] |- _ =>
(* Use [unless] to avoid creating a new hypothesis [H y : P y] if [P y]
already exists. *)
unless (P y) by assumption; pose proof (H y)
| H : ( x : ?A, @?P x) |- context [multiplicity ?y _] =>
(* Use [unless] to avoid creating a new hypothesis [H y : P y] if [P y]
already exists. *)
unless (P y) by assumption; pose proof (H y)
end;
(* Step 3.2: simplify singletons. *)
(* Note that we do not use [repeat case_decide] since that leads to an
exponential explosion in the number of singletons. *)
repeat match goal with
| H : context [multiplicity _ {[ _ ]}] |- _ =>
progress (rewrite ?multiplicity_singleton, ?multiplicity_singleton_ne in H by done)
| |- context [multiplicity _ {[ _ ]}] =>
progress (rewrite ?multiplicity_singleton, ?multiplicity_singleton_ne by done)
| H : ( x : ?A, _), _ : context [multiplicity ?y _] |- _ => learn_hyp (H y)
| H : ( x : ?A, _) |- context [multiplicity ?y _] => learn_hyp (H y)
end.
Ltac multiset_solver := set_solver by (multiset_instantiate; lia).
(** Step 4: simplify singletons *)
(** This lemma results in information loss if there are other occurrences of
[y] in the goal. In the tactic [multiset_simplify_singletons] we use [clear y]
to ensure we do not use the lemma if it leads to information loss. *)
Local Lemma multiplicity_singleton_forget `{Countable A} x y :
n, multiplicity (A:=A) x {[+ y +]} = n n 1.
Proof. rewrite multiplicity_singleton'. case_decide; eauto with lia. Qed.
Ltac multiset_simplify_singletons :=
repeat match goal with
| H : context [multiplicity ?x {[+ ?y +]}] |- _ =>
first
[progress rewrite ?multiplicity_singleton ?multiplicity_singleton_ne in H; [|done..]
(* This second case does *not* use ssreflect matching (due to [destruct]
and the [->] pattern). If the default Coq matching goes wrong it will
fail and fall back to the third case, which is strictly more general,
just slower. *)
|destruct (multiplicity_singleton_forget x y) as (?&->&?); clear y
|rewrite multiplicity_singleton' in H; destruct (decide (x = y)); simplify_eq/=]
| |- context [multiplicity ?x {[+ ?y +]}] =>
first
[progress rewrite ?multiplicity_singleton ?multiplicity_singleton_ne; [|done..]
(* Similar to above, this second case does *not* use ssreflect matching. *)
|destruct (multiplicity_singleton_forget x y) as (?&->&?); clear y
|rewrite multiplicity_singleton'; destruct (decide (x = y)); simplify_eq/=]
end.
End tactics.
(** Putting it all together *)
(** Similar to [set_solver] and [naive_solver], [multiset_solver] has a [by]
parameter whose default is [eauto]. *)
Tactic Notation "multiset_solver" "by" tactic3(tac) :=
set_solver by (multiset_instantiate;
multiset_simplify_singletons;
(* [fast_done] to solve trivial equalities or contradictions,
[lia] for the common case that involves arithmetic,
[tac] if all else fails *)
solve [fast_done|lia|tac]).
Tactic Notation "multiset_solver" := multiset_solver by eauto.
Section more_lemmas.
Context `{Countable A}.
......@@ -319,41 +433,97 @@ Section more_lemmas.
Lemma gmultiset_disj_union_union_r X Y Z : (X Y) Z = (X Z) (Y Z).
Proof. multiset_solver. Qed.
(** Element of operation *)
Lemma gmultiset_not_elem_of_empty x : x ∉@{gmultiset A} ∅.
Proof. multiset_solver. Qed.
Lemma gmultiset_not_elem_of_singleton x y : x ∉@{gmultiset A} {[+ y +]} x y.
Proof. multiset_solver. Qed.
Lemma gmultiset_not_elem_of_union x X Y : x X Y x X x Y.
Proof. multiset_solver. Qed.
Lemma gmultiset_not_elem_of_intersection x X Y : x X Y x X x Y.
Proof. multiset_solver. Qed.
(** Misc *)
Lemma gmultiset_non_empty_singleton x : {[ x ]} ≠@{gmultiset A} ∅.
Global Instance gmultiset_singleton_inj : Inj (=) (=@{gmultiset A}) singletonMS.
Proof.
intros x1 x2 Hx. rewrite gmultiset_eq in Hx. specialize (Hx x1).
rewrite multiplicity_singleton, multiplicity_singleton' in Hx.
case_decide; [done|lia].
Qed.
Lemma gmultiset_non_empty_singleton x : {[+ x +]} ≠@{gmultiset A} ∅.
Proof. multiset_solver. Qed.
(** Scalar *)
Lemma gmultiset_scalar_mul_0 X : 0 *: X = ∅.
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_S_l n X : S n *: X = X (n *: X).
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_S_r n X : S n *: X = (n *: X) X.
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_1 X : 1 *: X = X.
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_2 X : 2 *: X = X X.
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_empty n : n *: =@{gmultiset A} ∅.
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_disj_union n X Y :
n *: (X Y) =@{gmultiset A} (n *: X) (n *: Y).
Proof. multiset_solver. Qed.
Lemma gmultiset_scalar_mul_union n X Y :
n *: (X Y) =@{gmultiset A} (n *: X) (n *: Y).
Proof. set_unfold. intros x; by rewrite Nat.mul_max_distr_l. Qed.
Lemma gmultiset_scalar_mul_intersection n X Y :
n *: (X Y) =@{gmultiset A} (n *: X) (n *: Y).
Proof. set_unfold. intros x; by rewrite Nat.mul_min_distr_l. Qed.
Lemma gmultiset_scalar_mul_difference n X Y :
n *: (X Y) =@{gmultiset A} (n *: X) (n *: Y).
Proof. set_unfold. intros x; by rewrite Nat.mul_sub_distr_l. Qed.
Lemma gmultiset_scalar_mul_inj_ne_0 n X1 X2 :
n 0 n *: X1 = n *: X2 X1 = X2.
Proof. set_unfold. intros ? HX x. apply (Nat.mul_reg_l _ _ n); auto. Qed.
(** Specialized to [S n] so that type class search can find it. *)
Global Instance gmultiset_scalar_mul_inj_S n :
Inj (=) (=@{gmultiset A}) (S n *:.).
Proof. intros x1 x2. apply gmultiset_scalar_mul_inj_ne_0. lia. Qed.
(** Conversion from lists *)
Lemma list_to_set_disj_nil : list_to_set_disj [] =@{gmultiset A} ∅.
Proof. done. Qed.
Lemma list_to_set_disj_cons x l :
list_to_set_disj (x :: l) =@{gmultiset A} {[ x ]} list_to_set_disj l.
list_to_set_disj (x :: l) =@{gmultiset A} {[+ x +]} list_to_set_disj l.
Proof. done. Qed.
Lemma list_to_set_disj_app l1 l2 :
list_to_set_disj (l1 ++ l2) =@{gmultiset A} list_to_set_disj l1 list_to_set_disj l2.
Proof. induction l1; multiset_solver. Qed.
Lemma elem_of_list_to_set_disj x l :
x ∈@{gmultiset A} list_to_set_disj l x l.
Proof. induction l; set_solver. Qed.
Global Instance list_to_set_disj_perm :
Proper (() ==> (=)) (list_to_set_disj (C:=gmultiset A)).
Proof. induction 1; multiset_solver. Qed.
Lemma list_to_set_disj_replicate n x :
list_to_set_disj (replicate n x) =@{gmultiset A} n *: {[+ x +]}.
Proof. induction n; multiset_solver. Qed.
(** Properties of the elements operation *)
Lemma gmultiset_elements_empty : elements ( : gmultiset A) = [].
Proof.
unfold elements, gmultiset_elements; simpl. by rewrite map_to_list_empty.
Qed.
Lemma gmultiset_elements_empty_inv X : elements X = [] X = ∅.
Lemma gmultiset_elements_empty_iff X : elements X = [] X = ∅.
Proof.
split; [|intros ->; by rewrite gmultiset_elements_empty].
destruct X as [X]; unfold elements, gmultiset_elements; simpl.
intros; apply (f_equal GMultiSet). destruct (map_to_list X) as [|[]] eqn:?.
- by apply map_to_list_empty_inv.
- naive_solver.
Qed.
Lemma gmultiset_elements_empty' X : elements X = [] X = ∅.
Proof.
split; intros HX; [by apply gmultiset_elements_empty_inv|].
by rewrite HX, gmultiset_elements_empty.
intros; apply (f_equal GMultiSet).
destruct (map_to_list X) as [|[x p]] eqn:?; simpl in *.
- by apply map_to_list_empty_iff.
- pose proof (Pos2Nat.is_pos p). destruct (Pos.to_nat); naive_solver lia.
Qed.
Lemma gmultiset_elements_singleton x : elements ({[ x ]} : gmultiset A) = [ x ].
Lemma gmultiset_elements_empty_inv X : elements X = [] X = ∅.
Proof. apply gmultiset_elements_empty_iff. Qed.
Lemma gmultiset_elements_singleton x : elements ({[+ x +]} : gmultiset A) = [ x ].
Proof.
unfold elements, gmultiset_elements; simpl. by rewrite map_to_list_singleton.
Qed.
......@@ -361,25 +531,32 @@ Section more_lemmas.
elements (X Y) elements X ++ elements Y.
Proof.
destruct X as [X], Y as [Y]; unfold elements, gmultiset_elements.
set (f xn := let '(x, n) := xn in replicate (S n) x); simpl.
set (f xn := let '(x, n) := xn in replicate (Pos.to_nat n) x); simpl.
revert Y; induction X as [|x n X HX IH] using map_ind; intros Y.
{ by rewrite (left_id_L _ _ Y), map_to_list_empty. }
destruct (Y !! x) as [n'|] eqn:HY.
- rewrite <-(insert_id Y x n'), <-(insert_delete Y) by done.
- rewrite <-(insert_delete Y x n') by done.
erewrite <-insert_union_with by done.
rewrite !map_to_list_insert, !bind_cons
by (by rewrite ?lookup_union_with, ?lookup_delete, ?HX).
rewrite (assoc_L _), <-(comm (++) (f (_,n'))), <-!(assoc_L _), <-IH.
rewrite (assoc_L _). f_equiv.
rewrite (comm _); simpl. by rewrite replicate_plus, Permutation_middle.
rewrite (comm _); simpl. by rewrite Pos2Nat.inj_add, replicate_add.
- rewrite <-insert_union_with_l, !map_to_list_insert, !bind_cons
by (by rewrite ?lookup_union_with, ?HX, ?HY).
by rewrite <-(assoc_L (++)), <-IH.
Qed.
Lemma gmultiset_elements_scalar_mul n X :
elements (n *: X) mjoin (replicate n (elements X)).
Proof.
induction n as [|n IH]; simpl.
- by rewrite gmultiset_scalar_mul_0, gmultiset_elements_empty.
- by rewrite gmultiset_scalar_mul_S_l, gmultiset_elements_disj_union, IH.
Qed.
Lemma gmultiset_elem_of_elements x X : x elements X x X.
Proof.
destruct X as [X]. unfold elements, gmultiset_elements.
set (f xn := let '(x, n) := xn in replicate (S n) x); simpl.
set (f xn := let '(x, n) := xn in replicate (Pos.to_nat n) x); simpl.
unfold elem_of at 2, gmultiset_elem_of, multiplicity; simpl.
rewrite elem_of_list_bind. split.
- intros [[??] [[<- ?]%elem_of_replicate ->%elem_of_map_to_list]]; lia.
......@@ -387,49 +564,79 @@ Section more_lemmas.
exists (x,n); split; [|by apply elem_of_map_to_list].
apply elem_of_replicate; auto with lia.
Qed.
Lemma gmultiset_elem_of_dom x X : x dom (gset A) X x X.
Proof.
unfold dom, gmultiset_dom, elem_of at 2, gmultiset_elem_of, multiplicity.
destruct X as [X]; simpl; rewrite elem_of_dom, <-not_eq_None_Some.
destruct (X !! x); naive_solver lia.
Qed.
(** Properties of the set_fold operation *)
Lemma gmultiset_set_fold_empty {B} (f : A B B) (b : B) :
set_fold f b ( : gmultiset A) = b.
Proof. by unfold set_fold; simpl; rewrite gmultiset_elements_empty. Qed.
Lemma gmultiset_set_fold_singleton {B} (f : A B B) (b : B) (a : A) :
set_fold f b ({[a]} : gmultiset A) = f a b.
set_fold f b ({[+ a +]} : gmultiset A) = f a b.
Proof. by unfold set_fold; simpl; rewrite gmultiset_elements_singleton. Qed.
Lemma gmultiset_set_fold_disj_union_strong {B} (R : relation B) `{!PreOrder R}
(f : A B B) (b : B) X Y :
( x, Proper (R ==> R) (f x))
( x1 x2 c, x1 X Y x2 X Y R (f x1 (f x2 c)) (f x2 (f x1 c)))
R (set_fold f b (X Y)) (set_fold f (set_fold f b X) Y).
Proof.
intros ? Hf. unfold set_fold; simpl.
rewrite <-foldr_app. apply (foldr_permutation R f b).
- intros j1 a1 j2 a2 c ? Ha1%elem_of_list_lookup_2 Ha2%elem_of_list_lookup_2.
rewrite gmultiset_elem_of_elements in Ha1, Ha2. eauto.
- rewrite (comm (++)). apply gmultiset_elements_disj_union.
Qed.
Lemma gmultiset_set_fold_disj_union (f : A A A) (b : A) X Y :
Comm (=) f
Assoc (=) f
set_fold f b (X Y) = set_fold f (set_fold f b X) Y.
Proof.
intros Hcomm Hassoc. unfold set_fold; simpl.
by rewrite gmultiset_elements_disj_union, <- foldr_app, (comm (++)).
intros ??; apply gmultiset_set_fold_disj_union_strong; [apply _..|].
intros x1 x2 ? _ _. by rewrite 2!assoc, (comm f x1 x2).
Qed.
Lemma gmultiset_set_fold_scalar_mul (f : A A A) (b : A) n X :
Comm (=) f
Assoc (=) f
set_fold f b (n *: X) = Nat.iter n (flip (set_fold f) X) b.
Proof.
intros Hcomm Hassoc. induction n as [|n IH]; simpl.
- by rewrite gmultiset_scalar_mul_0, gmultiset_set_fold_empty.
- rewrite gmultiset_scalar_mul_S_r.
by rewrite (gmultiset_set_fold_disj_union _ _ _ _ _ _), IH.
Qed.
Lemma gmultiset_set_fold_comm_acc_strong {B} (R : relation B) `{!PreOrder R}
(f : A B B) (g : B B) b X :
( x, Proper (R ==> R) (f x))
( x (y : B), x X R (f x (g y)) (g (f x y)))
R (set_fold f (g b) X) (g (set_fold f b X)).
Proof.
intros ? Hfg. unfold set_fold; simpl.
apply foldr_comm_acc_strong; [done|solve_proper|].
intros. by apply Hfg, gmultiset_elem_of_elements.
Qed.
Lemma gmultiset_set_fold_comm_acc {B} (f : A B B) (g : B B) (b : B) X :
( x c, g (f x c) = f x (g c))
set_fold f (g b) X = g (set_fold f b X).
Proof.
intros. apply (gmultiset_set_fold_comm_acc_strong _); [solve_proper|done].
Qed.
(** Properties of the size operation *)
Lemma gmultiset_size_empty : size ( : gmultiset A) = 0.
Proof. done. Qed.
Lemma gmultiset_size_empty_inv X : size X = 0 X = ∅.
Proof.
unfold size, gmultiset_size; simpl. rewrite length_zero_iff_nil.
apply gmultiset_elements_empty_inv.
Qed.
Lemma gmultiset_size_empty_iff X : size X = 0 X = ∅.
Proof.
split; [apply gmultiset_size_empty_inv|].
by intros ->; rewrite gmultiset_size_empty.
unfold size, gmultiset_size; simpl.
by rewrite length_zero_iff_nil, gmultiset_elements_empty_iff.
Qed.
Lemma gmultiset_size_empty_inv X : size X = 0 X = ∅.
Proof. apply gmultiset_size_empty_iff. Qed.
Lemma gmultiset_size_non_empty_iff X : size X 0 X ∅.
Proof. by rewrite gmultiset_size_empty_iff. Qed.
Lemma gmultiset_choose_or_empty X : ( x, x X) X = ∅.
Proof.
destruct (elements X) as [|x l] eqn:HX; [right|left].
- by apply gmultiset_elements_empty_inv.
- by apply gmultiset_elements_empty_iff.
- exists x. rewrite <-gmultiset_elem_of_elements, HX. by left.
Qed.
Lemma gmultiset_choose X : X x, x X.
......@@ -440,37 +647,44 @@ Section more_lemmas.
contradict Hsz. rewrite HX, gmultiset_size_empty; lia.
Qed.
Lemma gmultiset_size_singleton x : size ({[ x ]} : gmultiset A) = 1.
Lemma gmultiset_size_singleton x : size ({[+ x +]} : gmultiset A) = 1.
Proof.
unfold size, gmultiset_size; simpl. by rewrite gmultiset_elements_singleton.
Qed.
Lemma gmultiset_size_disj_union X Y : size (X Y) = size X + size Y.
Proof.
unfold size, gmultiset_size; simpl.
by rewrite gmultiset_elements_disj_union, app_length.
by rewrite gmultiset_elements_disj_union, length_app.
Qed.
Lemma gmultiset_size_scalar_mul n X : size (n *: X) = n * size X.
Proof.
induction n as [|n IH].
- by rewrite gmultiset_scalar_mul_0, gmultiset_size_empty.
- rewrite gmultiset_scalar_mul_S_l, gmultiset_size_disj_union, IH. lia.
Qed.
(** Order stuff *)
Global Instance gmultiset_po : PartialOrder (⊆@{gmultiset A}).
Proof. repeat split; repeat intro; multiset_solver. Qed.
Lemma gmultiset_subseteq_alt X Y :
Local Lemma gmultiset_subseteq_alt X Y :
X Y
map_relation () (λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y).
map_relation (λ _, Pos.le) (λ _ _, False) (λ _ _, True)
(gmultiset_car X) (gmultiset_car Y).
Proof.
apply forall_proper; intros x. unfold multiplicity.
destruct (gmultiset_car X !! x), (gmultiset_car Y !! x); naive_solver lia.
Qed.
Global Instance gmultiset_subseteq_dec : RelDecision (⊆@{gmultiset A}).
Proof.
refine (λ X Y, cast_if (decide (map_relation ()
(λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y))));
refine (λ X Y, cast_if (decide (map_relation
(λ _, Pos.le) (λ _ _, False) (λ _ _, True)
(gmultiset_car X) (gmultiset_car Y))));
by rewrite gmultiset_subseteq_alt.
Defined.
Lemma gmultiset_subset_subseteq X Y : X Y X Y.
Proof. apply strict_include. Qed.
Hint Resolve gmultiset_subset_subseteq : core.
Proof. multiset_solver. Qed.
Lemma gmultiset_empty_subseteq X : X.
Proof. multiset_solver. Qed.
......@@ -500,32 +714,24 @@ Section more_lemmas.
Lemma gmultiset_subset X Y : X Y size X < size Y X Y.
Proof. intros. apply strict_spec_alt; split; naive_solver auto with lia. Qed.
Lemma gmultiset_disj_union_subset_l X Y : Y X X Y.
Proof.
intros HY%gmultiset_size_non_empty_iff.
apply gmultiset_subset; auto using gmultiset_disj_union_subseteq_l.
rewrite gmultiset_size_disj_union; lia.
Qed.
Proof. multiset_solver. Qed.
Lemma gmultiset_union_subset_r X Y : X Y X Y.
Proof. rewrite (comm_L ()). apply gmultiset_disj_union_subset_l. Qed.
Proof. multiset_solver. Qed.
Lemma gmultiset_elem_of_singleton_subseteq x X : x X {[ x ]} X.
Proof.
rewrite elem_of_multiplicity. split.
- intros Hx y. rewrite multiplicity_singleton'.
destruct (decide (y = x)); naive_solver lia.
- intros Hx. generalize (Hx x). rewrite multiplicity_singleton. lia.
Qed.
Lemma gmultiset_singleton_subseteq_l x X : {[+ x +]} X x X.
Proof. multiset_solver. Qed.
Lemma gmultiset_singleton_subseteq x y :
{[+ x +]} ⊆@{gmultiset A} {[+ y +]} x = y.
Proof. multiset_solver. Qed.
Lemma gmultiset_elem_of_subseteq X1 X2 x : x X1 X1 X2 x X2.
Proof. rewrite !gmultiset_elem_of_singleton_subseteq. by intros ->. Qed.
Proof. multiset_solver. Qed.
Lemma gmultiset_disj_union_difference X Y : X Y Y = X Y X.
Proof. multiset_solver. Qed.
Lemma gmultiset_disj_union_difference' x Y : x Y Y = {[ x ]} Y {[ x ]}.
Proof.
intros. by apply gmultiset_disj_union_difference,
gmultiset_elem_of_singleton_subseteq.
Qed.
Lemma gmultiset_disj_union_difference' x Y :
x Y Y = {[+ x +]} Y {[+ x +]}.
Proof. multiset_solver. Qed.
Lemma gmultiset_size_difference X Y : Y X size (X Y) = size X - size Y.
Proof.
......@@ -537,20 +743,19 @@ Section more_lemmas.
Proof. multiset_solver. Qed.
Lemma gmultiset_non_empty_difference X Y : X Y Y X ∅.
Proof.
intros [_ HXY2] Hdiff; destruct HXY2; intros x.
generalize (f_equal (multiplicity x) Hdiff).
rewrite multiplicity_difference, multiplicity_empty; lia.
Qed.
Proof. multiset_solver. Qed.
Lemma gmultiset_difference_diag X : X X = ∅.
Proof. multiset_solver. Qed.
Lemma gmultiset_difference_subset X Y : X X Y Y X Y.
Proof.
intros. eapply strict_transitive_l; [by apply gmultiset_union_subset_r|].
by rewrite <-(gmultiset_disj_union_difference X Y).
Qed.
Proof. multiset_solver. Qed.
Lemma gmultiset_difference_disj_union_r X Y Z : X Y = (X Z) (Y Z).
Proof. multiset_solver. Qed.
Lemma gmultiset_difference_disj_union_l X Y Z : X Y = (Z X) (Z Y).
Proof. multiset_solver. Qed.
(** Mononicity *)
Lemma gmultiset_elements_submseteq X Y : X Y elements X ⊆+ elements Y.
......@@ -567,22 +772,138 @@ Section more_lemmas.
intros HXY. assert (size (Y X) 0).
{ by apply gmultiset_size_non_empty_iff, gmultiset_non_empty_difference. }
rewrite (gmultiset_disj_union_difference X Y),
gmultiset_size_disj_union by auto. lia.
gmultiset_size_disj_union by auto using gmultiset_subset_subseteq. lia.
Qed.
(** Well-foundedness *)
Lemma gmultiset_wf : wf (⊂@{gmultiset A}).
Lemma gmultiset_wf : well_founded (⊂@{gmultiset A}).
Proof.
apply (wf_projected (<) size); auto using gmultiset_subset_size, lt_wf.
Qed.
Lemma gmultiset_ind (P : gmultiset A Prop) :
P ( x X, P X P ({[ x ]} X)) X, P X.
P ( x X, P X P ({[+ x +]} X)) X, P X.
Proof.
intros Hemp Hinsert X. induction (gmultiset_wf X) as [X _ IH].
destruct (gmultiset_choose_or_empty X) as [[x Hx]| ->]; auto.
rewrite (gmultiset_disj_union_difference' x X) by done.
apply Hinsert, IH, gmultiset_difference_subset,
gmultiset_elem_of_singleton_subseteq; auto using gmultiset_non_empty_singleton.
apply Hinsert, IH; multiset_solver.
Qed.
End more_lemmas.
(** * Map *)
Section map.
Context `{Countable A, Countable B}.
Context (f : A B).
Lemma gmultiset_map_alt X :
gmultiset_map f X = list_to_set_disj (f <$> elements X).
Proof.
destruct X as [m]. unfold elements, gmultiset_map. simpl.
induction m as [|x n m ?? IH] using map_first_key_ind; [done|].
rewrite map_to_list_insert_first_key, map_fold_insert_first_key by done.
csimpl. rewrite fmap_app, fmap_replicate, list_to_set_disj_app, <-IH.
apply gmultiset_eq; intros y.
rewrite multiplicity_disj_union, list_to_set_disj_replicate.
rewrite multiplicity_scalar_mul, multiplicity_singleton'.
unfold multiplicity; simpl. destruct (decide (y = f x)) as [->|].
- rewrite lookup_partial_alter; simpl. destruct (_ !! f x); simpl; lia.
- rewrite lookup_partial_alter_ne by done. lia.
Qed.
Lemma gmultiset_map_empty : gmultiset_map f = ∅.
Proof. done. Qed.
Lemma gmultiset_map_disj_union X Y :
gmultiset_map f (X Y) = gmultiset_map f X gmultiset_map f Y.
Proof.
apply gmultiset_eq; intros x.
rewrite !gmultiset_map_alt, gmultiset_elements_disj_union, fmap_app.
by rewrite list_to_set_disj_app.
Qed.
Lemma gmultiset_map_singleton x :
gmultiset_map f {[+ x +]} = {[+ f x +]}.
Proof.
rewrite gmultiset_map_alt, gmultiset_elements_singleton.
multiset_solver.
Qed.
Lemma elem_of_gmultiset_map X y :
y gmultiset_map f X x, y = f x x X.
Proof.
rewrite gmultiset_map_alt, elem_of_list_to_set_disj, elem_of_list_fmap.
by setoid_rewrite gmultiset_elem_of_elements.
Qed.
Lemma multiplicity_gmultiset_map X x :
Inj (=) (=) f
multiplicity (f x) (gmultiset_map f X) = multiplicity x X.
Proof.
intros. induction X as [|y X IH] using gmultiset_ind; [multiset_solver|].
rewrite gmultiset_map_disj_union, gmultiset_map_singleton,
!multiplicity_disj_union.
multiset_solver.
Qed.
Global Instance gmultiset_map_inj :
Inj (=) (=) f Inj (=) (=) (gmultiset_map f).
Proof.
intros ? X Y HXY. apply gmultiset_eq; intros x.
by rewrite <-!(multiplicity_gmultiset_map _ _ _), HXY.
Qed.
Global Instance set_unfold_gmultiset_map X (P : A Prop) y :
( x, SetUnfoldElemOf x X (P x))
SetUnfoldElemOf y (gmultiset_map f X) ( x, y = f x P x).
Proof. constructor. rewrite elem_of_gmultiset_map; naive_solver. Qed.
Global Instance multiset_unfold_map x X n :
Inj (=) (=) f
MultisetUnfold x X n
MultisetUnfold (f x) (gmultiset_map f X) n.
Proof.
intros ? [HX]; constructor. by rewrite multiplicity_gmultiset_map, HX.
Qed.
End map.
(** * Big disjoint unions *)
Section disj_union_list.
Context `{Countable A}.
Implicit Types X Y : gmultiset A.
Implicit Types Xs Ys : list (gmultiset A).
Lemma gmultiset_disj_union_list_nil :
⋃+ (@nil (gmultiset A)) = ∅.
Proof. done. Qed.
Lemma gmultiset_disj_union_list_cons X Xs :
⋃+ (X :: Xs) = X ⋃+ Xs.
Proof. done. Qed.
Lemma gmultiset_disj_union_list_singleton X :
⋃+ [X] = X.
Proof. simpl. by rewrite (right_id_L _). Qed.
Lemma gmultiset_disj_union_list_app Xs1 Xs2 :
⋃+ (Xs1 ++ Xs2) = ⋃+ Xs1 ⋃+ Xs2.
Proof.
induction Xs1 as [|X Xs1 IH]; simpl; [by rewrite (left_id_L _)|].
by rewrite IH, (assoc_L _).
Qed.
Lemma elem_of_gmultiset_disj_union_list Xs x :
x ⋃+ Xs X, X Xs x X.
Proof. induction Xs; multiset_solver. Qed.
Lemma multiplicity_gmultiset_disj_union_list x Xs :
multiplicity x (⋃+ Xs) = sum_list (multiplicity x <$> Xs).
Proof.
induction Xs as [|X Xs IH]; [done|]; simpl.
by rewrite multiplicity_disj_union, IH.
Qed.
Global Instance gmultiset_disj_union_list_proper :
Proper (() ==> (=)) (@disj_union_list (gmultiset A) _ _).
Proof. induction 1; multiset_solver. Qed.
End disj_union_list.
......@@ -10,8 +10,8 @@ Record hashset {A} (hash : A → Z) := Hashset {
hashset_prf :
map_Forall (λ n l, Forall (λ x, hash x = n) l NoDup l) hashset_car
}.
Arguments Hashset {_ _} _ _ : assert.
Arguments hashset_car {_ _} _ : assert.
Global Arguments Hashset {_ _} _ _ : assert.
Global Arguments hashset_car {_ _} _ : assert.
Section hashset.
Context `{EqDecision A} (hash : A Z).
......@@ -39,7 +39,7 @@ Qed.
Global Program Instance hashset_intersection: Intersection (hashset hash) := λ m1 m2,
let (m1,Hm1) := m1 in let (m2,Hm2) := m2 in
Hashset (intersection_with (λ l k,
let l' := list_intersection l k in guard (l' []); Some l') m1 m2) _.
let l' := list_intersection l k in guard (l' []);; Some l') m1 m2) _.
Next Obligation.
intros _ _ m1 Hm1 m2 Hm2 n l'. rewrite lookup_intersection_with_Some.
intros (?&?&?&?&?); simplify_option_eq.
......@@ -49,7 +49,7 @@ Qed.
Global Program Instance hashset_difference: Difference (hashset hash) := λ m1 m2,
let (m1,Hm1) := m1 in let (m2,Hm2) := m2 in
Hashset (difference_with (λ l k,
let l' := list_difference l k in guard (l' []); Some l') m1 m2) _.
let l' := list_difference l k in guard (l' []);; Some l') m1 m2) _.
Next Obligation.
intros _ _ m1 Hm1 m2 Hm2 n l'. rewrite lookup_difference_with_Some.
intros [[??]|(?&?&?&?&?)]; simplify_option_eq; auto.
......@@ -105,7 +105,7 @@ Proof.
- unfold elements, hashset_elements. intros [m Hm]; simpl.
rewrite map_Forall_to_list in Hm. generalize (NoDup_fst_map_to_list m).
induction Hm as [|[n l] m' [??] Hm];
csimpl; inversion_clear 1 as [|?? Hn]; [constructor|].
csimpl; inv 1 as [|?? Hn]; [constructor|].
apply NoDup_app; split_and?; eauto.
setoid_rewrite elem_of_list_bind; intros x ? ([n' l']&?&?); simpl in *.
assert (hash x = n hash x = n') as [??]; subst.
......@@ -116,7 +116,7 @@ Proof.
Qed.
End hashset.
Typeclasses Opaque hashset_elem_of.
Global Typeclasses Opaque hashset_elem_of.
Section remove_duplicates.
Context `{EqDecision A} (hash : A Z).
......@@ -126,7 +126,7 @@ Definition remove_dups_fast (l : list A) : list A :=
| [] => []
| [x] => [x]
| _ =>
let n : Z := length l in
let n : Z := Z.of_nat (length l) in
elements (foldr (λ x, ({[ x ]} ∪.)) l :
hashset (λ x, hash x `mod` (2 * n))%Z)
end.
......@@ -134,7 +134,7 @@ Lemma elem_of_remove_dups_fast l x : x ∈ remove_dups_fast l ↔ x ∈ l.
Proof.
destruct l as [|x1 [|x2 l]]; try reflexivity.
unfold remove_dups_fast; generalize (x1 :: x2 :: l); clear l; intros l.
generalize (λ x, hash x `mod` (2 * length l))%Z; intros f.
generalize (λ x, hash x `mod` (2 * Z.of_nat (length l)))%Z; intros f.
rewrite elem_of_elements; split.
- revert x. induction l as [|y l IH]; intros x; simpl.
{ by rewrite elem_of_empty. }
......@@ -152,6 +152,6 @@ Definition listset_normalize (X : listset A) : listset A :=
let (l) := X in Listset (remove_dups_fast l).
Lemma listset_normalize_correct X : listset_normalize X X.
Proof.
destruct X as [l]. apply elem_of_equiv; intro; apply elem_of_remove_dups_fast.
destruct X as [l]. apply set_equiv; intro; apply elem_of_remove_dups_fast.
Qed.
End remove_duplicates.
......@@ -22,12 +22,12 @@ Definition htail {A As} (xs : hlist (tcons A As)) : hlist As :=
Fixpoint hheads {As Bs} : hlist (tapp As Bs) hlist As :=
match As with
| tnil => λ _, hnil
| tcons A As => λ xs, hcons (hhead xs) (hheads (htail xs))
| tcons _ _ => λ xs, hcons (hhead xs) (hheads (htail xs))
end.
Fixpoint htails {As Bs} : hlist (tapp As Bs) hlist Bs :=
match As with
| tnil => id
| tcons A As => λ xs, htails (htail xs)
| tcons _ _ => λ xs, htails (htail xs)
end.
Fixpoint himpl (As : tlist) (B : Type) : Type :=
......@@ -35,23 +35,24 @@ Fixpoint himpl (As : tlist) (B : Type) : Type :=
Definition hinit {B} (y : B) : himpl tnil B := y.
Definition hlam {A As B} (f : A himpl As B) : himpl (tcons A As) B := f.
Arguments hlam _ _ _ _ _ / : assert.
Global Arguments hlam _ _ _ _ _ / : assert.
Definition hcurry {As B} (f : himpl As B) (xs : hlist As) : B :=
Definition huncurry {As B} (f : himpl As B) (xs : hlist As) : B :=
(fix go {As} xs :=
match xs in hlist As return himpl As B B with
| hnil => λ f, f
| hcons x xs => λ f, go xs (f x)
end) _ xs f.
Coercion hcurry : himpl >-> Funclass.
Coercion huncurry : himpl >-> Funclass.
Fixpoint huncurry {As B} : (hlist As B) himpl As B :=
Fixpoint hcurry {As B} : (hlist As B) himpl As B :=
match As with
| tnil => λ f, f hnil
| tcons x xs => λ f, hlam (λ x, huncurry (f hcons x))
| tcons x xs => λ f, hlam (λ x, hcurry (f hcons x))
end.
Lemma hcurry_uncurry {As B} (f : hlist As B) xs : huncurry f xs = f xs.
Lemma huncurry_curry {As B} (f : hlist As B) xs :
huncurry (hcurry f) xs = f xs.
Proof. by induction xs as [|A As x xs IH]; simpl; rewrite ?IH. Qed.
Fixpoint hcompose {As B C} (f : B C) {struct As} : himpl As B himpl As C :=
......
......@@ -37,7 +37,7 @@ Section search_infinite.
Context `{!Inj (=) (=) f, !EqDecision B}.
Lemma search_infinite_R_wf xs : wf (R xs).
Lemma search_infinite_R_wf xs : well_founded (R xs).
Proof.
revert xs. assert (help : xs n n',
Acc (R (filter (. f n') xs)) n n' < n Acc (R xs) n).
......@@ -45,7 +45,7 @@ Section search_infinite.
split; [done|]. apply elem_of_list_filter; naive_solver lia. }
intros xs. induction (well_founded_ltof _ length xs) as [xs _ IH].
intros n1; constructor; intros n2 [Hn Hs].
apply help with (n2 - 1); [|lia]. apply IH. eapply filter_length_lt; eauto.
apply help with (n2 - 1); [|lia]. apply IH. eapply length_filter_lt; eauto.
Qed.
Definition search_infinite_go (xs : list B) (n : nat)
......@@ -106,47 +106,47 @@ Section max_infinite.
End max_infinite.
(** Instances for number types *)
Program Instance nat_infinite : Infinite nat :=
Global Program Instance nat_infinite : Infinite nat :=
max_infinite (Nat.max S) 0 (<) _ _ _ _.
Solve Obligations with (intros; simpl; lia).
Program Instance N_infinite : Infinite N :=
Global Program Instance N_infinite : Infinite N :=
max_infinite (N.max N.succ) 0%N N.lt _ _ _ _.
Solve Obligations with (intros; simpl; lia).
Program Instance positive_infinite : Infinite positive :=
Global Program Instance positive_infinite : Infinite positive :=
max_infinite (Pos.max Pos.succ) 1%positive Pos.lt _ _ _ _.
Solve Obligations with (intros; simpl; lia).
Program Instance Z_infinite: Infinite Z :=
Global Program Instance Z_infinite: Infinite Z :=
max_infinite (Z.max Z.succ) 0%Z Z.lt _ _ _ _.
Solve Obligations with (intros; simpl; lia).
(** Instances for option, sum, products *)
Instance option_infinite `{Infinite A} : Infinite (option A) :=
Global Instance option_infinite `{Infinite A} : Infinite (option A) :=
inj_infinite Some id (λ _, eq_refl).
Instance sum_infinite_l `{Infinite A} {B} : Infinite (A + B) :=
Global Instance sum_infinite_l `{Infinite A} {B} : Infinite (A + B) :=
inj_infinite inl (maybe inl) (λ _, eq_refl).
Instance sum_infinite_r {A} `{Infinite B} : Infinite (A + B) :=
Global Instance sum_infinite_r {A} `{Infinite B} : Infinite (A + B) :=
inj_infinite inr (maybe inr) (λ _, eq_refl).
Instance prod_infinite_l `{Infinite A, Inhabited B} : Infinite (A * B) :=
Global Instance prod_infinite_l `{Infinite A, Inhabited B} : Infinite (A * B) :=
inj_infinite (., inhabitant) (Some fst) (λ _, eq_refl).
Instance prod_infinite_r `{Inhabited A, Infinite B} : Infinite (A * B) :=
Global Instance prod_infinite_r `{Inhabited A, Infinite B} : Infinite (A * B) :=
inj_infinite (inhabitant ,.) (Some snd) (λ _, eq_refl).
(** Instance for lists *)
Program Instance list_infinite `{Inhabited A} : Infinite (list A) := {|
Global Program Instance list_infinite `{Inhabited A} : Infinite (list A) := {|
infinite_fresh xxs := replicate (fresh (length <$> xxs)) inhabitant
|}.
Next Obligation.
intros A ? xs ?. destruct (infinite_is_fresh (length <$> xs)).
apply elem_of_list_fmap. eexists; split; [|done].
unfold fresh. by rewrite replicate_length.
unfold fresh. by rewrite length_replicate.
Qed.
Next Obligation. unfold fresh. by intros A ? xs1 xs2 ->. Qed.
(** Instance for strings *)
Program Instance string_infinite : Infinite string :=
Global Program Instance string_infinite : Infinite string :=
search_infinite pretty.
......@@ -10,17 +10,17 @@ Notation cast_trichotomy T :=
| inright _ => inright _
end.
Instance prod_lexico `{Lexico A, Lexico B} : Lexico (A * B) := λ p1 p2,
Global Instance prod_lexico `{Lexico A, Lexico B} : Lexico (A * B) := λ p1 p2,
(**i 1.) *) lexico (p1.1) (p2.1)
(**i 2.) *) p1.1 = p2.1 lexico (p1.2) (p2.2).
Instance bool_lexico : Lexico bool := λ b1 b2,
Global Instance bool_lexico : Lexico bool := λ b1 b2,
match b1, b2 with false, true => True | _, _ => False end.
Instance nat_lexico : Lexico nat := (<).
Instance N_lexico : Lexico N := (<)%N.
Instance Z_lexico : Lexico Z := (<)%Z.
Typeclasses Opaque bool_lexico nat_lexico N_lexico Z_lexico.
Instance list_lexico `{Lexico A} : Lexico (list A) :=
Global Instance nat_lexico : Lexico nat := (<).
Global Instance N_lexico : Lexico N := (<)%N.
Global Instance Z_lexico : Lexico Z := (<)%Z.
Global Typeclasses Opaque bool_lexico nat_lexico N_lexico Z_lexico.
Global Instance list_lexico `{Lexico A} : Lexico (list A) :=
fix go l1 l2 :=
let _ : Lexico (list A) := @go in
match l1, l2 with
......@@ -28,7 +28,7 @@ Instance list_lexico `{Lexico A} : Lexico (list A) :=
| x1 :: l1, x2 :: l2 => lexico (x1,l1) (x2,l2)
| _, _ => False
end.
Instance sig_lexico `{Lexico A} (P : A Prop) `{ x, ProofIrrel (P x)} :
Global Instance sig_lexico `{Lexico A} (P : A Prop) `{ x, ProofIrrel (P x)} :
Lexico (sig P) := λ x1 x2, lexico (`x1) (`x2).
Lemma prod_lexico_irreflexive `{Lexico A, Lexico B, !Irreflexive (@lexico A _)}
......@@ -44,7 +44,7 @@ Proof.
by left; trans x2.
Qed.
Instance prod_lexico_po `{Lexico A, Lexico B, !StrictOrder (@lexico A _)}
Global Instance prod_lexico_po `{Lexico A, Lexico B, !StrictOrder (@lexico A _)}
`{!StrictOrder (@lexico B _)} : StrictOrder (@lexico (A * B) _).
Proof.
split.
......@@ -53,7 +53,7 @@ Proof.
- intros [??] [??] [??] ??.
eapply prod_lexico_transitive; eauto. apply transitivity.
Qed.
Instance prod_lexico_trichotomyT `{Lexico A, tA : !TrichotomyT (@lexico A _)}
Global Instance prod_lexico_trichotomyT `{Lexico A, tA : !TrichotomyT (@lexico A _)}
`{Lexico B, tB : !TrichotomyT (@lexico B _)}: TrichotomyT (@lexico (A * B) _).
Proof.
red; refine (λ p1 p2,
......@@ -65,13 +65,13 @@ Proof.
abstract (unfold lexico, prod_lexico; auto using injective_projections).
Defined.
Instance bool_lexico_po : StrictOrder (@lexico bool _).
Global Instance bool_lexico_po : StrictOrder (@lexico bool _).
Proof.
split.
- by intros [] ?.
- by intros [] [] [] ??.
Qed.
Instance bool_lexico_trichotomy: TrichotomyT (@lexico bool _).
Global Instance bool_lexico_trichotomy: TrichotomyT (@lexico bool _).
Proof.
red; refine (λ b1 b2,
match b1, b2 with
......@@ -82,9 +82,9 @@ Proof.
end); abstract (unfold strict, lexico, bool_lexico; naive_solver).
Defined.
Instance nat_lexico_po : StrictOrder (@lexico nat _).
Global Instance nat_lexico_po : StrictOrder (@lexico nat _).
Proof. unfold lexico, nat_lexico. apply _. Qed.
Instance nat_lexico_trichotomy: TrichotomyT (@lexico nat _).
Global Instance nat_lexico_trichotomy: TrichotomyT (@lexico nat _).
Proof.
red; refine (λ n1 n2,
match Nat.compare n1 n2 as c return Nat.compare n1 n2 = c _ with
......@@ -94,9 +94,9 @@ Proof.
end eq_refl).
Defined.
Instance N_lexico_po : StrictOrder (@lexico N _).
Global Instance N_lexico_po : StrictOrder (@lexico N _).
Proof. unfold lexico, N_lexico. apply _. Qed.
Instance N_lexico_trichotomy: TrichotomyT (@lexico N _).
Global Instance N_lexico_trichotomy: TrichotomyT (@lexico N _).
Proof.
red; refine (λ n1 n2,
match N.compare n1 n2 as c return N.compare n1 n2 = c _ with
......@@ -106,9 +106,9 @@ Proof.
end eq_refl).
Defined.
Instance Z_lexico_po : StrictOrder (@lexico Z _).
Global Instance Z_lexico_po : StrictOrder (@lexico Z _).
Proof. unfold lexico, Z_lexico. apply _. Qed.
Instance Z_lexico_trichotomy: TrichotomyT (@lexico Z _).
Global Instance Z_lexico_trichotomy: TrichotomyT (@lexico Z _).
Proof.
red; refine (λ n1 n2,
match Z.compare n1 n2 as c return Z.compare n1 n2 = c _ with
......@@ -118,7 +118,7 @@ Proof.
end eq_refl).
Defined.
Instance list_lexico_po `{Lexico A, !StrictOrder (@lexico A _)} :
Global Instance list_lexico_po `{Lexico A, !StrictOrder (@lexico A _)} :
StrictOrder (@lexico (list A) _).
Proof.
split.
......@@ -126,7 +126,7 @@ Proof.
- intros l1. induction l1 as [|x1 l1]; intros [|x2 l2] [|x3 l3] ??; try done.
eapply prod_lexico_transitive; eauto.
Qed.
Instance list_lexico_trichotomy `{Lexico A, tA : !TrichotomyT (@lexico A _)} :
Global Instance list_lexico_trichotomy `{Lexico A, tA : !TrichotomyT (@lexico A _)} :
TrichotomyT (@lexico (list A) _).
Proof.
refine (
......@@ -138,17 +138,17 @@ Proof.
| _ :: _, [] => inright _
| x1 :: l1, x2 :: l2 => cast_trichotomy (trichotomyT lexico (x1,l1) (x2,l2))
end); clear tA go go';
abstract (repeat (done || constructor || congruence || by inversion 1)).
abstract (repeat (done || constructor || congruence || by inv 1)).
Defined.
Instance sig_lexico_po `{Lexico A, !StrictOrder (@lexico A _)}
Global Instance sig_lexico_po `{Lexico A, !StrictOrder (@lexico A _)}
(P : A Prop) `{ x, ProofIrrel (P x)} : StrictOrder (@lexico (sig P) _).
Proof.
unfold lexico, sig_lexico. split.
- intros [x ?] ?. by apply (irreflexivity lexico x).
- intros [x1 ?] [x2 ?] [x3 ?] ??. by trans x2.
Qed.
Instance sig_lexico_trichotomy `{Lexico A, tA : !TrichotomyT (@lexico A _)}
Global Instance sig_lexico_trichotomy `{Lexico A, tA : !TrichotomyT (@lexico A _)}
(P : A Prop) `{ x, ProofIrrel (P x)} : TrichotomyT (@lexico (sig P) _).
Proof.
red; refine (λ x1 x2, cast_trichotomy (trichotomyT lexico (`x1) (`x2)));
......
(** This file re-exports all the list lemmas in std++. Do *not* import the individual
[list_*] modules; their organization may cahnge over time. Always import [list]. *)
From stdpp Require Export list_basics list_relations list_monad list_misc list_tactics list_numbers.
From stdpp Require Import options.
From stdpp Require Export numbers base option.
From stdpp Require Import options.
Global Arguments length {_} _ : assert.
Global Arguments cons {_} _ _ : assert.
Global Arguments app {_} _ _ : assert.
Global Instance: Params (@length) 1 := {}.
Global Instance: Params (@cons) 1 := {}.
Global Instance: Params (@app) 1 := {}.
(** [head] and [tail] are defined as [parsing only] for [hd_error] and [tl] in
the Coq standard library. We redefine these notations to make sure they also
pretty print properly. *)
Notation head := hd_error.
Notation tail := tl.
Notation take := firstn.
Notation drop := skipn.
Global Arguments head {_} _ : assert.
Global Arguments tail {_} _ : assert.
Global Arguments take {_} !_ !_ / : assert.
Global Arguments drop {_} !_ !_ / : assert.
Global Instance: Params (@head) 1 := {}.
Global Instance: Params (@tail) 1 := {}.
Global Instance: Params (@take) 1 := {}.
Global Instance: Params (@drop) 1 := {}.
Notation "(::)" := cons (only parsing) : list_scope.
Notation "( x ::.)" := (cons x) (only parsing) : list_scope.
Notation "(.:: l )" := (λ x, cons x l) (only parsing) : list_scope.
Notation "(++)" := app (only parsing) : list_scope.
Notation "( l ++.)" := (app l) (only parsing) : list_scope.
Notation "(.++ k )" := (λ l, app l k) (only parsing) : list_scope.
Global Instance maybe_cons {A} : Maybe2 (@cons A) := λ l,
match l with x :: l => Some (x,l) | _ => None end.
(** The operation [l !! i] gives the [i]th element of the list [l], or [None]
in case [i] is out of bounds. *)
Global Instance list_lookup {A} : Lookup nat A (list A) :=
fix go i l {struct l} : option A := let _ : Lookup _ _ _ := @go in
match l with
| [] => None | x :: l => match i with 0 => Some x | S i => l !! i end
end.
(** The operation [l !!! i] is a total version of the lookup operation
[l !! i]. *)
Global Instance list_lookup_total `{!Inhabited A} : LookupTotal nat A (list A) :=
fix go i l {struct l} : A := let _ : LookupTotal _ _ _ := @go in
match l with
| [] => inhabitant
| x :: l => match i with 0 => x | S i => l !!! i end
end.
(** The operation [alter f i l] applies the function [f] to the [i]th element
of [l]. In case [i] is out of bounds, the list is returned unchanged. *)
Global Instance list_alter {A} : Alter nat A (list A) := λ f,
fix go i l {struct l} :=
match l with
| [] => []
| x :: l => match i with 0 => f x :: l | S i => x :: go i l end
end.
(** The operation [<[i:=x]> l] overwrites the element at position [i] with the
value [x]. In case [i] is out of bounds, the list is returned unchanged. *)
Global Instance list_insert {A} : Insert nat A (list A) :=
fix go i y l {struct l} := let _ : Insert _ _ _ := @go in
match l with
| [] => []
| x :: l => match i with 0 => y :: l | S i => x :: <[i:=y]>l end
end.
Fixpoint list_inserts {A} (i : nat) (k l : list A) : list A :=
match k with
| [] => l
| y :: k => <[i:=y]>(list_inserts (S i) k l)
end.
Global Instance: Params (@list_inserts) 1 := {}.
(** The operation [delete i l] removes the [i]th element of [l] and moves
all consecutive elements one position ahead. In case [i] is out of bounds,
the list is returned unchanged. *)
Global 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.
(** The function [option_list o] converts an element [Some x] into the
singleton list [[x]], and [None] into the empty list [[]]. *)
Definition option_list {A} : option A list A := option_rect _ (λ x, [x]) [].
Global Instance: Params (@option_list) 1 := {}.
Global Instance maybe_list_singleton {A} : Maybe (λ x : A, [x]) := λ l,
match l with [x] => Some x | _ => None end.
(** The function [filter P l] returns the list of elements of [l] that
satisfies [P]. The order remains unchanged. *)
Global Instance list_filter {A} : Filter A (list A) :=
fix go P _ l := let _ : Filter _ _ := @go in
match l with
| [] => []
| x :: l => if decide (P x) then x :: filter P l else filter P 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.
Global Instance: Params (@replicate) 2 := {}.
(** The function [reverse l] returns the elements of [l] in reverse order. *)
Definition reverse {A} (l : list A) : list A := rev_append l [].
Global Instance: Params (@reverse) 1 := {}.
(** The function [last l] returns the last element of the list [l], or [None]
if the list [l] is empty. *)
Fixpoint last {A} (l : list A) : option A :=
match l with [] => None | [x] => Some x | _ :: l => last l end.
Global Instance: Params (@last) 1 := {}.
Global Arguments last : simpl nomatch.
(** Functions to fold over a list. We redefine [foldl] with the arguments in
the same order as in Haskell. *)
Notation foldr := fold_right.
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.
(** Set operations on lists *)
Section list_set.
Context `{dec : EqDecision A}.
Global Instance elem_of_list_dec : RelDecision (∈@{list A}).
Proof using Type*.
refine (
fix go x l :=
match l return Decision (x l) with
| [] => right _
| y :: l => cast_if_or (decide (x = y)) (go x l)
end); clear go dec; subst; try (by constructor); abstract by inv 1.
Defined.
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.
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.
Definition list_union (l k : list A) : list A := list_difference l k ++ k.
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.
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.
End list_set.
(** * General theorems *)
Section general_properties.
Context {A : Type}.
Implicit Types x y z : A.
Implicit Types l k : list A.
(* TODO: Coq 8.20 has the same lemma under the same name, so remove our version
once we require Coq 8.20. In Coq 8.19 and before, this lemma is called
[app_length]. *)
Lemma length_app (l l' : list A) : length (l ++ l') = length l + length l'.
Proof. induction l; f_equal/=; auto. Qed.
Lemma app_inj_1 (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 app_inj_2 (l1 k1 l2 k2 : list A) :
length l2 = length k2 l1 ++ l2 = k1 ++ k2 l1 = k1 l2 = k2.
Proof.
intros ? Hl. apply app_inj_1; auto.
apply (f_equal length) in Hl. rewrite !length_app in Hl. lia.
Qed.
Global Instance cons_eq_inj : Inj2 (=) (=) (=) (@cons A).
Proof. by injection 1. Qed.
Global Instance: k, Inj (=) (=) (k ++.).
Proof. intros ???. apply app_inv_head. Qed.
Global Instance: k, Inj (=) (=) (.++ k).
Proof. intros ???. apply app_inv_tail. Qed.
Global Instance: Assoc (=) (@app A).
Proof. intros ???. apply app_assoc. Qed.
Global Instance: LeftId (=) [] (@app A).
Proof. done. Qed.
Global Instance: RightId (=) [] (@app A).
Proof. intro. apply app_nil_r. Qed.
Lemma app_nil l1 l2 : l1 ++ l2 = [] l1 = [] l2 = [].
Proof. split; [apply app_eq_nil|]. by intros [-> ->]. Qed.
Lemma app_singleton l1 l2 x :
l1 ++ l2 = [x] l1 = [] l2 = [x] l1 = [x] l2 = [].
Proof. split; [apply app_eq_unit|]. by intros [[-> ->]|[-> ->]]. Qed.
Lemma cons_middle x l1 l2 : l1 ++ x :: l2 = l1 ++ [x] ++ l2.
Proof. done. Qed.
Lemma list_eq l1 l2 : ( i, l1 !! i = l2 !! i) l1 = l2.
Proof.
revert l2. induction l1 as [|x l1 IH]; intros [|y l2] H.
- done.
- discriminate (H 0).
- discriminate (H 0).
- f_equal; [by injection (H 0)|]. apply (IH _ $ λ i, H (S i)).
Qed.
Global Instance list_eq_dec {dec : EqDecision A} : EqDecision (list A) :=
list_eq_dec dec.
Global Instance list_eq_nil_dec l : Decision (l = []).
Proof. by refine match l with [] => left _ | _ => right _ end. Defined.
Lemma list_singleton_reflect l :
option_reflect (λ x, l = [x]) (length l 1) (maybe (λ x, [x]) l).
Proof. by destruct l as [|? []]; constructor. Defined.
Lemma list_eq_Forall2 l1 l2 : l1 = l2 Forall2 eq l1 l2.
Proof.
split.
- intros <-. induction l1; eauto using Forall2.
- induction 1; naive_solver.
Qed.
Definition length_nil : length (@nil A) = 0 := eq_refl.
Definition length_cons x l : length (x :: l) = S (length l) := eq_refl.
Lemma nil_or_length_pos l : l = [] length l 0.
Proof. destruct l; simpl; auto with lia. Qed.
Lemma nil_length_inv l : length l = 0 l = [].
Proof. by destruct l. Qed.
Lemma lookup_cons_ne_0 l x i : i 0 (x :: l) !! i = l !! pred i.
Proof. by destruct i. Qed.
Lemma lookup_total_cons_ne_0 `{!Inhabited A} l x i :
i 0 (x :: l) !!! i = l !!! pred i.
Proof. by destruct i. Qed.
Lemma lookup_nil i : @nil A !! i = None.
Proof. by destruct i. Qed.
Lemma lookup_total_nil `{!Inhabited A} i : @nil A !!! i = inhabitant.
Proof. by destruct i. Qed.
Lemma lookup_tail l i : tail l !! i = l !! S i.
Proof. by destruct l. Qed.
Lemma lookup_total_tail `{!Inhabited A} l i : tail l !!! i = l !!! S i.
Proof. by destruct l. Qed.
Lemma lookup_lt_Some l i x : l !! i = Some x i < length l.
Proof. revert i. induction l; intros [|?] ?; naive_solver auto with arith. Qed.
Lemma lookup_lt_is_Some_1 l i : is_Some (l !! i) i < length l.
Proof. intros [??]; eauto using lookup_lt_Some. Qed.
Lemma lookup_lt_is_Some_2 l i : i < length l is_Some (l !! i).
Proof. revert i. induction l; intros [|?] ?; naive_solver auto with lia. Qed.
Lemma lookup_lt_is_Some l i : is_Some (l !! i) i < length l.
Proof. split; auto using lookup_lt_is_Some_1, lookup_lt_is_Some_2. Qed.
Lemma lookup_ge_None l i : l !! i = None length l i.
Proof. rewrite eq_None_not_Some, lookup_lt_is_Some. lia. Qed.
Lemma lookup_ge_None_1 l i : l !! i = None length l i.
Proof. by rewrite lookup_ge_None. Qed.
Lemma lookup_ge_None_2 l i : length l i l !! i = None.
Proof. by rewrite lookup_ge_None. Qed.
Lemma list_eq_same_length l1 l2 n :
length l2 = n length l1 = n
( i x y, i < n l1 !! i = Some x l2 !! i = Some y x = y) l1 = l2.
Proof.
intros <- Hlen Hl; apply list_eq; intros i. destruct (l2 !! i) as [x|] eqn:Hx.
- destruct (lookup_lt_is_Some_2 l1 i) as [y Hy].
{ rewrite Hlen; eauto using lookup_lt_Some. }
rewrite Hy; f_equal; apply (Hl i); eauto using lookup_lt_Some.
- by rewrite lookup_ge_None, Hlen, <-lookup_ge_None.
Qed.
Lemma nth_lookup l i d : nth i l d = default d (l !! i).
Proof. revert i. induction l as [|x l IH]; intros [|i]; simpl; auto. Qed.
Lemma nth_lookup_Some l i d x : l !! i = Some x nth i l d = x.
Proof. rewrite nth_lookup. by intros ->. Qed.
Lemma nth_lookup_or_length l i d : {l !! i = Some (nth i l d)} + {length l i}.
Proof.
rewrite nth_lookup. destruct (l !! i) eqn:?; eauto using lookup_ge_None_1.
Qed.
Lemma list_lookup_total_alt `{!Inhabited A} l i :
l !!! i = default inhabitant (l !! i).
Proof. revert i. induction l; intros []; naive_solver. Qed.
Lemma list_lookup_total_correct `{!Inhabited A} l i x :
l !! i = Some x l !!! i = x.
Proof. rewrite list_lookup_total_alt. by intros ->. Qed.
Lemma list_lookup_lookup_total `{!Inhabited A} l i :
is_Some (l !! i) l !! i = Some (l !!! i).
Proof. rewrite list_lookup_total_alt; by intros [x ->]. Qed.
Lemma list_lookup_lookup_total_lt `{!Inhabited A} l i :
i < length l l !! i = Some (l !!! i).
Proof. intros ?. by apply list_lookup_lookup_total, lookup_lt_is_Some_2. Qed.
Lemma list_lookup_alt `{!Inhabited A} l i x :
l !! i = Some x i < length l l !!! i = x.
Proof.
naive_solver eauto using list_lookup_lookup_total_lt,
list_lookup_total_correct, lookup_lt_Some.
Qed.
Lemma lookup_app l1 l2 i :
(l1 ++ l2) !! i =
match l1 !! i with Some x => Some x | None => l2 !! (i - length l1) end.
Proof. revert i. induction l1 as [|x l1 IH]; intros [|i]; naive_solver. Qed.
Lemma lookup_app_l l1 l2 i : i < length l1 (l1 ++ l2) !! i = l1 !! i.
Proof. rewrite lookup_app. by intros [? ->]%lookup_lt_is_Some. Qed.
Lemma lookup_total_app_l `{!Inhabited A} l1 l2 i :
i < length l1 (l1 ++ l2) !!! i = l1 !!! i.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_app_l. Qed.
Lemma lookup_app_l_Some l1 l2 i x : l1 !! i = Some x (l1 ++ l2) !! i = Some x.
Proof. rewrite lookup_app. by intros ->. Qed.
Lemma lookup_app_r l1 l2 i :
length l1 i (l1 ++ l2) !! i = l2 !! (i - length l1).
Proof. rewrite lookup_app. by intros ->%lookup_ge_None. Qed.
Lemma lookup_total_app_r `{!Inhabited A} l1 l2 i :
length l1 i (l1 ++ l2) !!! i = l2 !!! (i - length l1).
Proof. intros. by rewrite !list_lookup_total_alt, lookup_app_r. Qed.
Lemma lookup_app_Some l1 l2 i x :
(l1 ++ l2) !! i = Some x
l1 !! i = Some x length l1 i l2 !! (i - length l1) = Some x.
Proof.
rewrite lookup_app. destruct (l1 !! i) eqn:Hi.
- apply lookup_lt_Some in Hi. naive_solver lia.
- apply lookup_ge_None in Hi. naive_solver lia.
Qed.
Lemma lookup_cons x l i :
(x :: l) !! i =
match i with 0 => Some x | S i => l !! i end.
Proof. reflexivity. Qed.
Lemma lookup_cons_Some x l i y :
(x :: l) !! i = Some y
(i = 0 x = y) (1 i l !! (i - 1) = Some y).
Proof.
rewrite lookup_cons. destruct i as [|i].
- naive_solver lia.
- replace (S i - 1) with i by lia. naive_solver lia.
Qed.
Lemma list_lookup_singleton x i :
[x] !! i =
match i with 0 => Some x | S _ => None end.
Proof. reflexivity. Qed.
Lemma list_lookup_singleton_Some x i y :
[x] !! i = Some y i = 0 x = y.
Proof. rewrite lookup_cons_Some. naive_solver. Qed.
Lemma lookup_snoc_Some x l i y :
(l ++ [x]) !! i = Some y
(i < length l l !! i = Some y) (i = length l x = y).
Proof.
rewrite lookup_app_Some, list_lookup_singleton_Some.
naive_solver auto using lookup_lt_is_Some_1 with lia.
Qed.
Lemma list_lookup_middle l1 l2 x n :
n = length l1 (l1 ++ x :: l2) !! n = Some x.
Proof. intros ->. by induction l1. Qed.
Lemma list_lookup_total_middle `{!Inhabited A} l1 l2 x n :
n = length l1 (l1 ++ x :: l2) !!! n = x.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_middle. Qed.
Lemma list_insert_alter l i x : <[i:=x]>l = alter (λ _, x) i l.
Proof. by revert i; induction l; intros []; intros; f_equal/=. Qed.
Lemma length_alter f l i : length (alter f i l) = length l.
Proof. revert i. by induction l; intros [|?]; f_equal/=. Qed.
Lemma length_insert l i x : length (<[i:=x]>l) = length l.
Proof. revert i. by induction l; intros [|?]; f_equal/=. Qed.
Lemma list_lookup_alter f l i : alter f i l !! i = f <$> l !! i.
Proof.
revert i.
induction l as [|?? IHl]; [done|].
intros [|i]; [done|]. apply (IHl i).
Qed.
Lemma list_lookup_total_alter `{!Inhabited A} f l i :
i < length l alter f i l !!! i = f (l !!! i).
Proof.
intros [x Hx]%lookup_lt_is_Some_2.
by rewrite !list_lookup_total_alt, list_lookup_alter, Hx.
Qed.
Lemma list_lookup_alter_ne f l i j : i j alter f i l !! j = l !! j.
Proof. revert i j. induction l; [done|]. intros [] []; naive_solver. Qed.
Lemma list_lookup_total_alter_ne `{!Inhabited A} f l i j :
i j alter f i l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_alter_ne. Qed.
Lemma list_lookup_insert l i x : i < length l <[i:=x]>l !! i = Some x.
Proof. revert i. induction l; intros [|?] ?; f_equal/=; auto with lia. Qed.
Lemma list_lookup_total_insert `{!Inhabited A} l i x :
i < length l <[i:=x]>l !!! i = x.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_insert. Qed.
Lemma list_lookup_insert_ne l i j x : i j <[i:=x]>l !! j = l !! j.
Proof. revert i j. induction l; [done|]. intros [] []; naive_solver. Qed.
Lemma list_lookup_total_insert_ne `{!Inhabited A} l i j x :
i j <[i:=x]>l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_insert_ne. Qed.
Lemma list_lookup_insert_Some l i x j y :
<[i:=x]>l !! j = Some y
i = j x = y j < length l i j l !! j = Some y.
Proof.
destruct (decide (i = j)) as [->|];
[split|rewrite list_lookup_insert_ne by done; tauto].
- intros Hy. assert (j < length l).
{ rewrite <-(length_insert l j x); eauto using lookup_lt_Some. }
rewrite list_lookup_insert in Hy by done; naive_solver.
- intros [(?&?&?)|[??]]; rewrite ?list_lookup_insert; naive_solver.
Qed.
Lemma list_insert_commute l i j x y :
i j <[i:=x]>(<[j:=y]>l) = <[j:=y]>(<[i:=x]>l).
Proof. revert i j. by induction l; intros [|?] [|?] ?; f_equal/=; auto. Qed.
Lemma list_insert_id' l i x : (i < length l l !! i = Some x) <[i:=x]>l = l.
Proof. revert i. induction l; intros [|i] ?; f_equal/=; naive_solver lia. Qed.
Lemma list_insert_id l i x : l !! i = Some x <[i:=x]>l = l.
Proof. intros ?. by apply list_insert_id'. Qed.
Lemma list_insert_ge l i x : length l i <[i:=x]>l = l.
Proof. revert i. induction l; intros [|i] ?; f_equal/=; auto with lia. Qed.
Lemma list_insert_insert l i x y : <[i:=x]> (<[i:=y]> l) = <[i:=x]> l.
Proof. revert i. induction l; intros [|i]; f_equal/=; auto. Qed.
Lemma list_lookup_other l i x :
length l 1 l !! i = Some x j y, j i l !! j = Some y.
Proof.
intros. destruct i, l as [|x0 [|x1 l]]; simplify_eq/=.
- by exists 1, x1.
- by exists 0, x0.
Qed.
Lemma alter_app_l f l1 l2 i :
i < length l1 alter f i (l1 ++ l2) = alter f i l1 ++ l2.
Proof. revert i. induction l1; intros [|?] ?; f_equal/=; auto with lia. Qed.
Lemma alter_app_r f l1 l2 i :
alter f (length l1 + i) (l1 ++ l2) = l1 ++ alter f i l2.
Proof. revert i. induction l1; intros [|?]; f_equal/=; auto. Qed.
Lemma alter_app_r_alt f l1 l2 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 list_alter_id f l i : ( x, f x = x) alter f i l = l.
Proof. intros ?. revert i. induction l; intros [|?]; f_equal/=; auto. Qed.
Lemma list_alter_ext f g l k i :
( x, l !! i = Some x f x = g x) l = k alter f i l = alter g i k.
Proof. intros H ->. revert i H. induction k; intros [|?] ?; f_equal/=; auto. Qed.
Lemma list_alter_compose f g l i :
alter (f g) i l = alter f i (alter g i l).
Proof. revert i. induction l; intros [|?]; f_equal/=; auto. Qed.
Lemma list_alter_commute f g l i j :
i j alter f i (alter g j l) = alter g j (alter f i l).
Proof. revert i j. induction l; intros [|?][|?] ?; f_equal/=; auto with lia. Qed.
Lemma insert_app_l l1 l2 i x :
i < length l1 <[i:=x]>(l1 ++ l2) = <[i:=x]>l1 ++ l2.
Proof. revert i. induction l1; intros [|?] ?; f_equal/=; auto with lia. Qed.
Lemma insert_app_r l1 l2 i x : <[length l1+i:=x]>(l1 ++ l2) = l1 ++ <[i:=x]>l2.
Proof. revert i. induction l1; intros [|?]; f_equal/=; auto. Qed.
Lemma insert_app_r_alt l1 l2 i x :
length l1 i <[i:=x]>(l1 ++ l2) = l1 ++ <[i - length l1:=x]>l2.
Proof.
intros. assert (i = length l1 + (i - length l1)) as Hi by lia.
rewrite Hi at 1. by apply insert_app_r.
Qed.
Lemma delete_middle l1 l2 x : delete (length l1) (l1 ++ x :: l2) = l1 ++ l2.
Proof. induction l1; f_equal/=; auto. Qed.
Lemma length_delete l i :
is_Some (l !! i) length (delete i l) = length l - 1.
Proof.
rewrite lookup_lt_is_Some. revert i.
induction l as [|x l IH]; intros [|i] ?; simpl in *; [lia..|].
rewrite IH by lia. lia.
Qed.
Lemma lookup_delete_lt l i j : j < i delete i l !! j = l !! j.
Proof. revert i j; induction l; intros [] []; naive_solver eauto with lia. Qed.
Lemma lookup_total_delete_lt `{!Inhabited A} l i j :
j < i delete i l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_delete_lt. Qed.
Lemma lookup_delete_ge l i j : i j delete i l !! j = l !! S j.
Proof. revert i j; induction l; intros [] []; naive_solver eauto with lia. Qed.
Lemma lookup_total_delete_ge `{!Inhabited A} l i j :
i j delete i l !!! j = l !!! S j.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_delete_ge. Qed.
Lemma length_inserts l i k : length (list_inserts i k l) = length l.
Proof.
revert i. induction k; intros ?; csimpl; rewrite ?length_insert; auto.
Qed.
Lemma list_lookup_inserts l i k j :
i j < i + length k j < length l
list_inserts i k l !! j = k !! (j - i).
Proof.
revert i j. induction k as [|y k IH]; csimpl; intros i j ??; [lia|].
destruct (decide (i = j)) as [->|].
{ by rewrite list_lookup_insert, Nat.sub_diag
by (rewrite length_inserts; lia). }
rewrite list_lookup_insert_ne, IH by lia.
by replace (j - i) with (S (j - S i)) by lia.
Qed.
Lemma list_lookup_total_inserts `{!Inhabited A} l i k j :
i j < i + length k j < length l
list_inserts i k l !!! j = k !!! (j - i).
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts. Qed.
Lemma list_lookup_inserts_lt l i k j :
j < i list_inserts i k l !! j = l !! j.
Proof.
revert i j. induction k; intros i j ?; csimpl;
rewrite ?list_lookup_insert_ne by lia; auto with lia.
Qed.
Lemma list_lookup_total_inserts_lt `{!Inhabited A}l i k j :
j < i list_inserts i k l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts_lt. Qed.
Lemma list_lookup_inserts_ge l i k j :
i + length k j list_inserts i k l !! j = l !! j.
Proof.
revert i j. induction k; csimpl; intros i j ?;
rewrite ?list_lookup_insert_ne by lia; auto with lia.
Qed.
Lemma list_lookup_total_inserts_ge `{!Inhabited A} l i k j :
i + length k j list_inserts i k l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts_ge. Qed.
Lemma list_lookup_inserts_Some l i k j y :
list_inserts i k l !! j = Some y
(j < i i + length k j) l !! j = Some y
i j < i + length k j < length l k !! (j - i) = Some y.
Proof.
destruct (decide (j < i)).
{ rewrite list_lookup_inserts_lt by done; intuition lia. }
destruct (decide (i + length k j)).
{ rewrite list_lookup_inserts_ge by done; intuition lia. }
split.
- intros Hy. assert (j < length l).
{ rewrite <-(length_inserts l i k); eauto using lookup_lt_Some. }
rewrite list_lookup_inserts in Hy by lia. intuition lia.
- intuition. by rewrite list_lookup_inserts by lia.
Qed.
Lemma list_insert_inserts_lt l i j x k :
i < j <[i:=x]>(list_inserts j k l) = list_inserts j k (<[i:=x]>l).
Proof.
revert i j. induction k; intros i j ?; simpl;
rewrite 1?list_insert_commute by lia; auto with f_equal.
Qed.
Lemma list_inserts_app_l l1 l2 l3 i :
list_inserts i (l1 ++ l2) l3 = list_inserts (length l1 + i) l2 (list_inserts i l1 l3).
Proof.
revert i; induction l1 as [|x l1 IH]; [done|].
intro i. simpl. rewrite IH, Nat.add_succ_r. apply list_insert_inserts_lt. lia.
Qed.
Lemma list_inserts_app_r l1 l2 l3 i :
list_inserts (length l2 + i) l1 (l2 ++ l3) = l2 ++ list_inserts i l1 l3.
Proof.
revert i; induction l1 as [|x l1 IH]; [done|].
intros i. simpl. by rewrite plus_n_Sm, IH, insert_app_r.
Qed.
Lemma list_inserts_nil l1 i : list_inserts i l1 [] = [].
Proof.
revert i; induction l1 as [|x l1 IH]; [done|].
intro i. simpl. by rewrite IH.
Qed.
Lemma list_inserts_cons l1 l2 i x :
list_inserts (S i) l1 (x :: l2) = x :: list_inserts i l1 l2.
Proof.
revert i; induction l1 as [|y l1 IH]; [done|].
intro i. simpl. by rewrite IH.
Qed.
Lemma list_inserts_0_r l1 l2 l3 :
length l1 = length l2 list_inserts 0 l1 (l2 ++ l3) = l1 ++ l3.
Proof.
revert l2. induction l1 as [|x l1 IH]; intros [|y l2] ?; simplify_eq/=; [done|].
rewrite list_inserts_cons. simpl. by rewrite IH.
Qed.
Lemma list_inserts_0_l l1 l2 l3 :
length l1 = length l3 list_inserts 0 (l1 ++ l2) l3 = l1.
Proof.
revert l3. induction l1 as [|x l1 IH]; intros [|z l3] ?; simplify_eq/=.
{ by rewrite list_inserts_nil. }
rewrite list_inserts_cons. simpl. by rewrite IH.
Qed.
(** ** Properties of the [reverse] function *)
Lemma reverse_nil : reverse [] =@{list A} [].
Proof. done. Qed.
Lemma reverse_singleton x : reverse [x] = [x].
Proof. done. Qed.
Lemma reverse_cons l x : reverse (x :: l) = reverse l ++ [x].
Proof. unfold reverse. by rewrite <-!rev_alt. Qed.
Lemma reverse_snoc l x : reverse (l ++ [x]) = x :: reverse l.
Proof. unfold reverse. by rewrite <-!rev_alt, rev_unit. Qed.
Lemma reverse_app l1 l2 : reverse (l1 ++ l2) = reverse l2 ++ reverse l1.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_app_distr. Qed.
Lemma length_reverse l : length (reverse l) = length l.
Proof.
induction l as [|x l IH]; [done|].
rewrite reverse_cons, length_app, IH. simpl. lia.
Qed.
Lemma reverse_involutive l : reverse (reverse l) = l.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_involutive. Qed.
Lemma reverse_lookup l i :
i < length l
reverse l !! i = l !! (length l - S i).
Proof.
revert i. induction l as [|x l IH]; simpl; intros i Hi; [done|].
rewrite reverse_cons.
destruct (decide (i = length l)); subst.
+ by rewrite list_lookup_middle, Nat.sub_diag by by rewrite length_reverse.
+ rewrite lookup_app_l by (rewrite length_reverse; lia).
rewrite IH by lia.
by assert (length l - i = S (length l - S i)) as -> by lia.
Qed.
Lemma reverse_lookup_Some l i x :
reverse l !! i = Some x l !! (length l - S i) = Some x i < length l.
Proof.
split.
- destruct (decide (i < length l)); [ by rewrite reverse_lookup|].
rewrite lookup_ge_None_2; [done|]. rewrite length_reverse. lia.
- intros [??]. by rewrite reverse_lookup.
Qed.
Global Instance: Inj (=) (=) (@reverse A).
Proof.
intros l1 l2 Hl.
by rewrite <-(reverse_involutive l1), <-(reverse_involutive l2), Hl.
Qed.
(** ** Properties of the [elem_of] predicate *)
Lemma not_elem_of_nil x : x [].
Proof. by inv 1. Qed.
Lemma elem_of_nil x : x [] False.
Proof. intuition. by destruct (not_elem_of_nil x). Qed.
Lemma elem_of_nil_inv l : ( x, x l) l = [].
Proof. destruct l; [done|]. by edestruct 1; constructor. Qed.
Lemma elem_of_not_nil x l : x l l [].
Proof. intros ? ->. by apply (elem_of_nil x). Qed.
Lemma elem_of_cons l x y : x y :: l x = y x l.
Proof. by split; [inv 1; subst|intros [->|?]]; constructor. Qed.
Lemma not_elem_of_cons l x y : x y :: l x y x l.
Proof. rewrite elem_of_cons. tauto. Qed.
Lemma elem_of_app l1 l2 x : x l1 ++ l2 x l1 x l2.
Proof.
induction l1 as [|y l1 IH]; simpl.
- rewrite elem_of_nil. naive_solver.
- rewrite !elem_of_cons, IH. naive_solver.
Qed.
Lemma not_elem_of_app l1 l2 x : x l1 ++ l2 x l1 x l2.
Proof. rewrite elem_of_app. tauto. Qed.
Lemma elem_of_list_singleton x y : x [y] x = y.
Proof. rewrite elem_of_cons, elem_of_nil. tauto. Qed.
Lemma elem_of_reverse_2 x l : x l x reverse l.
Proof.
induction 1; rewrite reverse_cons, elem_of_app,
?elem_of_list_singleton; intuition.
Qed.
Lemma elem_of_reverse x l : x reverse l x l.
Proof.
split; auto using elem_of_reverse_2.
intros. rewrite <-(reverse_involutive l). by apply elem_of_reverse_2.
Qed.
Lemma elem_of_list_lookup_1 l 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_total_1 `{!Inhabited A} l x :
x l i, i < length l l !!! i = x.
Proof.
intros [i Hi]%elem_of_list_lookup_1.
eauto using lookup_lt_Some, list_lookup_total_correct.
Qed.
Lemma elem_of_list_lookup_2 l i x : l !! i = Some x x l.
Proof.
revert i. induction l; intros [|i] ?; simplify_eq/=; constructor; eauto.
Qed.
Lemma elem_of_list_lookup_total_2 `{!Inhabited A} l i :
i < length l l !!! i l.
Proof. intros. by eapply elem_of_list_lookup_2, list_lookup_lookup_total_lt. Qed.
Lemma elem_of_list_lookup l x : x l i, l !! i = Some x.
Proof. firstorder eauto using elem_of_list_lookup_1, elem_of_list_lookup_2. Qed.
Lemma elem_of_list_lookup_total `{!Inhabited A} l x :
x l i, i < length l l !!! i = x.
Proof.
naive_solver eauto using elem_of_list_lookup_total_1, elem_of_list_lookup_total_2.
Qed.
Lemma elem_of_list_split_length l i x :
l !! i = Some x l1 l2, l = l1 ++ x :: l2 i = length l1.
Proof.
revert i; induction l as [|y l IH]; intros [|i] Hl; simplify_eq/=.
- exists []; eauto.
- destruct (IH _ Hl) as (?&?&?&?); simplify_eq/=.
eexists (y :: _); eauto.
Qed.
Lemma elem_of_list_split l x : x l l1 l2, l = l1 ++ x :: l2.
Proof.
intros [? (?&?&?&_)%elem_of_list_split_length]%elem_of_list_lookup_1; eauto.
Qed.
Lemma elem_of_list_split_l `{EqDecision A} l x :
x l l1 l2, l = l1 ++ x :: l2 x l1.
Proof.
induction 1 as [x l|x y l ? IH].
{ exists [], l. rewrite elem_of_nil. naive_solver. }
destruct (decide (x = y)) as [->|?].
- exists [], l. rewrite elem_of_nil. naive_solver.
- destruct IH as (l1 & l2 & -> & ?).
exists (y :: l1), l2. rewrite elem_of_cons. naive_solver.
Qed.
Lemma elem_of_list_split_r `{EqDecision A} l x :
x l l1 l2, l = l1 ++ x :: l2 x l2.
Proof.
induction l as [|y l IH] using rev_ind.
{ by rewrite elem_of_nil. }
destruct (decide (x = y)) as [->|].
- exists l, []. rewrite elem_of_nil. naive_solver.
- rewrite elem_of_app, elem_of_list_singleton. intros [?| ->]; try done.
destruct IH as (l1 & l2 & -> & ?); auto.
exists l1, (l2 ++ [y]).
rewrite elem_of_app, elem_of_list_singleton, <-(assoc_L (++)). naive_solver.
Qed.
Lemma list_elem_of_insert l i x : i < length l x <[i:=x]>l.
Proof. intros. by eapply elem_of_list_lookup_2, list_lookup_insert. Qed.
Lemma nth_elem_of l i d : i < length l nth i l d l.
Proof.
intros; eapply elem_of_list_lookup_2.
destruct (nth_lookup_or_length l i d); [done | by lia].
Qed.
Lemma not_elem_of_app_cons_inv_l x y l1 l2 k1 k2 :
x k1 y l1
l1 ++ x :: l2 = k1 ++ y :: k2
l1 = k1 x = y l2 = k2.
Proof.
revert k1. induction l1 as [|x' l1 IH]; intros [|y' k1] Hx Hy ?; simplify_eq/=;
try apply not_elem_of_cons in Hx as [??];
try apply not_elem_of_cons in Hy as [??]; naive_solver.
Qed.
Lemma not_elem_of_app_cons_inv_r x y l1 l2 k1 k2 :
x k2 y l2
l1 ++ x :: l2 = k1 ++ y :: k2
l1 = k1 x = y l2 = k2.
Proof.
intros. destruct (not_elem_of_app_cons_inv_l x y (reverse l2) (reverse l1)
(reverse k2) (reverse k1)); [..|naive_solver].
- by rewrite elem_of_reverse.
- by rewrite elem_of_reverse.
- rewrite <-!reverse_snoc, <-!reverse_app, <-!(assoc_L (++)). by f_equal.
Qed.
(** ** Set operations on lists *)
Section list_set.
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 l|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.
Context `{!EqDecision A}.
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 elem_of_list_union l k x : x list_union l k x l x k.
Proof.
unfold list_union. rewrite elem_of_app, elem_of_list_difference.
intuition. case (decide (x k)); intuition.
Qed.
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.
End list_set.
(** ** Properties of the [last] function *)
Lemma last_nil : last [] =@{option A} None.
Proof. done. Qed.
Lemma last_singleton x : last [x] = Some x.
Proof. done. Qed.
Lemma last_cons_cons x1 x2 l : last (x1 :: x2 :: l) = last (x2 :: l).
Proof. done. Qed.
Lemma last_app_cons l1 l2 x :
last (l1 ++ x :: l2) = last (x :: l2).
Proof. induction l1 as [|y [|y' l1] IHl]; done. Qed.
Lemma last_snoc x l : last (l ++ [x]) = Some x.
Proof. induction l as [|? []]; simpl; auto. Qed.
Lemma last_None l : last l = None l = [].
Proof.
split; [|by intros ->].
induction l as [|x1 [|x2 l] IH]; naive_solver.
Qed.
Lemma last_Some l x : last l = Some x l', l = l' ++ [x].
Proof.
split.
- destruct l as [|x' l'] using rev_ind; [done|].
rewrite last_snoc. naive_solver.
- intros [l' ->]. by rewrite last_snoc.
Qed.
Lemma last_is_Some l : is_Some (last l) l [].
Proof. rewrite <-not_eq_None_Some, last_None. naive_solver. Qed.
Lemma last_app l1 l2 :
last (l1 ++ l2) = match last l2 with Some y => Some y | None => last l1 end.
Proof.
destruct l2 as [|x l2 _] using rev_ind.
- by rewrite (right_id_L _ (++)).
- by rewrite (assoc_L (++)), !last_snoc.
Qed.
Lemma last_app_Some l1 l2 x :
last (l1 ++ l2) = Some x last l2 = Some x last l2 = None last l1 = Some x.
Proof. rewrite last_app. destruct (last l2); naive_solver. Qed.
Lemma last_app_None l1 l2 :
last (l1 ++ l2) = None last l1 = None last l2 = None.
Proof. rewrite last_app. destruct (last l2); naive_solver. Qed.
Lemma last_cons x l :
last (x :: l) = match last l with Some y => Some y | None => Some x end.
Proof. by apply (last_app [x]). Qed.
Lemma last_cons_Some_ne x y l :
x y last (x :: l) = Some y last l = Some y.
Proof. rewrite last_cons. destruct (last l); naive_solver. Qed.
Lemma last_lookup l : last l = l !! pred (length l).
Proof. by induction l as [| ?[]]. Qed.
Lemma last_reverse l : last (reverse l) = head l.
Proof. destruct l as [|x l]; simpl; by rewrite ?reverse_cons, ?last_snoc. Qed.
Lemma last_Some_elem_of l x :
last l = Some x x l.
Proof.
rewrite last_Some. intros [l' ->]. apply elem_of_app. right.
by apply elem_of_list_singleton.
Qed.
(** ** Properties of the [head] function *)
Lemma head_nil : head [] =@{option A} None.
Proof. done. Qed.
Lemma head_cons x l : head (x :: l) = Some x.
Proof. done. Qed.
Lemma head_None l : head l = None l = [].
Proof. split; [|by intros ->]. by destruct l. Qed.
Lemma head_Some l x : head l = Some x l', l = x :: l'.
Proof. split; [destruct l as [|x' l]; naive_solver | by intros [l' ->]]. Qed.
Lemma head_is_Some l : is_Some (head l) l [].
Proof. rewrite <-not_eq_None_Some, head_None. naive_solver. Qed.
Lemma head_snoc x l :
head (l ++ [x]) = match head l with Some y => Some y | None => Some x end.
Proof. by destruct l. Qed.
Lemma head_snoc_snoc x1 x2 l :
head (l ++ [x1; x2]) = head (l ++ [x1]).
Proof. by destruct l. Qed.
Lemma head_lookup l : head l = l !! 0.
Proof. by destruct l. Qed.
Lemma head_app l1 l2 :
head (l1 ++ l2) = match head l1 with Some y => Some y | None => head l2 end.
Proof. by destruct l1. Qed.
Lemma head_app_Some l1 l2 x :
head (l1 ++ l2) = Some x head l1 = Some x head l1 = None head l2 = Some x.
Proof. rewrite head_app. destruct (head l1); naive_solver. Qed.
Lemma head_app_None l1 l2 :
head (l1 ++ l2) = None head l1 = None head l2 = None.
Proof. rewrite head_app. destruct (head l1); naive_solver. Qed.
Lemma head_reverse l : head (reverse l) = last l.
Proof. by rewrite <-last_reverse, reverse_involutive. Qed.
Lemma head_Some_elem_of l x :
head l = Some x x l.
Proof. rewrite head_Some. intros [l' ->]. left. Qed.
(** ** Properties of the [take] function *)
Definition take_drop i l : take i l ++ drop i l = l := firstn_skipn i l.
Lemma take_drop_middle l i x :
l !! i = Some x take i l ++ x :: drop (S i) l = l.
Proof.
revert i x. induction l; intros [|?] ??; simplify_eq/=; f_equal; auto.
Qed.
Lemma take_0 l : take 0 l = [].
Proof. reflexivity. Qed.
Lemma take_nil n : take n [] =@{list A} [].
Proof. by destruct n. Qed.
Lemma take_S_r l n x : l !! n = Some x take (S n) l = take n l ++ [x].
Proof. revert n. induction l; intros []; naive_solver eauto with f_equal. Qed.
Lemma take_ge l n : length l n take n l = l.
Proof. revert n. induction l; intros [|?] ?; f_equal/=; auto with lia. Qed.
(** [take_app] is the most general lemma for [take] and [app]. Below that we
establish a number of useful corollaries. *)
Lemma take_app l k n : take n (l ++ k) = take n l ++ take (n - length l) k.
Proof. apply firstn_app. Qed.
Lemma take_app_ge l k n :
length l n take n (l ++ k) = l ++ take (n - length l) k.
Proof. intros. by rewrite take_app, take_ge. Qed.
Lemma take_app_le l k n : n length l take n (l ++ k) = take n l.
Proof.
intros. by rewrite take_app, (proj2 (Nat.sub_0_le _ _)), take_0, (right_id _ _).
Qed.
Lemma take_app_add l k m :
take (length l + m) (l ++ k) = l ++ take m k.
Proof. rewrite take_app, take_ge by lia. repeat f_equal; lia. Qed.
Lemma take_app_add' l k n m :
n = length l take (n + m) (l ++ k) = l ++ take m k.
Proof. intros ->. apply take_app_add. Qed.
Lemma take_app_length l k : take (length l) (l ++ k) = l.
Proof. by rewrite take_app, take_ge, Nat.sub_diag, take_0, (right_id _ _). Qed.
Lemma take_app_length' l k n : n = length l take n (l ++ k) = l.
Proof. intros ->. by apply take_app_length. Qed.
Lemma take_app3_length l1 l2 l3 : take (length l1) ((l1 ++ l2) ++ l3) = l1.
Proof. by rewrite <-(assoc_L (++)), take_app_length. Qed.
Lemma take_app3_length' l1 l2 l3 n :
n = length l1 take n ((l1 ++ l2) ++ l3) = l1.
Proof. intros ->. by apply take_app3_length. Qed.
Lemma take_take l n m : take n (take m l) = take (min n m) l.
Proof. revert n m. induction l; intros [|?] [|?]; f_equal/=; auto. Qed.
Lemma take_idemp l n : take n (take n l) = take n l.
Proof. by rewrite take_take, Nat.min_id. Qed.
Lemma length_take l n : length (take n l) = min n (length l).
Proof. revert n. induction l; intros [|?]; f_equal/=; done. Qed.
Lemma length_take_le l n : n length l length (take n l) = n.
Proof. rewrite length_take. apply Nat.min_l. Qed.
Lemma length_take_ge l n : length l n length (take n l) = length l.
Proof. rewrite length_take. apply Nat.min_r. Qed.
Lemma take_drop_commute l n m : take n (drop m l) = drop m (take (m + n) l).
Proof.
revert n m. induction l; intros [|?][|?]; simpl; auto using take_nil with lia.
Qed.
Lemma lookup_take l n i : i < n take n l !! i = l !! i.
Proof. revert n i. induction l; intros [|n] [|i] ?; simpl; auto with lia. Qed.
Lemma lookup_total_take `{!Inhabited A} l n i : i < n take n l !!! i = l !!! i.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_take. Qed.
Lemma lookup_take_ge l n i : n i take n l !! i = None.
Proof. revert n i. induction l; intros [|?] [|?] ?; simpl; auto with lia. Qed.
Lemma lookup_total_take_ge `{!Inhabited A} l n i : n i take n l !!! i = inhabitant.
Proof. intros. by rewrite list_lookup_total_alt, lookup_take_ge. Qed.
Lemma lookup_take_Some l n i a : take n l !! i = Some a l !! i = Some a i < n.
Proof.
split.
- destruct (decide (i < n)).
+ rewrite lookup_take; naive_solver.
+ rewrite lookup_take_ge; [done|lia].
- intros [??]. by rewrite lookup_take.
Qed.
Lemma elem_of_take x n l : x take n l i, l !! i = Some x i < n.
Proof.
rewrite elem_of_list_lookup. setoid_rewrite lookup_take_Some. naive_solver.
Qed.
Lemma take_alter f 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 n i x : n i take n (<[i:=x]>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_insert_ne by lia.
Qed.
Lemma take_insert_lt l n i x : i < n take n (<[i:=x]>l) = <[i:=x]>(take n l).
Proof.
revert l i. induction n as [|? IHn]; auto; simpl.
intros [|] [|] ?; auto; simpl. by rewrite IHn by lia.
Qed.
(** ** Properties of the [drop] function *)
Lemma drop_0 l : drop 0 l = l.
Proof. done. Qed.
Lemma drop_nil n : drop n [] =@{list A} [].
Proof. by destruct n. Qed.
Lemma drop_S l x n :
l !! n = Some x drop n l = x :: drop (S n) l.
Proof. revert n. induction l; intros []; naive_solver. Qed.
Lemma length_drop l n : length (drop n l) = length l - n.
Proof. revert n. by induction l; intros [|i]; f_equal/=. Qed.
Lemma drop_ge l n : length l n drop n l = [].
Proof. revert n. induction l; intros [|?]; simpl in *; auto with lia. Qed.
Lemma drop_all l : drop (length l) l = [].
Proof. by apply drop_ge. Qed.
Lemma drop_drop l n1 n2 : drop n1 (drop n2 l) = drop (n2 + n1) l.
Proof. revert n2. induction l; intros [|?]; simpl; rewrite ?drop_nil; auto. Qed.
(** [drop_app] is the most general lemma for [drop] and [app]. Below we prove a
number of useful corollaries. *)
Lemma drop_app l k n : drop n (l ++ k) = drop n l ++ drop (n - length l) k.
Proof. apply skipn_app. Qed.
Lemma drop_app_ge l k n :
length l n drop n (l ++ k) = drop (n - length l) k.
Proof. intros. by rewrite drop_app, drop_ge. Qed.
Lemma drop_app_le l k n :
n length l drop n (l ++ k) = drop n l ++ k.
Proof. intros. by rewrite drop_app, (proj2 (Nat.sub_0_le _ _)), drop_0. Qed.
Lemma drop_app_add l k m :
drop (length l + m) (l ++ k) = drop m k.
Proof. rewrite drop_app, drop_ge by lia. simpl. f_equal; lia. Qed.
Lemma drop_app_add' l k n m :
n = length l drop (n + m) (l ++ k) = drop m k.
Proof. intros ->. apply drop_app_add. Qed.
Lemma drop_app_length l k : drop (length l) (l ++ k) = k.
Proof. by rewrite drop_app_le, drop_all. Qed.
Lemma drop_app_length' l k n : n = length l drop n (l ++ k) = k.
Proof. intros ->. by apply drop_app_length. Qed.
Lemma drop_app3_length l1 l2 l3 :
drop (length l1) ((l1 ++ l2) ++ l3) = l2 ++ l3.
Proof. by rewrite <-(assoc_L (++)), drop_app_length. Qed.
Lemma drop_app3_length' l1 l2 l3 n :
n = length l1 drop n ((l1 ++ l2) ++ l3) = l2 ++ l3.
Proof. intros ->. apply drop_app3_length. Qed.
Lemma lookup_drop l n i : drop n l !! i = l !! (n + i).
Proof. revert n i. induction l; intros [|i] ?; simpl; auto. Qed.
Lemma lookup_total_drop `{!Inhabited A} l n i : drop n l !!! i = l !!! (n + i).
Proof. by rewrite !list_lookup_total_alt, lookup_drop. Qed.
Lemma drop_alter f l n i : i < n drop n (alter f i l) = drop n l.
Proof.
intros. apply list_eq. intros j.
by rewrite !lookup_drop, !list_lookup_alter_ne by lia.
Qed.
Lemma drop_insert_le l n i x : n i drop n (<[i:=x]>l) = <[i-n:=x]>(drop n l).
Proof. revert i n. induction l; intros [] []; naive_solver lia. Qed.
Lemma drop_insert_gt l n i x : i < n drop n (<[i:=x]>l) = drop n l.
Proof.
intros. apply list_eq. intros j.
by rewrite !lookup_drop, !list_lookup_insert_ne by lia.
Qed.
Lemma delete_take_drop l i : delete i l = take i l ++ drop (S i) l.
Proof. revert i. induction l; intros [|?]; f_equal/=; auto. Qed.
Lemma take_take_drop l n m : take n l ++ take m (drop n l) = take (n + m) l.
Proof. revert n m. induction l; intros [|?] [|?]; f_equal/=; auto. Qed.
Lemma drop_take_drop l n m : n m drop n (take m l) ++ drop m l = drop n l.
Proof.
revert n m. induction l; intros [|?] [|?] ?;
f_equal/=; auto using take_drop with lia.
Qed.
Lemma insert_take_drop l i x :
i < length l
<[i:=x]> l = take i l ++ x :: drop (S i) l.
Proof.
intros Hi.
rewrite <-(take_drop_middle (<[i:=x]> l) i x).
2:{ by rewrite list_lookup_insert. }
rewrite take_insert by done.
rewrite drop_insert_gt by lia.
done.
Qed.
(** ** Interaction between the [take]/[drop]/[reverse] functions *)
Lemma take_reverse l n : take n (reverse l) = reverse (drop (length l - n) l).
Proof. unfold reverse; rewrite <-!rev_alt. apply firstn_rev. Qed.
Lemma drop_reverse l n : drop n (reverse l) = reverse (take (length l - n) l).
Proof. unfold reverse; rewrite <-!rev_alt. apply skipn_rev. Qed.
Lemma reverse_take l n : reverse (take n l) = drop (length l - n) (reverse l).
Proof.
rewrite drop_reverse. destruct (decide (n length l)).
- repeat f_equal; lia.
- by rewrite !take_ge by lia.
Qed.
Lemma reverse_drop l n : reverse (drop n l) = take (length l - n) (reverse l).
Proof.
rewrite take_reverse. destruct (decide (n length l)).
- repeat f_equal; lia.
- by rewrite !drop_ge by lia.
Qed.
(** ** Other lemmas that use [take]/[drop] in their proof. *)
Lemma app_eq_inv l1 l2 k1 k2 :
l1 ++ l2 = k1 ++ k2
( k, l1 = k1 ++ k k2 = k ++ l2) ( k, k1 = l1 ++ k l2 = k ++ k2).
Proof.
intros Hlk. destruct (decide (length l1 < length k1)).
- right. rewrite <-(take_drop (length l1) k1), <-(assoc_L _) in Hlk.
apply app_inj_1 in Hlk as [Hl1 Hl2]; [|rewrite length_take; lia].
exists (drop (length l1) k1). by rewrite Hl1 at 1; rewrite take_drop.
- left. rewrite <-(take_drop (length k1) l1), <-(assoc_L _) in Hlk.
apply app_inj_1 in Hlk as [Hk1 Hk2]; [|rewrite length_take; lia].
exists (drop (length k1) l1). by rewrite <-Hk1 at 1; rewrite take_drop.
Qed.
(** ** Properties of the [replicate] function *)
Lemma length_replicate n x : length (replicate n x) = n.
Proof. induction n; simpl; auto. Qed.
Lemma lookup_replicate n x y i :
replicate n x !! i = Some y y = x i < n.
Proof.
split.
- revert i. induction n; intros [|?]; naive_solver auto with lia.
- intros [-> Hi]. revert i Hi.
induction n; intros [|?]; naive_solver auto with lia.
Qed.
Lemma elem_of_replicate n x y : y replicate n x y = x n 0.
Proof.
rewrite elem_of_list_lookup, Nat.neq_0_lt_0.
setoid_rewrite lookup_replicate; naive_solver eauto with lia.
Qed.
Lemma lookup_replicate_1 n x y i :
replicate n x !! i = Some y y = x i < n.
Proof. by rewrite lookup_replicate. Qed.
Lemma lookup_replicate_2 n x i : i < n replicate n x !! i = Some x.
Proof. by rewrite lookup_replicate. Qed.
Lemma lookup_total_replicate_2 `{!Inhabited A} n x i :
i < n replicate n x !!! i = x.
Proof. intros. by rewrite list_lookup_total_alt, lookup_replicate_2. Qed.
Lemma lookup_replicate_None n x i : n i replicate n x !! i = None.
Proof.
rewrite eq_None_not_Some, Nat.le_ngt. split.
- intros Hin [x' Hx']; destruct Hin. rewrite lookup_replicate in Hx'; tauto.
- intros Hx ?. destruct Hx. exists x; auto using lookup_replicate_2.
Qed.
Lemma insert_replicate x n i : <[i:=x]>(replicate n x) = replicate n x.
Proof. revert i. induction n; intros [|?]; f_equal/=; auto. Qed.
Lemma insert_replicate_lt x y n i :
i < n
<[i:=y]>(replicate n x) = replicate i x ++ y :: replicate (n - S i) x.
Proof.
revert i. induction n as [|n IH]; intros [|i] Hi; simpl; [lia..| |].
- by rewrite Nat.sub_0_r.
- by rewrite IH by lia.
Qed.
Lemma elem_of_replicate_inv x n y : x replicate n y x = y.
Proof. induction n; simpl; rewrite ?elem_of_nil, ?elem_of_cons; intuition. Qed.
Lemma replicate_S n x : replicate (S n) x = x :: replicate n x.
Proof. done. Qed.
Lemma replicate_S_end n x : replicate (S n) x = replicate n x ++ [x].
Proof. induction n; f_equal/=; auto. Qed.
Lemma replicate_add n m x :
replicate (n + m) x = replicate n x ++ replicate m x.
Proof. induction n; f_equal/=; auto. Qed.
Lemma replicate_cons_app n x :
x :: replicate n x = replicate n x ++ [x].
Proof. induction n; f_equal/=; eauto. Qed.
Lemma take_replicate n m x : take n (replicate m x) = replicate (min n m) x.
Proof. revert m. by induction n; intros [|?]; f_equal/=. Qed.
Lemma take_replicate_add n m x : 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 : drop n (replicate m x) = replicate (m - n) x.
Proof. revert m. by induction n; intros [|?]; f_equal/=. Qed.
Lemma drop_replicate_add n m x : drop n (replicate (n + m) x) = replicate m x.
Proof. rewrite drop_replicate. f_equal. lia. Qed.
Lemma replicate_as_elem_of x n l :
replicate n x = l length l = n y, y l y = x.
Proof.
split; [intros <-; eauto using elem_of_replicate_inv, length_replicate|].
intros [<- Hl]. symmetry. induction l as [|y l IH]; f_equal/=.
- apply Hl. by left.
- apply IH. intros ??. apply Hl. by right.
Qed.
Lemma reverse_replicate n x : reverse (replicate n x) = replicate n x.
Proof.
symmetry. apply replicate_as_elem_of.
rewrite length_reverse, length_replicate. split; auto.
intros y. rewrite elem_of_reverse. by apply elem_of_replicate_inv.
Qed.
Lemma replicate_false βs n : length βs = n replicate n false =.>* βs.
Proof. intros <-. by induction βs; simpl; constructor. Qed.
Lemma tail_replicate x n : tail (replicate n x) = replicate (pred n) x.
Proof. by destruct n. Qed.
Lemma head_replicate_Some x n : head (replicate n x) = Some x 0 < n.
Proof. destruct n; naive_solver lia. Qed.
(** ** Properties of the [filter] function *)
Section filter.
Context (P : A Prop) `{ x, Decision (P x)}.
Local Arguments filter {_ _ _} _ {_} !_ /.
Lemma filter_nil : filter P [] = [].
Proof. done. Qed.
Lemma filter_cons x l :
filter P (x :: l) = if decide (P x) then x :: filter P l else filter P l.
Proof. done. Qed.
Lemma filter_cons_True x l : P x filter P (x :: l) = x :: filter P l.
Proof. intros. by rewrite filter_cons, decide_True. Qed.
Lemma filter_cons_False x l : ¬P x filter P (x :: l) = filter P l.
Proof. intros. by rewrite filter_cons, decide_False. Qed.
Lemma filter_app l1 l2 : filter P (l1 ++ l2) = filter P l1 ++ filter P l2.
Proof.
induction l1 as [|x l1 IH]; simpl; [done| ].
case_decide; [|done].
by rewrite IH.
Qed.
Lemma elem_of_list_filter l x : x filter P l P x x l.
Proof.
induction l; simpl; repeat case_decide;
rewrite ?elem_of_nil, ?elem_of_cons; naive_solver.
Qed.
Lemma length_filter l : length (filter P l) length l.
Proof. induction l; simpl; repeat case_decide; simpl; lia. Qed.
Lemma length_filter_lt l x : x l ¬P x length (filter P l) < length l.
Proof.
intros (l1 & l2 & ->)%elem_of_list_split ?.
rewrite filter_app, !length_app, filter_cons, decide_False by done.
pose proof (length_filter l1); pose proof (length_filter l2). simpl. lia.
Qed.
Lemma filter_nil_not_elem_of l x : filter P l = [] P x x l.
Proof. induction 3; simplify_eq/=; case_decide; naive_solver. Qed.
Lemma filter_reverse l : filter P (reverse l) = reverse (filter P l).
Proof.
induction l as [|x l IHl]; [done|].
rewrite reverse_cons, filter_app, IHl, !filter_cons.
case_decide; [by rewrite reverse_cons|by rewrite filter_nil, app_nil_r].
Qed.
End filter.
Lemma list_filter_iff (P1 P2 : A Prop)
`{!∀ x, Decision (P1 x), !∀ x, Decision (P2 x)} (l : list A) :
( x, P1 x P2 x)
filter P1 l = filter P2 l.
Proof.
intros HPiff. induction l as [|a l IH]; [done|].
destruct (decide (P1 a)).
- rewrite !filter_cons_True by naive_solver. by rewrite IH.
- rewrite !filter_cons_False by naive_solver. by rewrite IH.
Qed.
Lemma list_filter_filter (P1 P2 : A Prop)
`{!∀ x, Decision (P1 x), !∀ x, Decision (P2 x)} (l : list A) :
filter P1 (filter P2 l) = filter (λ a, P1 a P2 a) l.
Proof.
induction l as [|x l IH]; [done|].
rewrite !filter_cons. case (decide (P2 x)) as [HP2|HP2].
- rewrite filter_cons, IH. apply decide_ext. naive_solver.
- rewrite IH. symmetry. apply decide_False. by intros [_ ?].
Qed.
Lemma list_filter_filter_l (P1 P2 : A Prop)
`{!∀ x, Decision (P1 x), !∀ x, Decision (P2 x)} (l : list A) :
( x, P1 x P2 x)
filter P1 (filter P2 l) = filter P1 l.
Proof.
intros HPimp. rewrite list_filter_filter.
apply list_filter_iff. naive_solver.
Qed.
Lemma list_filter_filter_r (P1 P2 : A Prop)
`{!∀ x, Decision (P1 x), !∀ x, Decision (P2 x)} (l : list A) :
( x, P2 x P1 x)
filter P1 (filter P2 l) = filter P2 l.
Proof.
intros HPimp. rewrite list_filter_filter.
apply list_filter_iff. naive_solver.
Qed.
End general_properties.
(** * Basic tactics on lists *)
(** These are used already in [list_relations] so we cannot have them in
[list_tactics]. *)
(** The tactic [discriminate_list] discharges a goal if there is a hypothesis
[l1 = l2] where [l1] and [l2] have different lengths. The tactic is guaranteed
to work if [l1] and [l2] contain solely [ [] ], [(::)] and [(++)]. *)
Tactic Notation "discriminate_list" hyp(H) :=
apply (f_equal length) in H;
repeat (csimpl in H || rewrite length_app in H); exfalso; lia.
Tactic Notation "discriminate_list" :=
match goal with H : _ =@{list _} _ |- _ => discriminate_list H end.
(** The tactic [simplify_list_eq] simplifies hypotheses involving
equalities on lists using injectivity of [(::)] and [(++)]. Also, it simplifies
lookups in singleton lists. *)
Ltac simplify_list_eq :=
repeat match goal with
| _ => progress simplify_eq/=
| H : _ ++ _ = _ ++ _ |- _ => first
[ apply app_inv_head in H | apply app_inv_tail in H
| apply app_inj_1 in H; [destruct H|done]
| apply app_inj_2 in H; [destruct H|done] ]
| H : [?x] !! ?i = Some ?y |- _ =>
destruct i; [change (Some x = Some y) in H | discriminate]
end.
From Coq Require Export Permutation.
From stdpp Require Export numbers base option list_basics list_relations list_monad.
From stdpp Require Import options.
(** The function [list_find P l] returns the first index [i] whose element
satisfies the predicate [P]. *)
Definition list_find {A} P `{ x, Decision (P x)} : list A option (nat * A) :=
fix go l :=
match l with
| [] => None
| x :: l => if decide (P x) then Some (0,x) else prod_map S id <$> go l
end.
Global Instance: Params (@list_find) 3 := {}.
(** 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.
Global Arguments resize {_} !_ _ !_ : assert.
Global Instance: Params (@resize) 2 := {}.
(** The function [rotate n l] rotates the list [l] by [n], e.g., [rotate 1
[x0; x1; ...; xm]] becomes [x1; ...; xm; x0]. Rotating by a multiple of
[length l] is the identity function. **)
Definition rotate {A} (n : nat) (l : list A) : list A :=
drop (n `mod` length l) l ++ take (n `mod` length l) l.
Global Instance: Params (@rotate) 2 := {}.
(** The function [rotate_take s e l] returns the range between the
indices [s] (inclusive) and [e] (exclusive) of [l]. If [e ≤ s], all
elements after [s] and before [e] are returned. *)
Definition rotate_take {A} (s e : nat) (l : list A) : list A :=
take (rotate_nat_sub s e (length l)) (rotate s l).
Global Instance: Params (@rotate_take) 3 := {}.
(** The function [reshape k l] transforms [l] into a list of lists whose sizes
are specified by [k]. In case [l] is too short, the resulting list will be
padded with empty lists. In case [l] is too long, it will be truncated. *)
Fixpoint reshape {A} (szs : list nat) (l : list A) : list (list A) :=
match szs with
| [] => [] | sz :: szs => take sz l :: reshape szs (drop sz l)
end.
Global Instance: Params (@reshape) 2 := {}.
Definition sublist_lookup {A} (i n : nat) (l : list A) : option (list A) :=
guard (i + n length l);; Some (take n (drop i l)).
Definition sublist_alter {A} (f : list A list A)
(i n : nat) (l : list A) : list A :=
take i l ++ f (take n (drop i l)) ++ drop (i + n) l.
(** The function [mask f βs l] applies the function [f] to elements in [l] at
positions that are [true] in [βs]. *)
Fixpoint mask {A} (f : A A) (βs : list bool) (l : list A) : list A :=
match βs, l with
| β :: βs, x :: l => (if β then f x else x) :: mask f βs l
| _, _ => l
end.
(** These next functions allow to efficiently encode lists of positives (bit
strings) into a single positive and go in the other direction as well. This is
for example used for the countable instance of lists and in namespaces.
The main functions are [positives_flatten] and [positives_unflatten]. *)
Fixpoint positives_flatten_go (xs : list positive) (acc : positive) : positive :=
match xs with
| [] => acc
| x :: xs => positives_flatten_go xs (acc~1~0 ++ Pos.reverse (Pos.dup x))
end.
(** Flatten a list of positives into a single positive by duplicating the bits
of each element, so that:
- [0 -> 00]
- [1 -> 11]
and then separating each element with [10]. *)
Definition positives_flatten (xs : list positive) : positive :=
positives_flatten_go xs 1.
Fixpoint positives_unflatten_go
(p : positive)
(acc_xs : list positive)
(acc_elm : positive)
: option (list positive) :=
match p with
| 1 => Some acc_xs
| p'~0~0 => positives_unflatten_go p' acc_xs (acc_elm~0)
| p'~1~1 => positives_unflatten_go p' acc_xs (acc_elm~1)
| p'~1~0 => positives_unflatten_go p' (acc_elm :: acc_xs) 1
| _ => None
end%positive.
(** Unflatten a positive into a list of positives, assuming the encoding
used by [positives_flatten]. *)
Definition positives_unflatten (p : positive) : option (list positive) :=
positives_unflatten_go p [] 1.
Section general_properties.
Context {A : Type}.
Implicit Types x y z : A.
Implicit Types l k : list A.
(** * Properties of the [find] function *)
Section find.
Context (P : A Prop) `{!∀ x, Decision (P x)}.
Lemma list_find_Some l i x :
list_find P l = Some (i,x)
l !! i = Some x P x j y, l !! j = Some y j < i ¬P y.
Proof.
revert i. induction l as [|y l IH]; intros i; csimpl; [naive_solver|].
case_decide.
- split; [naive_solver lia|]. intros (Hi&HP&Hlt).
destruct i as [|i]; simplify_eq/=; [done|].
destruct (Hlt 0 y); naive_solver lia.
- split.
+ intros ([i' x']&Hl&?)%fmap_Some; simplify_eq/=.
apply IH in Hl as (?&?&Hlt). split_and!; [done..|].
intros [|j] ?; naive_solver lia.
+ intros (?&?&Hlt). destruct i as [|i]; simplify_eq/=; [done|].
rewrite (proj2 (IH i)); [done|]. split_and!; [done..|].
intros j z ???. destruct (Hlt (S j) z); naive_solver lia.
Qed.
Lemma list_find_elem_of l x : x l P x is_Some (list_find P l).
Proof.
induction 1 as [|x y l ? IH]; intros; simplify_option_eq; eauto.
by destruct IH as [[i x'] ->]; [|exists (S i, x')].
Qed.
Lemma list_find_None l :
list_find P l = None Forall (λ x, ¬P x) l.
Proof.
rewrite eq_None_not_Some, Forall_forall. split.
- intros Hl x Hx HP. destruct Hl. eauto using list_find_elem_of.
- intros HP [[i x] (?%elem_of_list_lookup_2&?&?)%list_find_Some]; naive_solver.
Qed.
Lemma list_find_app_None l1 l2 :
list_find P (l1 ++ l2) = None list_find P l1 = None list_find P l2 = None.
Proof. by rewrite !list_find_None, Forall_app. Qed.
Lemma list_find_app_Some l1 l2 i x :
list_find P (l1 ++ l2) = Some (i,x)
list_find P l1 = Some (i,x)
length l1 i list_find P l1 = None list_find P l2 = Some (i - length l1,x).
Proof.
split.
- intros ([?|[??]]%lookup_app_Some&?&Hleast)%list_find_Some.
+ left. apply list_find_Some; eauto using lookup_app_l_Some.
+ right. split; [lia|]. split.
{ apply list_find_None, Forall_lookup. intros j z ??.
assert (j < length l1) by eauto using lookup_lt_Some.
naive_solver eauto using lookup_app_l_Some with lia. }
apply list_find_Some. split_and!; [done..|].
intros j z ??. eapply (Hleast (length l1 + j)); [|lia].
by rewrite lookup_app_r, Nat.add_sub' by lia.
- intros [(?&?&Hleast)%list_find_Some|(?&Hl1&(?&?&Hleast)%list_find_Some)].
+ apply list_find_Some. split_and!; [by auto using lookup_app_l_Some..|].
assert (i < length l1) by eauto using lookup_lt_Some.
intros j y ?%lookup_app_Some; naive_solver eauto with lia.
+ rewrite list_find_Some, lookup_app_Some. split_and!; [by auto..|].
intros j y [?|?]%lookup_app_Some ?; [|naive_solver auto with lia].
by eapply (Forall_lookup_1 (not P) l1); [by apply list_find_None|..].
Qed.
Lemma list_find_app_l l1 l2 i x:
list_find P l1 = Some (i, x) list_find P (l1 ++ l2) = Some (i, x).
Proof. rewrite list_find_app_Some. auto. Qed.
Lemma list_find_app_r l1 l2:
list_find P l1 = None
list_find P (l1 ++ l2) = prod_map (λ n, n + length l1) id <$> list_find P l2.
Proof.
intros. apply option_eq; intros [j y]. rewrite list_find_app_Some. split.
- intros [?|(?&?&->)]; naive_solver auto with f_equal lia.
- intros ([??]&->&?)%fmap_Some; naive_solver auto with f_equal lia.
Qed.
Lemma list_find_insert_Some l i j x y :
list_find P (<[i:=x]> l) = Some (j,y)
(j < i list_find P l = Some (j,y))
(i = j x = y j < length l P x i' z, l !! i' = Some z i' < i ¬P z)
(i < j ¬P x list_find P l = Some (j,y) z, l !! i = Some z ¬P z)
( z, i < j ¬P x P y P z l !! i = Some z l !! j = Some y
i' z, l !! i' = Some z i' i i' < j ¬P z).
Proof.
split.
- intros ([(->&->&?)|[??]]%list_lookup_insert_Some&?&Hleast)%list_find_Some.
{ right; left. split_and!; [done..|]. intros k z ??.
apply (Hleast k); [|lia]. by rewrite list_lookup_insert_ne by lia. }
assert (j < i i < j) as [?|?] by lia.
{ left. rewrite list_find_Some. split_and!; [by auto..|]. intros k z ??.
apply (Hleast k); [|lia]. by rewrite list_lookup_insert_ne by lia. }
right; right. assert (j < length l) by eauto using lookup_lt_Some.
destruct (lookup_lt_is_Some_2 l i) as [z ?]; [lia|].
destruct (decide (P z)).
{ right. exists z. split_and!; [done| |done..|].
+ apply (Hleast i); [|done]. by rewrite list_lookup_insert by lia.
+ intros k z' ???.
apply (Hleast k); [|lia]. by rewrite list_lookup_insert_ne by lia. }
left. split_and!; [done|..|naive_solver].
+ apply (Hleast i); [|done]. by rewrite list_lookup_insert by lia.
+ apply list_find_Some. split_and!; [by auto..|]. intros k z' ??.
destruct (decide (k = i)) as [->|]; [naive_solver|].
apply (Hleast k); [|lia]. by rewrite list_lookup_insert_ne by lia.
- intros [[? Hl]|[(->&->&?&?&Hleast)|[(?&?&Hl&Hnot)|(z&?&?&?&?&?&?&?Hleast)]]];
apply list_find_Some.
+ apply list_find_Some in Hl as (?&?&Hleast).
rewrite list_lookup_insert_ne by lia. split_and!; [done..|].
intros k z [(->&->&?)|[??]]%list_lookup_insert_Some; eauto with lia.
+ rewrite list_lookup_insert by done. split_and!; [by auto..|].
intros k z [(->&->&?)|[??]]%list_lookup_insert_Some; eauto with lia.
+ apply list_find_Some in Hl as (?&?&Hleast).
rewrite list_lookup_insert_ne by lia. split_and!; [done..|].
intros k z [(->&->&?)|[??]]%list_lookup_insert_Some; eauto with lia.
+ rewrite list_lookup_insert_ne by lia. split_and!; [done..|].
intros k z' [(->&->&?)|[??]]%list_lookup_insert_Some; eauto with lia.
Qed.
Lemma list_find_ext (Q : A Prop) `{ x, Decision (Q x)} l :
( x, P x Q x)
list_find P l = list_find Q l.
Proof.
intros HPQ. induction l as [|x l IH]; simpl; [done|].
by rewrite (decide_ext (P x) (Q x)), IH by done.
Qed.
End find.
Lemma list_find_fmap {B} (f : A B) (P : B Prop)
`{!∀ y : B, Decision (P y)} (l : list A) :
list_find P (f <$> l) = prod_map id f <$> list_find (P f) l.
Proof.
induction l as [|x l IH]; [done|]; csimpl. (* csimpl re-folds fmap *)
case_decide; [done|].
rewrite IH. by destruct (list_find (P f) l).
Qed.
(** ** Properties of the [resize] function *)
Lemma resize_spec l n x : resize n x l = take n l ++ replicate (n - length l) x.
Proof. revert n. induction l; intros [|?]; f_equal/=; auto. Qed.
Lemma resize_0 l x : resize 0 x l = [].
Proof. by destruct l. Qed.
Lemma resize_nil n x : resize n x [] = replicate n x.
Proof. rewrite resize_spec. rewrite take_nil. f_equal/=. lia. Qed.
Lemma resize_ge l 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 n x : n length l resize n x l = take n l.
Proof.
intros. rewrite resize_spec, (proj2 (Nat.sub_0_le _ _)) by done.
simpl. by rewrite (right_id_L [] (++)).
Qed.
Lemma resize_all l x : resize (length l) x l = l.
Proof. intros. by rewrite resize_le, take_ge. Qed.
Lemma resize_all_alt l n x : n = length l resize n x l = l.
Proof. intros ->. by rewrite resize_all. Qed.
Lemma resize_add l 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 [|?] [|?]; f_equal/=; auto.
- by rewrite Nat.add_0_r, (right_id_L [] (++)).
- by rewrite replicate_add.
Qed.
Lemma resize_add_eq l n m x :
length l = n resize (n + m) x l = l ++ replicate m x.
Proof. intros <-. by rewrite resize_add, resize_all, drop_all, resize_nil. Qed.
Lemma resize_app_le l1 l2 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 ?length_app; lia).
Qed.
Lemma resize_app l1 l2 n x : n = length l1 resize n x (l1 ++ l2) = l1.
Proof. intros ->. by rewrite resize_app_le, resize_all. Qed.
Lemma resize_app_ge l1 l2 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, (assoc_L (++)) by done.
do 2 f_equal. rewrite length_app. lia.
Qed.
Lemma length_resize l n x : length (resize n x l) = n.
Proof. rewrite resize_spec, length_app, length_replicate, length_take. lia. Qed.
Lemma resize_replicate x n m : resize n x (replicate m x) = replicate n x.
Proof. revert m. induction n; intros [|?]; f_equal/=; auto. Qed.
Lemma resize_resize l 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 [|?] [|?] ?; f_equal/=; auto with lia.
Qed.
Lemma resize_idemp l n x : resize n x (resize n x l) = resize n x l.
Proof. by rewrite resize_resize. Qed.
Lemma resize_take_le l n m x : n m resize n x (take m l) = resize n x l.
Proof. revert n m. induction l; intros [|?][|?] ?; f_equal/=; auto with lia. Qed.
Lemma resize_take_eq l n x : resize n x (take n l) = resize n x l.
Proof. by rewrite resize_take_le. Qed.
Lemma take_resize l n m x : take n (resize m x l) = resize (min n m) x l.
Proof.
revert n m. induction l; intros [|?][|?]; f_equal/=; auto using take_replicate.
Qed.
Lemma take_resize_le l n m x : n m take n (resize m x l) = resize n x l.
Proof. intros. by rewrite take_resize, Nat.min_l. Qed.
Lemma take_resize_eq l n x : take n (resize n x l) = resize n x l.
Proof. intros. by rewrite take_resize, Nat.min_l. Qed.
Lemma take_resize_add l 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 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_add l 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.
Lemma lookup_resize l n x i : i < n i < length l resize n x l !! i = l !! i.
Proof.
intros ??. destruct (decide (n < length l)).
- by rewrite resize_le, lookup_take by lia.
- by rewrite resize_ge, lookup_app_l by lia.
Qed.
Lemma lookup_total_resize `{!Inhabited A} l n x i :
i < n i < length l resize n x l !!! i = l !!! i.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_resize. Qed.
Lemma lookup_resize_new l n x i :
length l i i < n resize n x l !! i = Some x.
Proof.
intros ??. rewrite resize_ge by lia.
replace i with (length l + (i - length l)) by lia.
by rewrite lookup_app_r, lookup_replicate_2 by lia.
Qed.
Lemma lookup_total_resize_new `{!Inhabited A} l n x i :
length l i i < n resize n x l !!! i = x.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_resize_new. Qed.
Lemma lookup_resize_old l n x i : n i resize n x l !! i = None.
Proof. intros ?. apply lookup_ge_None_2. by rewrite length_resize. Qed.
Lemma lookup_total_resize_old `{!Inhabited A} l n x i :
n i resize n x l !!! i = inhabitant.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_resize_old. Qed.
Lemma Forall_resize P n x l : P x Forall P l Forall P (resize n x l).
Proof.
intros ? Hl. revert n.
induction Hl; intros [|?]; simpl; auto using Forall_replicate.
Qed.
Lemma fmap_resize {B} (f : A B) n x l : f <$> resize n x l = resize n (f x) (f <$> l).
Proof.
revert n. induction l; intros [|?]; f_equal/=; auto using fmap_replicate.
Qed.
Lemma Forall_resize_inv P n x l :
length l n Forall P (resize n x l) Forall P l.
Proof. intros ?. rewrite resize_ge, Forall_app by done. by intros []. Qed.
(** ** Properties of the [rotate] function *)
Lemma rotate_replicate n1 n2 x:
rotate n1 (replicate n2 x) = replicate n2 x.
Proof.
unfold rotate. rewrite drop_replicate, take_replicate, <-replicate_add.
f_equal. lia.
Qed.
Lemma length_rotate l n:
length (rotate n l) = length l.
Proof. unfold rotate. rewrite length_app, length_drop, length_take. lia. Qed.
Lemma lookup_rotate_r l n i:
i < length l
rotate n l !! i = l !! rotate_nat_add n i (length l).
Proof.
intros Hlen. pose proof (Nat.mod_upper_bound n (length l)) as ?.
unfold rotate. rewrite rotate_nat_add_add_mod, rotate_nat_add_alt by lia.
remember (n `mod` length l) as n'.
case_decide.
- by rewrite lookup_app_l, lookup_drop by (rewrite length_drop; lia).
- rewrite lookup_app_r, lookup_take, length_drop by (rewrite length_drop; lia).
f_equal. lia.
Qed.
Lemma lookup_rotate_r_Some l n i x:
rotate n l !! i = Some x
l !! rotate_nat_add n i (length l) = Some x i < length l.
Proof.
split.
- intros Hl. pose proof (lookup_lt_Some _ _ _ Hl) as Hlen.
rewrite length_rotate in Hlen. by rewrite <-lookup_rotate_r.
- intros [??]. by rewrite lookup_rotate_r.
Qed.
Lemma lookup_rotate_l l n i:
i < length l rotate n l !! rotate_nat_sub n i (length l) = l !! i.
Proof.
intros ?. rewrite lookup_rotate_r, rotate_nat_add_sub;[done..|].
apply rotate_nat_sub_lt. lia.
Qed.
Lemma elem_of_rotate l n x:
x rotate n l x l.
Proof.
unfold rotate. rewrite <-(take_drop (n `mod` length l) l) at 5.
rewrite !elem_of_app. naive_solver.
Qed.
Lemma rotate_insert_l l n i x:
i < length l
rotate n (<[rotate_nat_add n i (length l):=x]> l) = <[i:=x]> (rotate n l).
Proof.
intros Hlen. pose proof (Nat.mod_upper_bound n (length l)) as ?. unfold rotate.
rewrite length_insert, rotate_nat_add_add_mod, rotate_nat_add_alt by lia.
remember (n `mod` length l) as n'.
case_decide.
- rewrite take_insert, drop_insert_le, insert_app_l
by (rewrite ?length_drop; lia). do 2 f_equal. lia.
- rewrite take_insert_lt, drop_insert_gt, insert_app_r_alt, length_drop
by (rewrite ?length_drop; lia). do 2 f_equal. lia.
Qed.
Lemma rotate_insert_r l n i x:
i < length l
rotate n (<[i:=x]> l) = <[rotate_nat_sub n i (length l):=x]> (rotate n l).
Proof.
intros ?. rewrite <-rotate_insert_l, rotate_nat_add_sub;[done..|].
apply rotate_nat_sub_lt. lia.
Qed.
(** ** Properties of the [rotate_take] function *)
Lemma rotate_take_insert l s e i x:
i < length l
rotate_take s e (<[i:=x]>l) =
if decide (rotate_nat_sub s i (length l) < rotate_nat_sub s e (length l)) then
<[rotate_nat_sub s i (length l):=x]> (rotate_take s e l) else rotate_take s e l.
Proof.
intros ?. unfold rotate_take. rewrite rotate_insert_r, length_insert by done.
case_decide; [rewrite take_insert_lt | rewrite take_insert]; naive_solver lia.
Qed.
Lemma rotate_take_add l b i :
i < length l
rotate_take b (rotate_nat_add b i (length l)) l = take i (rotate b l).
Proof. intros ?. unfold rotate_take. by rewrite rotate_nat_sub_add. Qed.
(** ** Properties of the [reshape] function *)
Lemma length_reshape szs l : length (reshape szs l) = length szs.
Proof. revert l. by induction szs; intros; f_equal/=. Qed.
Lemma Forall_reshape P l szs : Forall P l Forall (Forall P) (reshape szs l).
Proof.
revert l. induction szs; simpl; auto using Forall_take, Forall_drop.
Qed.
(** ** Properties of [sublist_lookup] and [sublist_alter] *)
Lemma sublist_lookup_length l i n k :
sublist_lookup i n l = Some k length k = n.
Proof.
unfold sublist_lookup; intros; simplify_option_eq.
rewrite length_take, length_drop; lia.
Qed.
Lemma sublist_lookup_all l n : length l = n sublist_lookup 0 n l = Some l.
Proof.
intros. unfold sublist_lookup; case_guard; [|lia].
by rewrite take_ge by (rewrite length_drop; lia).
Qed.
Lemma sublist_lookup_Some l i n :
i + n length l sublist_lookup i n l = Some (take n (drop i l)).
Proof. by unfold sublist_lookup; intros; simplify_option_eq. Qed.
Lemma sublist_lookup_Some' l i n l' :
sublist_lookup i n l = Some l' l' = take n (drop i l) i + n length l.
Proof. unfold sublist_lookup. case_guard; naive_solver lia. Qed.
Lemma sublist_lookup_None l i n :
length l < i + n sublist_lookup i n l = None.
Proof. by unfold sublist_lookup; intros; simplify_option_eq by lia. Qed.
Lemma sublist_eq l k n :
(n | length l) (n | length k)
( i, sublist_lookup (i * n) n l = sublist_lookup (i * n) n k) l = k.
Proof.
revert l k. assert ( l i,
n 0 (n | length l) ¬n * i `div` n + n length l length l i).
{ intros l i ? [j ->] Hjn. apply Nat.nlt_ge; contradict Hjn.
rewrite <-Nat.mul_succ_r, (Nat.mul_comm n).
apply Nat.mul_le_mono_r, Nat.le_succ_l, Nat.div_lt_upper_bound; lia. }
intros l k Hl Hk Hlookup. destruct (decide (n = 0)) as [->|].
{ by rewrite (nil_length_inv l),
(nil_length_inv k) by eauto using Nat.divide_0_l. }
apply list_eq; intros i. specialize (Hlookup (i `div` n)).
rewrite (Nat.mul_comm _ n) in Hlookup.
unfold sublist_lookup in *; simplify_option_eq;
[|by rewrite !lookup_ge_None_2 by auto].
apply (f_equal (.!! i `mod` n)) in Hlookup.
by rewrite !lookup_take, !lookup_drop, <-!Nat.div_mod in Hlookup
by (auto using Nat.mod_upper_bound with lia).
Qed.
Lemma sublist_eq_same_length l k j n :
length l = j * n length k = j * n
( i,i < j sublist_lookup (i * n) n l = sublist_lookup (i * n) n k) l = k.
Proof.
intros Hl Hk ?. destruct (decide (n = 0)) as [->|].
{ by rewrite (nil_length_inv l), (nil_length_inv k) by lia. }
apply sublist_eq with n; [by exists j|by exists j|].
intros i. destruct (decide (i < j)); [by auto|].
assert ( m, m = j * n m < i * n + n).
{ intros ? ->. replace (i * n + n) with (S i * n) by lia.
apply Nat.mul_lt_mono_pos_r; lia. }
by rewrite !sublist_lookup_None by auto.
Qed.
Lemma sublist_lookup_reshape l i n m :
0 < n length l = m * n
reshape (replicate m n) l !! i = sublist_lookup (i * n) n l.
Proof.
intros Hn Hl. unfold sublist_lookup. apply option_eq; intros x; split.
- intros Hx. case_guard as Hi; simplify_eq/=.
{ f_equal. clear Hi. revert i l Hl Hx.
induction m as [|m IH]; intros [|i] l ??; simplify_eq/=; auto.
rewrite <-drop_drop. apply IH; rewrite ?length_drop; auto with lia. }
destruct Hi. rewrite Hl, <-Nat.mul_succ_l.
apply Nat.mul_le_mono_r, Nat.le_succ_l. apply lookup_lt_Some in Hx.
by rewrite length_reshape, length_replicate in Hx.
- intros Hx. case_guard as Hi; simplify_eq/=.
revert i l Hl Hi. induction m as [|m IH]; [auto with lia|].
intros [|i] l ??; simpl; [done|]. rewrite <-drop_drop.
rewrite IH; rewrite ?length_drop; auto with lia.
Qed.
Lemma sublist_lookup_compose l1 l2 l3 i n j m :
sublist_lookup i n l1 = Some l2 sublist_lookup j m l2 = Some l3
sublist_lookup (i + j) m l1 = Some l3.
Proof.
unfold sublist_lookup; intros; simplify_option_eq;
repeat match goal with
| H : _ length _ |- _ => rewrite length_take, length_drop in H
end; rewrite ?take_drop_commute, ?drop_drop, ?take_take,
?Nat.min_l, Nat.add_assoc by lia; auto with lia.
Qed.
Lemma length_sublist_alter f l i n k :
sublist_lookup i n l = Some k length (f k) = n
length (sublist_alter f i n l) = length l.
Proof.
unfold sublist_alter, sublist_lookup. intros Hk ?; simplify_option_eq.
rewrite !length_app, Hk, !length_take, !length_drop; lia.
Qed.
Lemma sublist_lookup_alter f l i n k :
sublist_lookup i n l = Some k length (f k) = n
sublist_lookup i n (sublist_alter f i n l) = f <$> sublist_lookup i n l.
Proof.
unfold sublist_lookup. intros Hk ?. erewrite length_sublist_alter by eauto.
unfold sublist_alter; simplify_option_eq.
by rewrite Hk, drop_app_length', take_app_length' by (rewrite ?length_take; lia).
Qed.
Lemma sublist_lookup_alter_ne f l i j n k :
sublist_lookup j n l = Some k length (f k) = n i + n j j + n i
sublist_lookup i n (sublist_alter f j n l) = sublist_lookup i n l.
Proof.
unfold sublist_lookup. intros Hk Hi ?. erewrite length_sublist_alter by eauto.
unfold sublist_alter; simplify_option_eq; f_equal; rewrite Hk.
apply list_eq; intros ii.
destruct (decide (ii < length (f k))); [|by rewrite !lookup_take_ge by lia].
rewrite !lookup_take, !lookup_drop by done. destruct (decide (i + ii < j)).
{ by rewrite lookup_app_l, lookup_take by (rewrite ?length_take; lia). }
rewrite lookup_app_r by (rewrite length_take; lia).
rewrite length_take_le, lookup_app_r, lookup_drop by lia. f_equal; lia.
Qed.
Lemma sublist_alter_all f l n : length l = n sublist_alter f 0 n l = f l.
Proof.
intros <-. unfold sublist_alter; simpl.
by rewrite drop_all, (right_id_L [] (++)), take_ge.
Qed.
Lemma sublist_alter_compose f g l i n k :
sublist_lookup i n l = Some k length (f k) = n length (g k) = n
sublist_alter (f g) i n l = sublist_alter f i n (sublist_alter g i n l).
Proof.
unfold sublist_alter, sublist_lookup. intros Hk ??; simplify_option_eq.
by rewrite !take_app_length', drop_app_length', !(assoc_L (++)), drop_app_length',
take_app_length' by (rewrite ?length_app, ?length_take, ?Hk; lia).
Qed.
Lemma Forall_sublist_lookup P l i n k :
sublist_lookup i n l = Some k Forall P l Forall P k.
Proof.
unfold sublist_lookup. intros; simplify_option_eq.
auto using Forall_take, Forall_drop.
Qed.
Lemma Forall_sublist_alter P f l i n k :
Forall P l sublist_lookup i n l = Some k Forall P (f k)
Forall P (sublist_alter f i n l).
Proof.
unfold sublist_alter, sublist_lookup. intros; simplify_option_eq.
auto using Forall_app_2, Forall_drop, Forall_take.
Qed.
Lemma Forall_sublist_alter_inv P f l i n k :
sublist_lookup i n l = Some k
Forall P (sublist_alter f i n l) Forall P (f k).
Proof.
unfold sublist_alter, sublist_lookup. intros ?; simplify_option_eq.
rewrite !Forall_app; tauto.
Qed.
End general_properties.
Lemma zip_with_sublist_alter {A B} (f : A B A) g l k i n l' k' :
length l = length k
sublist_lookup i n l = Some l' sublist_lookup i n k = Some k'
length (g l') = length k' zip_with f (g l') k' = g (zip_with f l' k')
zip_with f (sublist_alter g i n l) k = sublist_alter g i n (zip_with f l k).
Proof.
unfold sublist_lookup, sublist_alter. intros Hlen; rewrite Hlen.
intros ?? Hl' Hk'. simplify_option_eq.
by rewrite !zip_with_app_l, !zip_with_drop, Hl', drop_drop, !zip_with_take,
!length_take_le, Hk' by (rewrite ?length_drop; auto with lia).
Qed.
(** Interaction of [Forall2] with the above operations (needs to be outside the
section since the operations are used at different types). *)
Section Forall2.
Context {A B} (P : A B Prop).
Implicit Types x : A.
Implicit Types y : B.
Implicit Types l : list A.
Implicit Types k : list B.
Lemma Forall2_resize l (k : list B) x (y : B) n :
P x y Forall2 P l k Forall2 P (resize n x l) (resize n y k).
Proof.
intros. rewrite !resize_spec, (Forall2_length P l k) by done.
auto using Forall2_app, Forall2_take, Forall2_replicate.
Qed.
Lemma Forall2_resize_l l k x y n m :
P x y Forall (flip P y) l
Forall2 P (resize n x l) k Forall2 P (resize m x l) (resize m y k).
Proof.
intros. destruct (decide (m n)).
{ rewrite <-(resize_resize l m n) by done. by apply Forall2_resize. }
intros. assert (n = length k); subst.
{ by rewrite <-(Forall2_length P (resize n x l) k), length_resize. }
rewrite (Nat.le_add_sub (length k) m), !resize_add,
resize_all, drop_all, resize_nil by lia.
auto using Forall2_app, Forall2_replicate_r,
Forall_resize, Forall_drop, length_resize.
Qed.
Lemma Forall2_resize_r l k x y n m :
P x y Forall (P x) k
Forall2 P l (resize n y k) Forall2 P (resize m x l) (resize m y k).
Proof.
intros. destruct (decide (m n)).
{ rewrite <-(resize_resize k m n) by done. by apply Forall2_resize. }
assert (n = length l); subst.
{ by rewrite (Forall2_length P l (resize n y k)), length_resize. }
rewrite (Nat.le_add_sub (length l) m), !resize_add,
resize_all, drop_all, resize_nil by lia.
auto using Forall2_app, Forall2_replicate_l,
Forall_resize, Forall_drop, length_resize.
Qed.
Lemma Forall2_resize_r_flip l k x y n m :
P x y Forall (P x) k
length k = m Forall2 P l (resize n y k) Forall2 P (resize m x l) k.
Proof.
intros ?? <- ?. rewrite <-(resize_all k y) at 2.
apply Forall2_resize_r with n; auto using Forall_true.
Qed.
Lemma Forall2_rotate n l k :
Forall2 P l k Forall2 P (rotate n l) (rotate n k).
Proof.
intros HAll. unfold rotate. rewrite (Forall2_length _ _ _ HAll).
eauto using Forall2_app, Forall2_take, Forall2_drop.
Qed.
Lemma Forall2_rotate_take s e l k :
Forall2 P l k Forall2 P (rotate_take s e l) (rotate_take s e k).
Proof.
intros HAll. unfold rotate_take. rewrite (Forall2_length _ _ _ HAll).
eauto using Forall2_take, Forall2_rotate.
Qed.
Lemma Forall2_sublist_lookup_l l k n i l' :
Forall2 P l k sublist_lookup n i l = Some l'
k', sublist_lookup n i k = Some k' Forall2 P l' k'.
Proof.
unfold sublist_lookup. intros Hlk Hl.
exists (take i (drop n k)); simplify_option_eq.
- auto using Forall2_take, Forall2_drop.
- apply Forall2_length in Hlk; lia.
Qed.
Lemma Forall2_sublist_lookup_r l k n i k' :
Forall2 P l k sublist_lookup n i k = Some k'
l', sublist_lookup n i l = Some l' Forall2 P l' k'.
Proof.
intro. unfold sublist_lookup.
erewrite (Forall2_length P) by eauto; intros; simplify_option_eq.
eauto using Forall2_take, Forall2_drop.
Qed.
Lemma Forall2_sublist_alter f g l k i n l' k' :
Forall2 P l k sublist_lookup i n l = Some l'
sublist_lookup i n k = Some k' Forall2 P (f l') (g k')
Forall2 P (sublist_alter f i n l) (sublist_alter g i n k).
Proof.
intro. unfold sublist_alter, sublist_lookup.
erewrite Forall2_length by eauto; intros; simplify_option_eq.
auto using Forall2_app, Forall2_drop, Forall2_take.
Qed.
Lemma Forall2_sublist_alter_l f l k i n l' k' :
Forall2 P l k sublist_lookup i n l = Some l'
sublist_lookup i n k = Some k' Forall2 P (f l') k'
Forall2 P (sublist_alter f i n l) k.
Proof.
intro. unfold sublist_lookup, sublist_alter.
erewrite <-(Forall2_length P) by eauto; intros; simplify_option_eq.
apply Forall2_app_l;
rewrite ?length_take_le by lia; auto using Forall2_take.
apply Forall2_app_l; erewrite Forall2_length, length_take,
length_drop, <-Forall2_length, Nat.min_l by eauto with lia; [done|].
rewrite drop_drop; auto using Forall2_drop.
Qed.
End Forall2.
Section Forall2_proper.
Context {A} (R : relation A).
Global Instance: n, Proper (R ==> Forall2 R ==> Forall2 R) (resize n).
Proof. repeat intro. eauto using Forall2_resize. Qed.
Global Instance resize_proper `{!Equiv A} n : Proper (() ==> () ==> (≡@{list A})) (resize n).
Proof.
induction n; destruct 2; simpl; repeat (constructor || f_equiv); auto.
Qed.
Global Instance : n, Proper (Forall2 R ==> Forall2 R) (rotate n).
Proof. repeat intro. eauto using Forall2_rotate. Qed.
Global Instance rotate_proper `{!Equiv A} n : Proper ((≡@{list A}) ==> ()) (rotate n).
Proof. intros ??. rewrite !list_equiv_Forall2. by apply Forall2_rotate. Qed.
Global Instance: s e, Proper (Forall2 R ==> Forall2 R) (rotate_take s e).
Proof. repeat intro. eauto using Forall2_rotate_take. Qed.
Global Instance rotate_take_proper `{!Equiv A} s e : Proper ((≡@{list A}) ==> ()) (rotate_take s e).
Proof. intros ??. rewrite !list_equiv_Forall2. by apply Forall2_rotate_take. Qed.
End Forall2_proper.
Section more_general_properties.
Context {A : Type}.
Implicit Types x y z : A.
Implicit Types l k : list A.
(** ** Properties of the [mask] function *)
Lemma mask_nil f βs : mask f βs [] =@{list A} [].
Proof. by destruct βs. Qed.
Lemma length_mask f βs l : length (mask f βs l) = length l.
Proof. revert βs. induction l; intros [|??]; f_equal/=; auto. Qed.
Lemma mask_true f l n : length l n mask f (replicate n true) l = f <$> l.
Proof. revert n. induction l; intros [|?] ?; f_equal/=; auto with lia. Qed.
Lemma mask_false f l n : mask f (replicate n false) l = l.
Proof. revert l. induction n; intros [|??]; f_equal/=; auto. Qed.
Lemma mask_app f βs1 βs2 l :
mask f (βs1 ++ βs2) l
= mask f βs1 (take (length βs1) l) ++ mask f βs2 (drop (length βs1) l).
Proof. revert l. induction βs1;intros [|??]; f_equal/=; auto using mask_nil. Qed.
Lemma mask_app_2 f βs l1 l2 :
mask f βs (l1 ++ l2)
= mask f (take (length l1) βs) l1 ++ mask f (drop (length l1) βs) l2.
Proof. revert βs. induction l1; intros [|??]; f_equal/=; auto. Qed.
Lemma take_mask f βs l n : take n (mask f βs l) = mask f (take n βs) (take n l).
Proof. revert n βs. induction l; intros [|?] [|[] ?]; f_equal/=; auto. Qed.
Lemma drop_mask f βs l n : drop n (mask f βs l) = mask f (drop n βs) (drop n l).
Proof.
revert n βs. induction l; intros [|?] [|[] ?]; f_equal/=; auto using mask_nil.
Qed.
Lemma sublist_lookup_mask f βs l i n :
sublist_lookup i n (mask f βs l)
= mask f (take n (drop i βs)) <$> sublist_lookup i n l.
Proof.
unfold sublist_lookup; rewrite length_mask; simplify_option_eq; auto.
by rewrite drop_mask, take_mask.
Qed.
Lemma mask_mask f g βs1 βs2 l :
( x, f (g x) = f x) βs1 =.>* βs2
mask f βs2 (mask g βs1 l) = mask f βs2 l.
Proof.
intros ? Hβs. revert l. by induction Hβs as [|[] []]; intros [|??]; f_equal/=.
Qed.
Lemma lookup_mask f βs l i :
βs !! i = Some true mask f βs l !! i = f <$> l !! i.
Proof.
revert i βs. induction l; intros [] [] ?; simplify_eq/=; f_equal; auto.
Qed.
Lemma lookup_mask_notin f βs l i :
βs !! i Some true mask f βs l !! i = l !! i.
Proof.
revert i βs. induction l; intros [] [|[]] ?; simplify_eq/=; auto.
Qed.
End more_general_properties.
(** Lemmas about [positives_flatten] and [positives_unflatten]. *)
Section positives_flatten_unflatten.
Local Open Scope positive_scope.
Lemma positives_flatten_go_app xs acc :
positives_flatten_go xs acc = acc ++ positives_flatten_go xs 1.
Proof.
revert acc.
induction xs as [|x xs IH]; intros acc; simpl.
- reflexivity.
- rewrite IH.
rewrite (IH (6 ++ _)).
rewrite 2!(assoc_L (++)).
reflexivity.
Qed.
Lemma positives_unflatten_go_app p suffix xs acc :
positives_unflatten_go (suffix ++ Pos.reverse (Pos.dup p)) xs acc =
positives_unflatten_go suffix xs (acc ++ p).
Proof.
revert suffix acc.
induction p as [p IH|p IH|]; intros acc suffix; simpl.
- rewrite 2!Pos.reverse_xI.
rewrite 2!(assoc_L (++)).
rewrite IH.
reflexivity.
- rewrite 2!Pos.reverse_xO.
rewrite 2!(assoc_L (++)).
rewrite IH.
reflexivity.
- reflexivity.
Qed.
Lemma positives_unflatten_flatten_go suffix xs acc :
positives_unflatten_go (suffix ++ positives_flatten_go xs 1) acc 1 =
positives_unflatten_go suffix (xs ++ acc) 1.
Proof.
revert suffix acc.
induction xs as [|x xs IH]; intros suffix acc; simpl.
- reflexivity.
- rewrite positives_flatten_go_app.
rewrite (assoc_L (++)).
rewrite IH.
rewrite (assoc_L (++)).
rewrite positives_unflatten_go_app.
simpl.
rewrite (left_id_L 1 (++)).
reflexivity.
Qed.
Lemma positives_unflatten_flatten xs :
positives_unflatten (positives_flatten xs) = Some xs.
Proof.
unfold positives_flatten, positives_unflatten.
replace (positives_flatten_go xs 1)
with (1 ++ positives_flatten_go xs 1)
by apply (left_id_L 1 (++)).
rewrite positives_unflatten_flatten_go.
simpl.
rewrite (right_id_L [] (++)%list).
reflexivity.
Qed.
Lemma positives_flatten_app xs ys :
positives_flatten (xs ++ ys) = positives_flatten xs ++ positives_flatten ys.
Proof.
unfold positives_flatten.
revert ys.
induction xs as [|x xs IH]; intros ys; simpl.
- rewrite (left_id_L 1 (++)).
reflexivity.
- rewrite positives_flatten_go_app, (positives_flatten_go_app xs).
rewrite IH.
rewrite (assoc_L (++)).
reflexivity.
Qed.
Lemma positives_flatten_cons x xs :
positives_flatten (x :: xs)
= 1~1~0 ++ Pos.reverse (Pos.dup x) ++ positives_flatten xs.
Proof.
change (x :: xs) with ([x] ++ xs)%list.
rewrite positives_flatten_app.
rewrite (assoc_L (++)).
reflexivity.
Qed.
Lemma positives_flatten_suffix (l k : list positive) :
l `suffix_of` k q, positives_flatten k = q ++ positives_flatten l.
Proof.
intros [l' ->].
exists (positives_flatten l').
apply positives_flatten_app.
Qed.
Lemma positives_flatten_suffix_eq p1 p2 (xs ys : list positive) :
length xs = length ys
p1 ++ positives_flatten xs = p2 ++ positives_flatten ys
xs = ys.
Proof.
revert p1 p2 ys; induction xs as [|x xs IH];
intros p1 p2 [|y ys] ?; simplify_eq/=; auto.
rewrite !positives_flatten_cons, !(assoc _); intros Hl.
assert (xs = ys) as <- by eauto; clear IH; f_equal.
apply (inj (.++ positives_flatten xs)) in Hl.
rewrite 2!Pos.reverse_dup in Hl.
apply (Pos.dup_suffix_eq _ _ p1 p2) in Hl.
by apply (inj Pos.reverse).
Qed.
End positives_flatten_unflatten.