From 8d5876268d6b061b85476ac5bc7ca6ab750c76b4 Mon Sep 17 00:00:00 2001 From: Ralf Jung <jung@mpi-sws.org> Date: Wed, 20 Jun 2018 15:40:33 +0200 Subject: [PATCH] use lia instead of omega --- theories/collections.v | 10 +++++----- theories/gmultiset.v | 44 +++++++++++++++++++++--------------------- theories/infinite.v | 12 ++++++------ theories/list.v | 2 +- theories/nat_cancel.v | 12 ++++++------ 5 files changed, 40 insertions(+), 40 deletions(-) diff --git a/theories/collections.v b/theories/collections.v index e4869e00..082d16d1 100644 --- a/theories/collections.v +++ b/theories/collections.v @@ -1051,22 +1051,22 @@ Section seq_set. x ∈ seq_set (C:=C) start len ↔ start ≤ x < start + len. Proof. revert start. induction len as [|len IH]; intros start; simpl. - - rewrite elem_of_empty. omega. - - rewrite elem_of_union, elem_of_singleton, IH. omega. + - rewrite elem_of_empty. lia. + - rewrite elem_of_union, elem_of_singleton, IH. lia. Qed. Lemma seq_set_start_disjoint start len : {[ start ]} ## seq_set (C:=C) (S start) len. - Proof. intros x. rewrite elem_of_singleton, elem_of_seq_set. omega. Qed. + Proof. intros x. rewrite elem_of_singleton, elem_of_seq_set. lia. Qed. Lemma seq_set_S_disjoint start len : {[ start + len ]} ## seq_set (C:=C) start len. - Proof. intros x. rewrite elem_of_singleton, elem_of_seq_set. omega. Qed. + Proof. intros x. rewrite elem_of_singleton, elem_of_seq_set. lia. Qed. Lemma seq_set_S_union start len : seq_set start (S len) ≡@{C} {[ start + len ]} ∪ seq_set start len. Proof. - intros x. rewrite elem_of_union, elem_of_singleton, !elem_of_seq_set. omega. + intros x. rewrite elem_of_union, elem_of_singleton, !elem_of_seq_set. lia. Qed. Lemma seq_set_S_union_L `{!LeibnizEquiv C} start len : diff --git a/theories/gmultiset.v b/theories/gmultiset.v index f7040a81..a7c9c110 100644 --- a/theories/gmultiset.v +++ b/theories/gmultiset.v @@ -80,14 +80,14 @@ Lemma multiplicity_union X Y x : multiplicity x (X ∪ Y) = multiplicity x X + multiplicity x Y. Proof. destruct X as [X], Y as [Y]; unfold multiplicity; simpl. - rewrite lookup_union_with. destruct (X !! _), (Y !! _); simpl; omega. + rewrite lookup_union_with. destruct (X !! _), (Y !! _); simpl; lia. Qed. Lemma multiplicity_difference X Y x : multiplicity x (X ∖ Y) = multiplicity x X - multiplicity x Y. Proof. destruct X as [X], Y as [Y]; unfold multiplicity; simpl. rewrite lookup_difference_with. - destruct (X !! _), (Y !! _); simplify_option_eq; omega. + destruct (X !! _), (Y !! _); simplify_option_eq; lia. Qed. (* Collection *) @@ -97,12 +97,12 @@ Proof. done. Qed. Global Instance gmultiset_simple_collection : SimpleCollection A (gmultiset A). Proof. split. - - intros x. rewrite elem_of_multiplicity, multiplicity_empty. omega. + - intros x. rewrite elem_of_multiplicity, multiplicity_empty. lia. - intros x y. destruct (decide (x = y)) as [->|]. + rewrite elem_of_multiplicity, multiplicity_singleton. split; auto with lia. + rewrite elem_of_multiplicity, multiplicity_singleton_ne by done. by split; auto with lia. - - intros X Y x. rewrite !elem_of_multiplicity, multiplicity_union. omega. + - intros X Y x. rewrite !elem_of_multiplicity, multiplicity_union. lia. Qed. Global Instance gmultiset_elem_of_dec : RelDecision (∈@{gmultiset A}). Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined. @@ -110,11 +110,11 @@ Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined. (* Algebraic laws *) Global Instance gmultiset_comm : Comm (=@{gmultiset A}) (∪). Proof. - intros X Y. apply gmultiset_eq; intros x. rewrite !multiplicity_union; omega. + intros X Y. apply gmultiset_eq; intros x. rewrite !multiplicity_union; lia. Qed. Global Instance gmultiset_assoc : Assoc (=@{gmultiset A}) (∪). Proof. - intros X Y Z. apply gmultiset_eq; intros x. rewrite !multiplicity_union; omega. + intros X Y Z. apply gmultiset_eq; intros x. rewrite !multiplicity_union; lia. Qed. Global Instance gmultiset_left_id : LeftId (=@{gmultiset A}) ∅ (∪). Proof. @@ -127,7 +127,7 @@ Proof. intros X. by rewrite (comm_L (∪)), (left_id_L _ _). Qed. Global Instance gmultiset_union_inj_1 X : Inj (=) (=) (X ∪). Proof. intros Y1 Y2. rewrite !gmultiset_eq. intros HX x; generalize (HX x). - rewrite !multiplicity_union. omega. + rewrite !multiplicity_union. lia. Qed. Global Instance gmultiset_union_inj_2 X : Inj (=) (=) (∪ X). Proof. intros Y1 Y2. rewrite <-!(comm_L _ X). apply (inj _). Qed. @@ -185,15 +185,15 @@ Proof. 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. - - intros. destruct (X !! x) as [n|] eqn:Hx; [|omega]. + - intros. destruct (X !! x) as [n|] eqn:Hx; [|lia]. exists (x,n); split; [|by apply elem_of_map_to_list]. - apply elem_of_replicate; auto with omega. + 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 omega. + destruct (X !! x); naive_solver lia. Qed. (* Properties of the size operation *) @@ -250,7 +250,7 @@ Lemma gmultiset_subseteq_alt X Y : map_relation (≤) (λ _, 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 omega. + destruct (gmultiset_car X !! x), (gmultiset_car Y !! x); naive_solver lia. Qed. Global Instance gmultiset_subseteq_dec : RelDecision (⊆@{gmultiset A}). Proof. @@ -264,12 +264,12 @@ Proof. apply strict_include. Qed. Hint Resolve gmultiset_subset_subseteq. Lemma gmultiset_empty_subseteq X : ∅ ⊆ X. -Proof. intros x. rewrite multiplicity_empty. omega. Qed. +Proof. intros x. rewrite multiplicity_empty. lia. Qed. Lemma gmultiset_union_subseteq_l X Y : X ⊆ X ∪ Y. -Proof. intros x. rewrite multiplicity_union. omega. Qed. +Proof. intros x. rewrite multiplicity_union. lia. Qed. Lemma gmultiset_union_subseteq_r X Y : Y ⊆ X ∪ Y. -Proof. intros x. rewrite multiplicity_union. omega. Qed. +Proof. intros x. rewrite multiplicity_union. lia. Qed. Lemma gmultiset_union_mono X1 X2 Y1 Y2 : X1 ⊆ X2 → Y1 ⊆ Y2 → X1 ∪ Y1 ⊆ X2 ∪ Y2. Proof. intros ?? x. rewrite !multiplicity_union. by apply Nat.add_le_mono. Qed. Lemma gmultiset_union_mono_l X Y1 Y2 : Y1 ⊆ Y2 → X ∪ Y1 ⊆ X ∪ Y2. @@ -278,12 +278,12 @@ Lemma gmultiset_union_mono_r X1 X2 Y : X1 ⊆ X2 → X1 ∪ Y ⊆ X2 ∪ Y. Proof. intros. by apply gmultiset_union_mono. Qed. Lemma gmultiset_subset X Y : X ⊆ Y → size X < size Y → X ⊂ Y. -Proof. intros. apply strict_spec_alt; split; naive_solver auto with omega. Qed. +Proof. intros. apply strict_spec_alt; split; naive_solver auto with lia. Qed. Lemma gmultiset_union_subset_l X Y : Y ≠∅ → X ⊂ X ∪ Y. Proof. intros HY%gmultiset_size_non_empty_iff. apply gmultiset_subset; auto using gmultiset_union_subseteq_l. - rewrite gmultiset_size_union; omega. + rewrite gmultiset_size_union; lia. Qed. Lemma gmultiset_union_subset_r X Y : X ≠∅ → Y ⊂ X ∪ Y. Proof. rewrite (comm_L (∪)). apply gmultiset_union_subset_l. Qed. @@ -292,9 +292,9 @@ Lemma gmultiset_elem_of_singleton_subseteq x X : x ∈ X ↔ {[ x ]} ⊆ X. Proof. rewrite elem_of_multiplicity. split. - intros Hx y; destruct (decide (x = y)) as [->|]. - + rewrite multiplicity_singleton; omega. - + rewrite multiplicity_singleton_ne by done; omega. - - intros Hx. generalize (Hx x). rewrite multiplicity_singleton. omega. + + rewrite multiplicity_singleton; lia. + + rewrite multiplicity_singleton_ne by done; lia. + - intros Hx. generalize (Hx x). rewrite multiplicity_singleton. lia. Qed. Lemma gmultiset_elem_of_subseteq X1 X2 x : x ∈ X1 → X1 ⊆ X2 → x ∈ X2. @@ -303,7 +303,7 @@ Proof. rewrite !gmultiset_elem_of_singleton_subseteq. by intros ->. Qed. Lemma gmultiset_union_difference X Y : X ⊆ Y → Y = X ∪ Y ∖ X. Proof. intros HXY. apply gmultiset_eq; intros x; specialize (HXY x). - rewrite multiplicity_union, multiplicity_difference; omega. + rewrite multiplicity_union, multiplicity_difference; lia. Qed. Lemma gmultiset_union_difference' x Y : x ∈ Y → Y = {[ x ]} ∪ Y ∖ {[ x ]}. Proof. @@ -314,14 +314,14 @@ Qed. Lemma gmultiset_size_difference X Y : Y ⊆ X → size (X ∖ Y) = size X - size Y. Proof. intros HX%gmultiset_union_difference. - rewrite HX at 2; rewrite gmultiset_size_union. omega. + rewrite HX at 2; rewrite gmultiset_size_union. lia. 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; omega. + rewrite multiplicity_difference, multiplicity_empty; lia. Qed. Lemma gmultiset_difference_subset X Y : X ≠∅ → X ⊆ Y → Y ∖ X ⊂ Y. diff --git a/theories/infinite.v b/theories/infinite.v index 81863ebc..e5c5ceb3 100644 --- a/theories/infinite.v +++ b/theories/infinite.v @@ -52,15 +52,15 @@ Section Fresh. revert n. induction s as [s IH] using (well_founded_ind collection_wf); intros n. setoid_rewrite fresh_generic_fixpoint_unfold; unfold fresh_generic_body. - destruct decide as [Hcase|Hcase]; [|by eauto with omega]. + destruct decide as [Hcase|Hcase]; [|by eauto with lia]. destruct (IH _ (subset_difference_elem_of Hcase) (S n)) as (m & Hmbound & Heqfix & Hnotin & Hinbelow). - exists m; repeat split; auto with omega. + exists m; repeat split; auto with lia. - rewrite not_elem_of_difference, elem_of_singleton in Hnotin. - destruct Hnotin as [?|?%inject_injective]; auto with omega. + destruct Hnotin as [?|?%inject_injective]; auto with lia. - intros i Hibound. destruct (decide (i = n)) as [<-|Hneq]; [by auto|]. - assert (inject i ∈ s ∖ {[inject n]}) by auto with omega. + assert (inject i ∈ s ∖ {[inject n]}) by auto with lia. set_solver. Qed. @@ -76,8 +76,8 @@ Section Fresh. destruct (fresh_generic_fixpoint_spec Y 0) as (mY & _ & -> & HnotinY & HbelowinY). destruct (Nat.lt_trichotomy mX mY) as [case|[->|case]]; auto. - + contradict HnotinX. rewrite HeqXY. apply HbelowinY; omega. - + contradict HnotinY. rewrite <-HeqXY. apply HbelowinX; omega. + + contradict HnotinX. rewrite HeqXY. apply HbelowinY; lia. + + contradict HnotinY. rewrite <-HeqXY. apply HbelowinX; lia. - intros X. unfold fresh, fresh_generic. destruct (fresh_generic_fixpoint_spec X 0) as (m & _ & -> & HnotinX & HbelowinX); auto. diff --git a/theories/list.v b/theories/list.v index 367594a6..72dee9a4 100644 --- a/theories/list.v +++ b/theories/list.v @@ -1484,7 +1484,7 @@ Proof. eapply Permutation_cons, (IH _ g). + rewrite length_delete by (rewrite <-Hl; eauto); simpl in *; lia. + unfold g. intros i j Hg. - repeat destruct (lt_eq_lt_dec _ _) as [[?|?]|?]; simplify_eq/=; try omega. + repeat destruct (lt_eq_lt_dec _ _) as [[?|?]|?]; simplify_eq/=; try lia. apply (inj S), (inj f); lia. + intros i. unfold g. destruct (lt_eq_lt_dec _ _) as [[?|?]|?]. * by rewrite lookup_delete_lt, <-Hl. diff --git a/theories/nat_cancel.v b/theories/nat_cancel.v index c9aea3cc..23028c93 100644 --- a/theories/nat_cancel.v +++ b/theories/nat_cancel.v @@ -78,25 +78,25 @@ Module nat_cancel. Instance nat_cancel_leaf_plus m m' m'' n1 n2 n1' n2' n1'n2' : NatCancelR m n1 m' n1' → NatCancelR m' n2 m'' n2' → MakeNatPlus n1' n2' n1'n2' → NatCancelR m (n1 + n2) m'' n1'n2' | 2. - Proof. unfold NatCancelR, NatCancelL, MakeNatPlus. omega. Qed. + Proof. unfold NatCancelR, NatCancelL, MakeNatPlus. lia. Qed. Instance nat_cancel_leaf_S_here m n m' n' : NatCancelR m n m' n' → NatCancelR (S m) (S n) m' n' | 3. - Proof. unfold NatCancelR, NatCancelL. omega. Qed. + Proof. unfold NatCancelR, NatCancelL. lia. Qed. Instance nat_cancel_leaf_S_else m n m' n' : NatCancelR m n m' n' → NatCancelR m (S n) m' (S n') | 4. - Proof. unfold NatCancelR, NatCancelL. omega. Qed. + Proof. unfold NatCancelR, NatCancelL. lia. Qed. (** The instance [nat_cancel_S_both] is redundant, but may reduce proof search quite a bit, e.g. when canceling constants in constants. *) Instance nat_cancel_S_both m n m' n' : NatCancelL m n m' n' → NatCancelL (S m) (S n) m' n' | 1. - Proof. unfold NatCancelL. omega. Qed. + Proof. unfold NatCancelL. lia. Qed. Instance nat_cancel_plus m1 m2 m1' m2' m1'm2' n n' n'' : NatCancelL m1 n m1' n' → NatCancelL m2 n' m2' n'' → MakeNatPlus m1' m2' m1'm2' → NatCancelL (m1 + m2) n m1'm2' n'' | 2. - Proof. unfold NatCancelL, MakeNatPlus. omega. Qed. + Proof. unfold NatCancelL, MakeNatPlus. lia. Qed. Instance nat_cancel_S m m' m'' Sm' n n' n'' : NatCancelL m n m' n' → NatCancelR 1 n' m'' n'' → MakeNatS m'' m' Sm' → NatCancelL (S m) n Sm' n'' | 3. - Proof. unfold NatCancelR, NatCancelL, MakeNatS. omega. Qed. + Proof. unfold NatCancelR, NatCancelL, MakeNatS. lia. Qed. End nat_cancel. -- GitLab