Skip to content
Snippets Groups Projects
Commit 376e5e05 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Merge branch 'robbert/multiset_solver' into 'master'

Add solver `multiset_solver` for multisets

See merge request iris/stdpp!167
parents 6ee0e09e 45b64e21
No related branches found
No related tags found
No related merge requests found
...@@ -20,6 +20,7 @@ API-breaking change is listed. ...@@ -20,6 +20,7 @@ API-breaking change is listed.
- Rename `drop_insert` into `drop_insert_gt` and add `drop_insert_le`. - Rename `drop_insert` into `drop_insert_gt` and add `drop_insert_le`.
- Added `Countable` instance for `Ascii.ascii`. - Added `Countable` instance for `Ascii.ascii`.
- Make lemma `list_find_Some` more apply friendly. - Make lemma `list_find_Some` more apply friendly.
- Add tactic `multiset_solver` for solving goals involving multisets.
## std++ 1.3 (released 2020-03-18) ## std++ 1.3 (released 2020-03-18)
......
From stdpp Require Import gmultiset.
Section test.
Context `{Countable A}.
Implicit Types x y : A.
Implicit Types X Y : gmultiset A.
Lemma test1 x y X : {[x]} ({[y]} X) ∅.
Proof. multiset_solver. Qed.
Lemma test2 x y X : {[x]} ({[y]} X) = {[y]} ({[x]} X).
Proof. multiset_solver. Qed.
Lemma test3 x : {[x]} ≠@{gmultiset A} ∅.
Proof. multiset_solver. Qed.
Lemma test4 x y z X Y :
{[z]} X = {[y]} Y
{[x]} ({[z]} X) = {[y]} ({[x]} Y).
Proof. multiset_solver. Qed.
Lemma test5 X x : X = {[x]} X ≠@{gmultiset A} ∅.
Proof. multiset_solver. Qed.
End test.
From stdpp Require Export countable.
From stdpp Require Import gmap. From stdpp Require Import gmap.
Set Default Proof Using "Type". Set Default Proof Using "Type".
...@@ -57,482 +58,531 @@ Typeclasses Opaque gmultiset_elements gmultiset_size gmultiset_empty. ...@@ -57,482 +58,531 @@ Typeclasses Opaque gmultiset_elements gmultiset_size gmultiset_empty.
Typeclasses Opaque gmultiset_singleton gmultiset_union gmultiset_difference. Typeclasses Opaque gmultiset_singleton gmultiset_union gmultiset_difference.
Typeclasses Opaque gmultiset_dom. Typeclasses Opaque gmultiset_dom.
Section lemmas. Section basic_lemmas.
Context `{Countable A}. Context `{Countable A}.
Implicit Types x y : A. Implicit Types x y : A.
Implicit Types X Y : gmultiset A. Implicit Types X Y : gmultiset A.
Lemma gmultiset_eq X Y : X = Y x, multiplicity x X = multiplicity x Y. Lemma gmultiset_eq X Y : X = Y x, multiplicity x X = multiplicity x Y.
Proof. Proof.
split; [by intros ->|intros HXY]. split; [by intros ->|intros HXY].
destruct X as [X], Y as [Y]; f_equal; apply map_eq; intros x. destruct X as [X], Y as [Y]; f_equal; apply map_eq; intros x.
specialize (HXY x); unfold multiplicity in *; simpl in *. specialize (HXY x); unfold multiplicity in *; simpl in *.
repeat case_match; naive_solver. repeat case_match; naive_solver.
Qed. Qed.
Global Instance gmultiset_leibniz : LeibnizEquiv (gmultiset A). Global Instance gmultiset_leibniz : LeibnizEquiv (gmultiset A).
Proof. intros X Y. by rewrite gmultiset_eq. Qed. Proof. intros X Y. by rewrite gmultiset_eq. Qed.
Global Instance gmultiset_equiv_equivalence : Equivalence (≡@{gmultiset A}). Global Instance gmultiset_equiv_equivalence : Equivalence (≡@{gmultiset A}).
Proof. constructor; repeat intro; naive_solver. Qed. Proof. constructor; repeat intro; naive_solver. Qed.
(* Multiplicity *) (* Multiplicity *)
Lemma multiplicity_empty x : multiplicity x = 0. Lemma multiplicity_empty x : multiplicity x = 0.
Proof. done. Qed. 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. 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. Proof. intros. unfold multiplicity; simpl. by rewrite lookup_singleton_ne. Qed.
Lemma multiplicity_singleton' x y : 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. Proof.
destruct (decide _) as [->|]. destruct (decide _) as [->|].
- by rewrite multiplicity_singleton. - by rewrite multiplicity_singleton.
- by rewrite multiplicity_singleton_ne. - by rewrite multiplicity_singleton_ne.
Qed. Qed.
Lemma multiplicity_union X Y x : Lemma multiplicity_union X Y x :
multiplicity x (X Y) = multiplicity x X `max` multiplicity x Y. multiplicity x (X Y) = multiplicity x X `max` multiplicity x Y.
Proof. Proof.
destruct X as [X], Y as [Y]; unfold multiplicity; simpl. destruct X as [X], Y as [Y]; unfold multiplicity; simpl.
rewrite lookup_union_with. destruct (X !! _), (Y !! _); simpl; lia. rewrite lookup_union_with. destruct (X !! _), (Y !! _); simpl; lia.
Qed. Qed.
Lemma multiplicity_intersection X Y x : Lemma multiplicity_intersection X Y x :
multiplicity x (X Y) = multiplicity x X `min` multiplicity x Y. multiplicity x (X Y) = multiplicity x X `min` multiplicity x Y.
Proof. Proof.
destruct X as [X], Y as [Y]; unfold multiplicity; simpl. destruct X as [X], Y as [Y]; unfold multiplicity; simpl.
rewrite lookup_intersection_with. destruct (X !! _), (Y !! _); simpl; lia. rewrite lookup_intersection_with. destruct (X !! _), (Y !! _); simpl; lia.
Qed. Qed.
Lemma multiplicity_disj_union X Y x : Lemma multiplicity_disj_union X Y x :
multiplicity x (X Y) = multiplicity x X + multiplicity x Y. multiplicity x (X Y) = multiplicity x X + multiplicity x Y.
Proof. Proof.
destruct X as [X], Y as [Y]; unfold multiplicity; simpl. destruct X as [X], Y as [Y]; unfold multiplicity; simpl.
rewrite lookup_union_with. destruct (X !! _), (Y !! _); simpl; lia. rewrite lookup_union_with. destruct (X !! _), (Y !! _); simpl; lia.
Qed. Qed.
Lemma multiplicity_difference X Y x : Lemma multiplicity_difference X Y x :
multiplicity x (X Y) = multiplicity x X - multiplicity x Y. multiplicity x (X Y) = multiplicity x X - multiplicity x Y.
Proof. Proof.
destruct X as [X], Y as [Y]; unfold multiplicity; simpl. destruct X as [X], Y as [Y]; unfold multiplicity; simpl.
rewrite lookup_difference_with. rewrite lookup_difference_with.
destruct (X !! _), (Y !! _); simplify_option_eq; lia. destruct (X !! _), (Y !! _); simplify_option_eq; lia.
Qed. Qed.
(* Set_ *) (* Set *)
Lemma elem_of_multiplicity x X : x X 0 < multiplicity x X. Lemma elem_of_multiplicity x X : x X 0 < multiplicity x X.
Proof. done. Qed. Proof. done. Qed.
Global Instance gmultiset_simple_set : SemiSet A (gmultiset A). Global Instance gmultiset_simple_set : SemiSet A (gmultiset A).
Proof. Proof.
split. split.
- intros x. rewrite elem_of_multiplicity, multiplicity_empty. lia. - intros x. rewrite elem_of_multiplicity, multiplicity_empty. lia.
- intros x y. - intros x y.
rewrite elem_of_multiplicity, multiplicity_singleton'. rewrite elem_of_multiplicity, multiplicity_singleton'.
destruct (decide (x = y)); intuition lia. destruct (decide (x = y)); intuition lia.
- intros X Y x. rewrite !elem_of_multiplicity, multiplicity_union. lia. - intros X Y x. rewrite !elem_of_multiplicity, multiplicity_union. lia.
Qed. Qed.
Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}). Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}).
Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined. Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined.
Lemma gmultiset_elem_of_disj_union X Y x : x X Y x X x Y. 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. Proof. rewrite !elem_of_multiplicity, multiplicity_disj_union. lia. Qed.
Global Instance set_unfold_gmultiset_disj_union x X Y P Q : Global Instance set_unfold_gmultiset_disj_union x X Y P Q :
SetUnfoldElemOf x X P SetUnfoldElemOf x Y Q SetUnfoldElemOf x X P SetUnfoldElemOf x Y Q
SetUnfoldElemOf x (X Y) (P Q). SetUnfoldElemOf x (X Y) (P Q).
Proof. Proof.
intros ??; constructor. rewrite gmultiset_elem_of_disj_union. intros ??; constructor. rewrite gmultiset_elem_of_disj_union.
by rewrite <-(set_unfold_elem_of x X P), <-(set_unfold_elem_of x Y Q). by rewrite <-(set_unfold_elem_of x X P), <-(set_unfold_elem_of x Y Q).
Qed. Qed.
End basic_lemmas.
(* Algebraic laws *)
(** For union *) (** * A solver for multisets *)
Global Instance gmultiset_union_comm : Comm (=@{gmultiset A}) (). (** We define a tactic [multiset_solver] that solves goals involving multisets.
Proof. The strategy of this tactic is as follows:
intros X Y. apply gmultiset_eq; intros x. rewrite !multiplicity_union; lia.
Qed. 1. Unfold all equalities ([=]), equivalences ([≡]), and inclusions ([⊆]) using
Global Instance gmultiset_union_assoc : Assoc (=@{gmultiset A}) (). the laws of [multiplicity] for the multiset operations. Note that strict
Proof. inclusion ([⊂]) is not supported.
intros X Y Z. apply gmultiset_eq; intros x. rewrite !multiplicity_union; lia. 2. Use [naive_solver] to decompose the goal into smaller subgoals.
Qed. 3. Instantiate all universally quantified hypotheses in the subgoals generated
Global Instance gmultiset_union_left_id : LeftId (=@{gmultiset A}) (). by [naive_solver] to obtain goals that can be solved using [lia].
Proof.
intros X. apply gmultiset_eq; intros x. Step (1) is implemented using a type class [MultisetUnfold] that hooks into
by rewrite multiplicity_union, multiplicity_empty. the [SetUnfold] mechanism of [set_solver]. Since [SetUnfold] already propagates
Qed. through logical connectives, we obtain the same behavior for our multiset
Global Instance gmultiset_union_right_id : RightId (=@{gmultiset A}) (). solver. Note that no [MultisetUnfold] instance is defined for the (non-trivial)
Proof. intros X. by rewrite (comm_L ()), (left_id_L _ _). Qed. singleton [{[ y ]}] since the singleton receives special treatment in step (3).
Global Instance gmultiset_union_idemp : IdemP (=@{gmultiset A}) ().
Proof. Step (3) is achieved using the tactic [multiset_instantiate], which instantiates
intros X. apply gmultiset_eq; intros x. rewrite !multiplicity_union; lia. universally quantified hypotheses [H : ∀ x : A, P x] in two ways:
Qed.
- If [P] contains a multiset singleton [{[ y ]}] it adds the hypothesis [H y].
(** For intersection *) - If the goal or some hypothesis contains [multiplicity y X] it adds the
Global Instance gmultiset_intersection_comm : Comm (=@{gmultiset A}) (). hypothesis [H y].
Proof. *)
intros X Y. apply gmultiset_eq; intros x. rewrite !multiplicity_intersection; lia. Class MultisetUnfold `{Countable A} (x : A) (X : gmultiset A) (n : nat) :=
Qed. { multiset_unfold : multiplicity x X = n }.
Global Instance gmultiset_intersection_assoc : Assoc (=@{gmultiset A}) (). Arguments multiset_unfold {_ _ _} _ _ _ {_} : assert.
Proof. Hint Mode MultisetUnfold + + + - + - : typeclass_instances.
intros X Y Z. apply gmultiset_eq; intros x. rewrite !multiplicity_intersection; lia.
Qed. Section multiset_unfold.
Global Instance gmultiset_intersection_left_absorb : LeftAbsorb (=@{gmultiset A}) (). Context `{Countable A}.
Proof. Implicit Types x y : A.
intros X. apply gmultiset_eq; intros x. Implicit Types X Y : gmultiset A.
by rewrite multiplicity_intersection, multiplicity_empty.
Qed. Global Instance multiset_unfold_default x X :
Global Instance gmultiset_intersection_right_absorb : RightAbsorb (=@{gmultiset A}) (). MultisetUnfold x X (multiplicity x X) | 1000.
Proof. intros X. by rewrite (comm_L ()), (left_absorb_L _ _). Qed. Proof. done. Qed.
Global Instance gmultiset_intersection_idemp : IdemP (=@{gmultiset A}) (). Global Instance multiset_unfold_empty x : MultisetUnfold x 0.
Proof. Proof. constructor. by rewrite multiplicity_empty. Qed.
intros X. apply gmultiset_eq; intros x. rewrite !multiplicity_intersection; lia. Global Instance multiset_unfold_singleton x y :
Qed. MultisetUnfold x {[ x ]} 1.
Proof. constructor. by rewrite multiplicity_singleton. Qed.
Lemma gmultiset_union_intersection_l X Y Z : X (Y Z) = (X Y) (X Z). Global Instance multiset_unfold_union x X Y n m :
Proof. MultisetUnfold x X n MultisetUnfold x Y m
apply gmultiset_eq; intros y. MultisetUnfold x (X Y) (n `max` m).
rewrite multiplicity_union, !multiplicity_intersection, !multiplicity_union. lia. Proof. intros [HX] [HY]; constructor. by rewrite multiplicity_union, HX, HY. Qed.
Qed. Global Instance multiset_unfold_intersection x X Y n m :
Lemma gmultiset_union_intersection_r X Y Z : (X Y) Z = (X Z) (Y Z). MultisetUnfold x X n MultisetUnfold x Y m
Proof. by rewrite <-!(comm_L _ Z), gmultiset_union_intersection_l. Qed. MultisetUnfold x (X Y) (n `min` m).
Lemma gmultiset_intersection_union_l X Y Z : X (Y Z) = (X Y) (X Z). Proof. intros [HX] [HY]; constructor. by rewrite multiplicity_intersection, HX, HY. Qed.
Proof. Global Instance multiset_unfold_disj_union x X Y n m :
apply gmultiset_eq; intros y. MultisetUnfold x X n MultisetUnfold x Y m
rewrite multiplicity_union, !multiplicity_intersection, !multiplicity_union. lia. MultisetUnfold x (X Y) (n + m).
Qed. Proof. intros [HX] [HY]; constructor. by rewrite multiplicity_disj_union, HX, HY. Qed.
Lemma gmultiset_intersection_union_r X Y Z : (X Y) Z = (X Z) (Y Z). Global Instance multiset_unfold_difference x X Y n m :
Proof. by rewrite <-!(comm_L _ Z), gmultiset_intersection_union_l. Qed. MultisetUnfold x X n MultisetUnfold x Y m
MultisetUnfold x (X Y) (n - m).
(** For disjoint union (aka sum) *) Proof. intros [HX] [HY]; constructor. by rewrite multiplicity_difference, HX, HY. Qed.
Global Instance gmultiset_disj_union_comm : Comm (=@{gmultiset A}) ().
Proof. Global Instance set_unfold_multiset_equiv X Y f g :
intros X Y. apply gmultiset_eq; intros x. rewrite !multiplicity_disj_union; lia. ( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
Qed. SetUnfold (X Y) ( x, f x = g x).
Global Instance gmultiset_disj_union_assoc : Assoc (=@{gmultiset A}) (). Proof.
Proof. constructor. apply forall_proper; intros x.
intros X Y Z. apply gmultiset_eq; intros x. rewrite !multiplicity_disj_union; lia. by rewrite (multiset_unfold x X (f x)), (multiset_unfold x Y (g x)).
Qed. Qed.
Global Instance gmultiset_disj_union_left_id : LeftId (=@{gmultiset A}) (). Global Instance set_unfold_multiset_eq X Y f g :
Proof. ( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
intros X. apply gmultiset_eq; intros x. SetUnfold (X = Y) ( x, f x = g x).
by rewrite multiplicity_disj_union, multiplicity_empty. Proof. constructor. unfold_leibniz. by apply set_unfold_multiset_equiv. Qed.
Qed. Global Instance set_unfold_multiset_subseteq X Y f g :
Global Instance gmultiset_disj_union_right_id : RightId (=@{gmultiset A}) (). ( x, MultisetUnfold x X (f x)) ( x, MultisetUnfold x Y (g x))
Proof. intros X. by rewrite (comm_L ()), (left_id_L _ _). Qed. SetUnfold (X Y) ( x, f x g x).
Proof.
Global Instance gmultiset_disj_union_inj_1 X : Inj (=) (=) (X ⊎.). constructor. apply forall_proper; intros x.
Proof. by rewrite (multiset_unfold x X (f x)), (multiset_unfold x Y (g x)).
intros Y1 Y2. rewrite !gmultiset_eq. intros HX x; generalize (HX x). Qed.
rewrite !multiplicity_disj_union. lia. End multiset_unfold.
Qed.
Global Instance gmultiset_disj_union_inj_2 X : Inj (=) (=) (. X). Ltac multiset_instantiate :=
Proof. intros Y1 Y2. rewrite <-!(comm_L _ X). apply (inj (X ⊎.)). Qed. (* Step 3.1: instantiate hypotheses *)
repeat match goal with
Lemma gmultiset_disj_union_intersection_l X Y Z : X (Y Z) = (X Y) (X Z). | H : ( x : ?A, @?P x) |- _ =>
Proof. let e := fresh in evar (e:A);
apply gmultiset_eq; intros y. let e' := eval unfold e in e in clear e;
rewrite multiplicity_disj_union, !multiplicity_intersection, lazymatch constr:(P e') with
!multiplicity_disj_union. lia. | context [ {[ ?y ]} ] =>
Qed. (* Use [unless] to avoid creating a new hypothesis [H y : P y] if [P y]
Lemma gmultiset_disj_union_intersection_r X Y Z : (X Y) Z = (X Z) (Y Z). already exists. *)
Proof. by rewrite <-!(comm_L _ Z), gmultiset_disj_union_intersection_l. Qed. unify y e'; unless (P y) by assumption; pose proof (H y)
end
Lemma gmultiset_disj_union_union_l X Y Z : X (Y Z) = (X Y) (X Z). | H : ( x : ?A, @?P x), _ : context [multiplicity ?y _] |- _ =>
Proof. (* Use [unless] to avoid creating a new hypothesis [H y : P y] if [P y]
apply gmultiset_eq; intros y. already exists. *)
rewrite multiplicity_disj_union, !multiplicity_union, unless (P y) by assumption; pose proof (H y)
!multiplicity_disj_union. lia. | H : ( x : ?A, @?P x) |- context [multiplicity ?y _] =>
Qed. (* Use [unless] to avoid creating a new hypothesis [H y : P y] if [P y]
Lemma gmultiset_disj_union_union_r X Y Z : (X Y) Z = (X Z) (Y Z). already exists. *)
Proof. by rewrite <-!(comm_L _ Z), gmultiset_disj_union_union_l. Qed. unless (P y) by assumption; pose proof (H y)
end;
(** Misc *) (* Step 3.2: simplify singletons. *)
Lemma gmultiset_non_empty_singleton x : {[ x ]} ≠@{gmultiset A} ∅. (* Note that we do not use [repeat case_decide] since that leads to an
Proof. exponential explosion in the number of singletons. *)
rewrite gmultiset_eq. intros Hx; generalize (Hx x). repeat match goal with
by rewrite multiplicity_singleton, multiplicity_empty. | H : context [multiplicity _ {[ _ ]}] |- _ =>
Qed. progress (rewrite ?multiplicity_singleton, ?multiplicity_singleton_ne in H by done)
| |- context [multiplicity _ {[ _ ]}] =>
(** Conversion from lists *) progress (rewrite ?multiplicity_singleton, ?multiplicity_singleton_ne by done)
Lemma list_to_set_disj_nil : list_to_set_disj [] =@{gmultiset A} ∅. end.
Proof. done. Qed.
Lemma list_to_set_disj_cons x l : Ltac multiset_solver := set_solver by (multiset_instantiate; lia).
list_to_set_disj (x :: l) =@{gmultiset A} {[ x ]} list_to_set_disj l.
Proof. done. Qed. Section more_lemmas.
Lemma list_to_set_disj_app l1 l2 : Context `{Countable A}.
list_to_set_disj (l1 ++ l2) =@{gmultiset A} list_to_set_disj l1 list_to_set_disj l2. Implicit Types x y : A.
Proof. Implicit Types X Y : gmultiset A.
induction l1 as [|x l1 IH]; simpl.
- by rewrite (left_id_L _ _). (* Algebraic laws *)
- by rewrite IH, (assoc_L _). (** For union *)
Qed. Global Instance gmultiset_union_comm : Comm (=@{gmultiset A}) ().
Global Instance list_to_set_disj_perm : Proof. unfold Comm. multiset_solver. Qed.
Proper (() ==> (=)) (list_to_set_disj (C:=gmultiset A)). Global Instance gmultiset_union_assoc : Assoc (=@{gmultiset A}) ().
Proof. Proof. unfold Assoc. multiset_solver. Qed.
induction 1 as [|x l l' _ IH|x y l|l l' l'' _ IH1 _ IH2]; simpl; auto. Global Instance gmultiset_union_left_id : LeftId (=@{gmultiset A}) ().
- by rewrite IH. Proof. unfold LeftId. multiset_solver. Qed.
- by rewrite !(assoc_L _), (comm_L _ {[ x ]}). Global Instance gmultiset_union_right_id : RightId (=@{gmultiset A}) ().
- by rewrite IH1. Proof. unfold RightId. multiset_solver. Qed.
Qed. Global Instance gmultiset_union_idemp : IdemP (=@{gmultiset A}) ().
Proof. unfold IdemP. multiset_solver. Qed.
(** Properties of the elements operation *)
Lemma gmultiset_elements_empty : elements ( : gmultiset A) = []. (** For intersection *)
Proof. Global Instance gmultiset_intersection_comm : Comm (=@{gmultiset A}) ().
unfold elements, gmultiset_elements; simpl. by rewrite map_to_list_empty. Proof. unfold Comm. multiset_solver. Qed.
Qed. Global Instance gmultiset_intersection_assoc : Assoc (=@{gmultiset A}) ().
Lemma gmultiset_elements_empty_inv X : elements X = [] X = ∅. Proof. unfold Assoc. multiset_solver. Qed.
Proof. Global Instance gmultiset_intersection_left_absorb : LeftAbsorb (=@{gmultiset A}) ().
destruct X as [X]; unfold elements, gmultiset_elements; simpl. Proof. unfold LeftAbsorb. multiset_solver. Qed.
intros; apply (f_equal GMultiSet). destruct (map_to_list X) as [|[]] eqn:?. Global Instance gmultiset_intersection_right_absorb : RightAbsorb (=@{gmultiset A}) ().
- by apply map_to_list_empty_inv. Proof. unfold RightAbsorb. multiset_solver. Qed.
- naive_solver. Global Instance gmultiset_intersection_idemp : IdemP (=@{gmultiset A}) ().
Qed. Proof. unfold IdemP. multiset_solver. Qed.
Lemma gmultiset_elements_empty' X : elements X = [] X = ∅.
Proof. Lemma gmultiset_union_intersection_l X Y Z : X (Y Z) = (X Y) (X Z).
split; intros HX; [by apply gmultiset_elements_empty_inv|]. Proof. multiset_solver. Qed.
by rewrite HX, gmultiset_elements_empty. Lemma gmultiset_union_intersection_r X Y Z : (X Y) Z = (X Z) (Y Z).
Qed. Proof. multiset_solver. Qed.
Lemma gmultiset_elements_singleton x : elements ({[ x ]} : gmultiset A) = [ x ]. Lemma gmultiset_intersection_union_l X Y Z : X (Y Z) = (X Y) (X Z).
Proof. Proof. multiset_solver. Qed.
unfold elements, gmultiset_elements; simpl. by rewrite map_to_list_singleton. Lemma gmultiset_intersection_union_r X Y Z : (X Y) Z = (X Z) (Y Z).
Qed. Proof. multiset_solver. Qed.
Lemma gmultiset_elements_disj_union X Y :
elements (X Y) elements X ++ elements Y. (** For disjoint union (aka sum) *)
Proof. Global Instance gmultiset_disj_union_comm : Comm (=@{gmultiset A}) ().
destruct X as [X], Y as [Y]; unfold elements, gmultiset_elements. Proof. unfold Comm. multiset_solver. Qed.
set (f xn := let '(x, n) := xn in replicate (S n) x); simpl. Global Instance gmultiset_disj_union_assoc : Assoc (=@{gmultiset A}) ().
revert Y; induction X as [|x n X HX IH] using map_ind; intros Y. Proof. unfold Assoc. multiset_solver. Qed.
{ by rewrite (left_id_L _ _ Y), map_to_list_empty. } Global Instance gmultiset_disj_union_left_id : LeftId (=@{gmultiset A}) ().
destruct (Y !! x) as [n'|] eqn:HY. Proof. unfold LeftId. multiset_solver. Qed.
- rewrite <-(insert_id Y x n'), <-(insert_delete Y) by done. Global Instance gmultiset_disj_union_right_id : RightId (=@{gmultiset A}) ().
erewrite <-insert_union_with by done. Proof. unfold RightId. multiset_solver. Qed.
rewrite !map_to_list_insert, !bind_cons
by (by rewrite ?lookup_union_with, ?lookup_delete, ?HX). Global Instance gmultiset_disj_union_inj_1 X : Inj (=) (=) (X ⊎.).
rewrite (assoc_L _), <-(comm (++) (f (_,n'))), <-!(assoc_L _), <-IH. Proof. unfold Inj. multiset_solver. Qed.
rewrite (assoc_L _). f_equiv. Global Instance gmultiset_disj_union_inj_2 X : Inj (=) (=) (. X).
rewrite (comm _); simpl. by rewrite replicate_plus, Permutation_middle. Proof. unfold Inj. multiset_solver. Qed.
- rewrite <-insert_union_with_l, !map_to_list_insert, !bind_cons
by (by rewrite ?lookup_union_with, ?HX, ?HY). Lemma gmultiset_disj_union_intersection_l X Y Z : X (Y Z) = (X Y) (X Z).
by rewrite <-(assoc_L (++)), <-IH. Proof. multiset_solver. Qed.
Qed. Lemma gmultiset_disj_union_intersection_r X Y Z : (X Y) Z = (X Z) (Y Z).
Lemma gmultiset_elem_of_elements x X : x elements X x X. Proof. multiset_solver. Qed.
Proof.
destruct X as [X]. unfold elements, gmultiset_elements. Lemma gmultiset_disj_union_union_l X Y Z : X (Y Z) = (X Y) (X Z).
set (f xn := let '(x, n) := xn in replicate (S n) x); simpl. Proof. multiset_solver. Qed.
unfold elem_of at 2, gmultiset_elem_of, multiplicity; simpl. Lemma gmultiset_disj_union_union_r X Y Z : (X Y) Z = (X Z) (Y Z).
rewrite elem_of_list_bind. split. Proof. multiset_solver. Qed.
- intros [[??] [[<- ?]%elem_of_replicate ->%elem_of_map_to_list]]; lia.
- intros. destruct (X !! x) as [n|] eqn:Hx; [|lia]. (** Misc *)
exists (x,n); split; [|by apply elem_of_map_to_list]. Lemma gmultiset_non_empty_singleton x : {[ x ]} ≠@{gmultiset A} ∅.
apply elem_of_replicate; auto with lia. Proof. multiset_solver. Qed.
Qed.
Lemma gmultiset_elem_of_dom x X : x dom (gset A) X x X. (** Conversion from lists *)
Proof. Lemma list_to_set_disj_nil : list_to_set_disj [] =@{gmultiset A} ∅.
unfold dom, gmultiset_dom, elem_of at 2, gmultiset_elem_of, multiplicity. Proof. done. Qed.
destruct X as [X]; simpl; rewrite elem_of_dom, <-not_eq_None_Some. Lemma list_to_set_disj_cons x l :
destruct (X !! x); naive_solver lia. list_to_set_disj (x :: l) =@{gmultiset A} {[ x ]} list_to_set_disj l.
Qed. Proof. done. Qed.
Lemma list_to_set_disj_app l1 l2 :
(** Properties of the set_fold operation *) list_to_set_disj (l1 ++ l2) =@{gmultiset A} list_to_set_disj l1 list_to_set_disj l2.
Lemma gmultiset_set_fold_empty {B} (f : A B B) (b : B) : Proof. induction l1; multiset_solver. Qed.
set_fold f b ( : gmultiset A) = b. Global Instance list_to_set_disj_perm :
Proof. by unfold set_fold; simpl; rewrite gmultiset_elements_empty. Qed. Proper (() ==> (=)) (list_to_set_disj (C:=gmultiset A)).
Lemma gmultiset_set_fold_singleton {B} (f : A B B) (b : B) (a : A) : Proof. induction 1; multiset_solver. Qed.
set_fold f b ({[a]} : gmultiset A) = f a b.
Proof. by unfold set_fold; simpl; rewrite gmultiset_elements_singleton. Qed. (** Properties of the elements operation *)
Lemma gmultiset_set_fold_disj_union (f : A A A) (b : A) X Y : Lemma gmultiset_elements_empty : elements ( : gmultiset A) = [].
Comm (=) f Proof.
Assoc (=) f unfold elements, gmultiset_elements; simpl. by rewrite map_to_list_empty.
set_fold f b (X Y) = set_fold f (set_fold f b X) Y. Qed.
Proof. Lemma gmultiset_elements_empty_inv X : elements X = [] X = ∅.
intros Hcomm Hassoc. unfold set_fold; simpl. Proof.
by rewrite gmultiset_elements_disj_union, <- foldr_app, (comm (++)). destruct X as [X]; unfold elements, gmultiset_elements; simpl.
Qed. intros; apply (f_equal GMultiSet). destruct (map_to_list X) as [|[]] eqn:?.
- by apply map_to_list_empty_inv.
(** Properties of the size operation *) - naive_solver.
Lemma gmultiset_size_empty : size ( : gmultiset A) = 0. Qed.
Proof. done. Qed. Lemma gmultiset_elements_empty' X : elements X = [] X = ∅.
Lemma gmultiset_size_empty_inv X : size X = 0 X = ∅. Proof.
Proof. split; intros HX; [by apply gmultiset_elements_empty_inv|].
unfold size, gmultiset_size; simpl. rewrite length_zero_iff_nil. by rewrite HX, gmultiset_elements_empty.
apply gmultiset_elements_empty_inv. Qed.
Qed. Lemma gmultiset_elements_singleton x : elements ({[ x ]} : gmultiset A) = [ x ].
Lemma gmultiset_size_empty_iff X : size X = 0 X = ∅. Proof.
Proof. unfold elements, gmultiset_elements; simpl. by rewrite map_to_list_singleton.
split; [apply gmultiset_size_empty_inv|]. Qed.
by intros ->; rewrite gmultiset_size_empty. Lemma gmultiset_elements_disj_union X Y :
Qed. elements (X Y) elements X ++ elements Y.
Lemma gmultiset_size_non_empty_iff X : size X 0 X ∅. Proof.
Proof. by rewrite gmultiset_size_empty_iff. Qed. destruct X as [X], Y as [Y]; unfold elements, gmultiset_elements.
set (f xn := let '(x, n) := xn in replicate (S n) x); simpl.
Lemma gmultiset_choose_or_empty X : ( x, x X) X = ∅. revert Y; induction X as [|x n X HX IH] using map_ind; intros Y.
Proof. { by rewrite (left_id_L _ _ Y), map_to_list_empty. }
destruct (elements X) as [|x l] eqn:HX; [right|left]. destruct (Y !! x) as [n'|] eqn:HY.
- by apply gmultiset_elements_empty_inv. - rewrite <-(insert_id Y x n'), <-(insert_delete Y) by done.
- exists x. rewrite <-gmultiset_elem_of_elements, HX. by left. erewrite <-insert_union_with by done.
Qed. rewrite !map_to_list_insert, !bind_cons
Lemma gmultiset_choose X : X x, x X. by (by rewrite ?lookup_union_with, ?lookup_delete, ?HX).
Proof. intros. by destruct (gmultiset_choose_or_empty X). Qed. rewrite (assoc_L _), <-(comm (++) (f (_,n'))), <-!(assoc_L _), <-IH.
Lemma gmultiset_size_pos_elem_of X : 0 < size X x, x X. rewrite (assoc_L _). f_equiv.
Proof. rewrite (comm _); simpl. by rewrite replicate_plus, Permutation_middle.
intros Hsz. destruct (gmultiset_choose_or_empty X) as [|HX]; [done|]. - rewrite <-insert_union_with_l, !map_to_list_insert, !bind_cons
contradict Hsz. rewrite HX, gmultiset_size_empty; lia. by (by rewrite ?lookup_union_with, ?HX, ?HY).
Qed. by rewrite <-(assoc_L (++)), <-IH.
Qed.
Lemma gmultiset_size_singleton x : size ({[ x ]} : gmultiset A) = 1. Lemma gmultiset_elem_of_elements x X : x elements X x X.
Proof. Proof.
unfold size, gmultiset_size; simpl. by rewrite gmultiset_elements_singleton. destruct X as [X]. unfold elements, gmultiset_elements.
Qed. set (f xn := let '(x, n) := xn in replicate (S n) x); simpl.
Lemma gmultiset_size_disj_union X Y : size (X Y) = size X + size Y. unfold elem_of at 2, gmultiset_elem_of, multiplicity; simpl.
Proof. rewrite elem_of_list_bind. split.
unfold size, gmultiset_size; simpl. - intros [[??] [[<- ?]%elem_of_replicate ->%elem_of_map_to_list]]; lia.
by rewrite gmultiset_elements_disj_union, app_length. - intros. destruct (X !! x) as [n|] eqn:Hx; [|lia].
Qed. exists (x,n); split; [|by apply elem_of_map_to_list].
apply elem_of_replicate; auto with lia.
(** Order stuff *) Qed.
Global Instance gmultiset_po : PartialOrder (⊆@{gmultiset A}). Lemma gmultiset_elem_of_dom x X : x dom (gset A) X x X.
Proof. Proof.
split; [split|]. unfold dom, gmultiset_dom, elem_of at 2, gmultiset_elem_of, multiplicity.
- by intros X x. destruct X as [X]; simpl; rewrite elem_of_dom, <-not_eq_None_Some.
- intros X Y Z HXY HYZ x. by trans (multiplicity x Y). destruct (X !! x); naive_solver lia.
- intros X Y HXY HYX; apply gmultiset_eq; intros x. by apply (anti_symm ()). Qed.
Qed.
(** Properties of the set_fold operation *)
Lemma gmultiset_subseteq_alt X Y : Lemma gmultiset_set_fold_empty {B} (f : A B B) (b : B) :
X Y set_fold f b ( : gmultiset A) = b.
map_relation () (λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y). Proof. by unfold set_fold; simpl; rewrite gmultiset_elements_empty. Qed.
Proof. Lemma gmultiset_set_fold_singleton {B} (f : A B B) (b : B) (a : A) :
apply forall_proper; intros x. unfold multiplicity. set_fold f b ({[a]} : gmultiset A) = f a b.
destruct (gmultiset_car X !! x), (gmultiset_car Y !! x); naive_solver lia. Proof. by unfold set_fold; simpl; rewrite gmultiset_elements_singleton. Qed.
Qed. Lemma gmultiset_set_fold_disj_union (f : A A A) (b : A) X Y :
Global Instance gmultiset_subseteq_dec : RelDecision (⊆@{gmultiset A}). Comm (=) f
Proof. Assoc (=) f
refine (λ X Y, cast_if (decide (map_relation () set_fold f b (X Y) = set_fold f (set_fold f b X) Y.
(λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y)))); Proof.
by rewrite gmultiset_subseteq_alt. intros Hcomm Hassoc. unfold set_fold; simpl.
Defined. by rewrite gmultiset_elements_disj_union, <- foldr_app, (comm (++)).
Qed.
Lemma gmultiset_subset_subseteq X Y : X Y X Y.
Proof. apply strict_include. Qed. (** Properties of the size operation *)
Hint Resolve gmultiset_subset_subseteq : core. Lemma gmultiset_size_empty : size ( : gmultiset A) = 0.
Proof. done. Qed.
Lemma gmultiset_empty_subseteq X : X. Lemma gmultiset_size_empty_inv X : size X = 0 X = ∅.
Proof. intros x. rewrite multiplicity_empty. lia. Qed. Proof.
unfold size, gmultiset_size; simpl. rewrite length_zero_iff_nil.
Lemma gmultiset_union_subseteq_l X Y : X X Y. apply gmultiset_elements_empty_inv.
Proof. intros x. rewrite multiplicity_union. lia. Qed. Qed.
Lemma gmultiset_union_subseteq_r X Y : Y X Y. Lemma gmultiset_size_empty_iff X : size X = 0 X = ∅.
Proof. intros x. rewrite multiplicity_union. lia. Qed. Proof.
Lemma gmultiset_union_mono X1 X2 Y1 Y2 : X1 X2 Y1 Y2 X1 Y1 X2 Y2. split; [apply gmultiset_size_empty_inv|].
Proof. by intros ->; rewrite gmultiset_size_empty.
intros HX HY x. rewrite !multiplicity_union. Qed.
specialize (HX x); specialize (HY x); lia. Lemma gmultiset_size_non_empty_iff X : size X 0 X ∅.
Qed. Proof. by rewrite gmultiset_size_empty_iff. Qed.
Lemma gmultiset_union_mono_l X Y1 Y2 : Y1 Y2 X Y1 X Y2.
Proof. intros. by apply gmultiset_union_mono. Qed. Lemma gmultiset_choose_or_empty X : ( x, x X) X = ∅.
Lemma gmultiset_union_mono_r X1 X2 Y : X1 X2 X1 Y X2 Y. Proof.
Proof. intros. by apply gmultiset_union_mono. Qed. destruct (elements X) as [|x l] eqn:HX; [right|left].
- by apply gmultiset_elements_empty_inv.
Lemma gmultiset_disj_union_subseteq_l X Y : X X Y. - exists x. rewrite <-gmultiset_elem_of_elements, HX. by left.
Proof. intros x. rewrite multiplicity_disj_union. lia. Qed. Qed.
Lemma gmultiset_disj_union_subseteq_r X Y : Y X Y. Lemma gmultiset_choose X : X x, x X.
Proof. intros x. rewrite multiplicity_disj_union. lia. Qed. Proof. intros. by destruct (gmultiset_choose_or_empty X). Qed.
Lemma gmultiset_disj_union_mono X1 X2 Y1 Y2 : X1 X2 Y1 Y2 X1 Y1 X2 Y2. Lemma gmultiset_size_pos_elem_of X : 0 < size X x, x X.
Proof. intros ?? x. rewrite !multiplicity_disj_union. by apply Nat.add_le_mono. Qed. Proof.
Lemma gmultiset_disj_union_mono_l X Y1 Y2 : Y1 Y2 X Y1 X Y2. intros Hsz. destruct (gmultiset_choose_or_empty X) as [|HX]; [done|].
Proof. intros. by apply gmultiset_disj_union_mono. Qed. contradict Hsz. rewrite HX, gmultiset_size_empty; lia.
Lemma gmultiset_disj_union_mono_r X1 X2 Y : X1 X2 X1 Y X2 Y. Qed.
Proof. intros. by apply gmultiset_disj_union_mono. Qed.
Lemma gmultiset_size_singleton x : size ({[ x ]} : gmultiset A) = 1.
Lemma gmultiset_subset X Y : X Y size X < size Y X Y. Proof.
Proof. intros. apply strict_spec_alt; split; naive_solver auto with lia. Qed. unfold size, gmultiset_size; simpl. by rewrite gmultiset_elements_singleton.
Lemma gmultiset_disj_union_subset_l X Y : Y X X Y. Qed.
Proof. Lemma gmultiset_size_disj_union X Y : size (X Y) = size X + size Y.
intros HY%gmultiset_size_non_empty_iff. Proof.
apply gmultiset_subset; auto using gmultiset_disj_union_subseteq_l. unfold size, gmultiset_size; simpl.
rewrite gmultiset_size_disj_union; lia. by rewrite gmultiset_elements_disj_union, app_length.
Qed. Qed.
Lemma gmultiset_union_subset_r X Y : X Y X Y.
Proof. rewrite (comm_L ()). apply gmultiset_disj_union_subset_l. Qed. (** Order stuff *)
Global Instance gmultiset_po : PartialOrder (⊆@{gmultiset A}).
Lemma gmultiset_elem_of_singleton_subseteq x X : x X {[ x ]} X. Proof. repeat split; repeat intro; multiset_solver. Qed.
Proof.
rewrite elem_of_multiplicity. split. Lemma gmultiset_subseteq_alt X Y :
- intros Hx y. rewrite multiplicity_singleton'. X Y
destruct (decide (y = x)); naive_solver lia. map_relation () (λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y).
- intros Hx. generalize (Hx x). rewrite multiplicity_singleton. lia. Proof.
Qed. apply forall_proper; intros x. unfold multiplicity.
destruct (gmultiset_car X !! x), (gmultiset_car Y !! x); naive_solver lia.
Lemma gmultiset_elem_of_subseteq X1 X2 x : x X1 X1 X2 x X2. Qed.
Proof. rewrite !gmultiset_elem_of_singleton_subseteq. by intros ->. Qed. Global Instance gmultiset_subseteq_dec : RelDecision (⊆@{gmultiset A}).
Proof.
Lemma gmultiset_disj_union_difference X Y : X Y Y = X Y X. refine (λ X Y, cast_if (decide (map_relation ()
Proof. (λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y))));
intros HXY. apply gmultiset_eq; intros x; specialize (HXY x). by rewrite gmultiset_subseteq_alt.
rewrite multiplicity_disj_union, multiplicity_difference; lia. Defined.
Qed.
Lemma gmultiset_disj_union_difference' x Y : x Y Y = {[ x ]} Y {[ x ]}. Lemma gmultiset_subset_subseteq X Y : X Y X Y.
Proof. Proof. apply strict_include. Qed.
intros. by apply gmultiset_disj_union_difference, Hint Resolve gmultiset_subset_subseteq : core.
gmultiset_elem_of_singleton_subseteq.
Qed. Lemma gmultiset_empty_subseteq X : X.
Proof. multiset_solver. Qed.
Lemma gmultiset_size_difference X Y : Y X size (X Y) = size X - size Y.
Proof. Lemma gmultiset_union_subseteq_l X Y : X X Y.
intros HX%gmultiset_disj_union_difference. Proof. multiset_solver. Qed.
rewrite HX at 2; rewrite gmultiset_size_disj_union. lia. Lemma gmultiset_union_subseteq_r X Y : Y X Y.
Qed. Proof. multiset_solver. Qed.
Lemma gmultiset_union_mono X1 X2 Y1 Y2 : X1 X2 Y1 Y2 X1 Y1 X2 Y2.
Lemma gmultiset_empty_difference X Y : Y X Y X = ∅. Proof. multiset_solver. Qed.
Proof. Lemma gmultiset_union_mono_l X Y1 Y2 : Y1 Y2 X Y1 X Y2.
intros HYX. unfold_leibniz. intros x. Proof. multiset_solver. Qed.
rewrite multiplicity_difference, multiplicity_empty. Lemma gmultiset_union_mono_r X1 X2 Y : X1 X2 X1 Y X2 Y.
specialize (HYX x). lia. Proof. multiset_solver. Qed.
Qed.
Lemma gmultiset_disj_union_subseteq_l X Y : X X Y.
Lemma gmultiset_non_empty_difference X Y : X Y Y X ∅. Proof. multiset_solver. Qed.
Proof. Lemma gmultiset_disj_union_subseteq_r X Y : Y X Y.
intros [_ HXY2] Hdiff; destruct HXY2; intros x. Proof. multiset_solver. Qed.
generalize (f_equal (multiplicity x) Hdiff). Lemma gmultiset_disj_union_mono X1 X2 Y1 Y2 : X1 X2 Y1 Y2 X1 Y1 X2 Y2.
rewrite multiplicity_difference, multiplicity_empty; lia. Proof. multiset_solver. Qed.
Qed. Lemma gmultiset_disj_union_mono_l X Y1 Y2 : Y1 Y2 X Y1 X Y2.
Proof. multiset_solver. Qed.
Lemma gmultiset_difference_diag X : X X = ∅. Lemma gmultiset_disj_union_mono_r X1 X2 Y : X1 X2 X1 Y X2 Y.
Proof. by apply gmultiset_empty_difference. Qed. Proof. multiset_solver. Qed.
Lemma gmultiset_difference_subset X Y : X X Y Y X Y. Lemma gmultiset_subset X Y : X Y size X < size Y X Y.
Proof. Proof. intros. apply strict_spec_alt; split; naive_solver auto with lia. Qed.
intros. eapply strict_transitive_l; [by apply gmultiset_union_subset_r|]. Lemma gmultiset_disj_union_subset_l X Y : Y X X Y.
by rewrite <-(gmultiset_disj_union_difference X Y). Proof.
Qed. intros HY%gmultiset_size_non_empty_iff.
apply gmultiset_subset; auto using gmultiset_disj_union_subseteq_l.
(** Mononicity *) rewrite gmultiset_size_disj_union; lia.
Lemma gmultiset_elements_submseteq X Y : X Y elements X ⊆+ elements Y. Qed.
Proof. Lemma gmultiset_union_subset_r X Y : X Y X Y.
intros ->%gmultiset_disj_union_difference. rewrite gmultiset_elements_disj_union. Proof. rewrite (comm_L ()). apply gmultiset_disj_union_subset_l. Qed.
by apply submseteq_inserts_r.
Qed. Lemma gmultiset_elem_of_singleton_subseteq x X : x X {[ x ]} X.
Proof.
Lemma gmultiset_subseteq_size X Y : X Y size X size Y. rewrite elem_of_multiplicity. split.
Proof. intros. by apply submseteq_length, gmultiset_elements_submseteq. Qed. - intros Hx y. rewrite multiplicity_singleton'.
destruct (decide (y = x)); naive_solver lia.
Lemma gmultiset_subset_size X Y : X Y size X < size Y. - intros Hx. generalize (Hx x). rewrite multiplicity_singleton. lia.
Proof. Qed.
intros HXY. assert (size (Y X) 0).
{ by apply gmultiset_size_non_empty_iff, gmultiset_non_empty_difference. } Lemma gmultiset_elem_of_subseteq X1 X2 x : x X1 X1 X2 x X2.
rewrite (gmultiset_disj_union_difference X Y), Proof. rewrite !gmultiset_elem_of_singleton_subseteq. by intros ->. Qed.
gmultiset_size_disj_union by auto. lia.
Qed. Lemma gmultiset_disj_union_difference X Y : X Y Y = X Y X.
Proof. multiset_solver. Qed.
(** Well-foundedness *) Lemma gmultiset_disj_union_difference' x Y : x Y Y = {[ x ]} Y {[ x ]}.
Lemma gmultiset_wf : wf (⊂@{gmultiset A}). Proof.
Proof. intros. by apply gmultiset_disj_union_difference,
apply (wf_projected (<) size); auto using gmultiset_subset_size, lt_wf. gmultiset_elem_of_singleton_subseteq.
Qed. Qed.
Lemma gmultiset_ind (P : gmultiset A Prop) : Lemma gmultiset_size_difference X Y : Y X size (X Y) = size X - size Y.
P ( x X, P X P ({[ x ]} X)) X, P X. Proof.
Proof. intros HX%gmultiset_disj_union_difference.
intros Hemp Hinsert X. induction (gmultiset_wf X) as [X _ IH]. rewrite HX at 2; rewrite gmultiset_size_disj_union. lia.
destruct (gmultiset_choose_or_empty X) as [[x Hx]| ->]; auto. Qed.
rewrite (gmultiset_disj_union_difference' x X) by done.
apply Hinsert, IH, gmultiset_difference_subset, Lemma gmultiset_empty_difference X Y : Y X Y X = ∅.
gmultiset_elem_of_singleton_subseteq; auto using gmultiset_non_empty_singleton. Proof. multiset_solver. Qed.
Qed.
End lemmas. 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.
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.
(** Mononicity *)
Lemma gmultiset_elements_submseteq X Y : X Y elements X ⊆+ elements Y.
Proof.
intros ->%gmultiset_disj_union_difference. rewrite gmultiset_elements_disj_union.
by apply submseteq_inserts_r.
Qed.
Lemma gmultiset_subseteq_size X Y : X Y size X size Y.
Proof. intros. by apply submseteq_length, gmultiset_elements_submseteq. Qed.
Lemma gmultiset_subset_size X Y : X Y size X < size Y.
Proof.
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.
Qed.
(** Well-foundedness *)
Lemma gmultiset_wf : wf (⊂@{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.
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.
Qed.
End more_lemmas.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment