aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--theories/Classes/CMorphisms.v2
-rw-r--r--theories/Classes/Morphisms.v2
-rw-r--r--theories/Classes/Morphisms_Prop.v2
-rw-r--r--theories/FSets/FMapAVL.v2
-rw-r--r--theories/FSets/FMapFacts.v94
-rw-r--r--theories/FSets/FMapPositive.v4
-rw-r--r--theories/FSets/FSetCompat.v10
-rw-r--r--theories/FSets/FSetEqProperties.v2
-rw-r--r--theories/FSets/FSetProperties.v44
9 files changed, 87 insertions, 75 deletions
diff --git a/theories/Classes/CMorphisms.v b/theories/Classes/CMorphisms.v
index c247bc11dc..598bd8b9c5 100644
--- a/theories/Classes/CMorphisms.v
+++ b/theories/Classes/CMorphisms.v
@@ -638,7 +638,7 @@ Instance PartialOrder_proper_type `(PartialOrder A eqA R) :
Proper (eqA==>eqA==>iffT) R.
Proof.
intros.
-apply proper_sym_arrow_iffT_2; auto with *.
+apply proper_sym_arrow_iffT_2. 1-2: auto with crelations.
intros x x' Hx y y' Hy Hr.
transitivity x.
- generalize (partial_order_equivalence x x'); compute; intuition.
diff --git a/theories/Classes/Morphisms.v b/theories/Classes/Morphisms.v
index 5adb927357..43adb0b69f 100644
--- a/theories/Classes/Morphisms.v
+++ b/theories/Classes/Morphisms.v
@@ -640,7 +640,7 @@ Instance PartialOrder_proper `(PartialOrder A eqA R) :
Proper (eqA==>eqA==>iff) R.
Proof.
intros.
-apply proper_sym_impl_iff_2; auto with *.
+apply proper_sym_impl_iff_2. 1-2: auto with relations.
intros x x' Hx y y' Hy Hr.
transitivity x.
- generalize (partial_order_equivalence x x'); compute; intuition.
diff --git a/theories/Classes/Morphisms_Prop.v b/theories/Classes/Morphisms_Prop.v
index 418f7a8767..7d0382ec43 100644
--- a/theories/Classes/Morphisms_Prop.v
+++ b/theories/Classes/Morphisms_Prop.v
@@ -75,7 +75,7 @@ Instance Acc_pt_morphism {A:Type}(E R : A->A->Prop)
`(Equivalence _ E) `(Proper _ (E==>E==>iff) R) :
Proper (E==>iff) (Acc R).
Proof.
- apply proper_sym_impl_iff; auto with *.
+ apply proper_sym_impl_iff. auto with relations.
intros x y EQ WF. apply Acc_intro; intros z Hz.
rewrite <- EQ in Hz. now apply Acc_inv with x.
Qed.
diff --git a/theories/FSets/FMapAVL.v b/theories/FSets/FMapAVL.v
index 8cdc6e54c5..82055c4752 100644
--- a/theories/FSets/FMapAVL.v
+++ b/theories/FSets/FMapAVL.v
@@ -1335,7 +1335,7 @@ Proof.
apply Hl; auto.
constructor.
apply Hr; eauto.
- apply InA_InfA with (eqA:=eqke); auto with *. intros (y',e') H6.
+ apply InA_InfA with (eqA:=eqke). auto with typeclass_instances. intros (y',e') H6.
destruct (elements_aux_mapsto r acc y' e'); intuition.
red; simpl; eauto.
red; simpl; eauto with ordered_type.
diff --git a/theories/FSets/FMapFacts.v b/theories/FSets/FMapFacts.v
index 1eb6a3f372..2001201ec3 100644
--- a/theories/FSets/FMapFacts.v
+++ b/theories/FSets/FMapFacts.v
@@ -760,7 +760,7 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
Instance eqke_equiv : Equivalence eqke.
Proof.
unfold eq_key_elt; split; repeat red; firstorder.
- eauto with *.
+ eauto.
congruence.
Qed.
@@ -910,9 +910,9 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
assert (Hstep' : forall k e a m' m'', InA eqke (k,e) l -> ~In k m' ->
Add k e m' m'' -> P m' a -> P m'' (F (k,e) a)).
intros k e a m' m'' H ? ? ?; eapply Hstep; eauto.
- revert H; unfold l; rewrite InA_rev, elements_mapsto_iff; auto with *.
+ revert H; unfold l; rewrite InA_rev, elements_mapsto_iff. auto.
assert (Hdup : NoDupA eqk l).
- unfold l. apply NoDupA_rev; try red; unfold eq_key ; eauto with *.
+ unfold l. apply NoDupA_rev; try red; unfold eq_key. auto with typeclass_instances.
apply elements_3w.
assert (Hsame : forall k, find k m = findA (eqb k) l).
intros k. unfold l. rewrite elements_o, findA_rev; auto.
@@ -993,7 +993,7 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
rewrite 2 fold_spec_right. set (l:=rev (elements m)).
assert (Rstep' : forall k e a b, InA eqke (k,e) l ->
R a b -> R (f k e a) (g k e b)) by
- (intros; apply Rstep; auto; rewrite elements_mapsto_iff, <- InA_rev; auto with *).
+ (intros; apply Rstep; auto; rewrite elements_mapsto_iff, <- InA_rev; assumption).
clearbody l; clear Rstep m.
induction l; simpl; auto.
apply Rstep'; auto.
@@ -1107,17 +1107,20 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
Proof.
intros.
rewrite 2 fold_spec_right.
- assert (NoDupA eqk (rev (elements m1))) by (auto with *).
- assert (NoDupA eqk (rev (elements m2))) by (auto with *).
- apply fold_right_equivlistA_restr with (R:=complement eqk)(eqA:=eqke);
- auto with *.
+ assert (NoDupA eqk (rev (elements m1))) by auto with map typeclass_instances.
+ assert (NoDupA eqk (rev (elements m2))) by auto with map typeclass_instances.
+ apply fold_right_equivlistA_restr with (R:=complement eqk)(eqA:=eqke).
+ auto with typeclass_instances.
+ auto.
+ 2: auto with crelations.
+ 4, 5: auto with map.
intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; simpl in *; apply Comp; auto.
unfold complement, eq_key, eq_key_elt; repeat red. intuition eauto.
intros (k,e) (k',e'); unfold eq_key, uncurry; simpl; auto.
rewrite <- NoDupA_altdef; auto.
intros (k,e).
- rewrite 2 InA_rev, <- 2 elements_mapsto_iff, 2 find_mapsto_iff, H;
- auto with *.
+ rewrite 2 InA_rev, <- 2 elements_mapsto_iff, 2 find_mapsto_iff, H.
+ auto with crelations.
Qed.
Lemma fold_Equal2 : forall m1 m2 i j, Equal m1 m2 -> eqA i j ->
@@ -1125,10 +1128,13 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
Proof.
intros.
rewrite 2 fold_spec_right.
- assert (NoDupA eqk (rev (elements m1))) by (auto with * ).
- assert (NoDupA eqk (rev (elements m2))) by (auto with * ).
- apply fold_right_equivlistA_restr2 with (R:=complement eqk)(eqA:=eqke)
- ; auto with *.
+ assert (NoDupA eqk (rev (elements m1))) by auto with map typeclass_instances.
+ assert (NoDupA eqk (rev (elements m2))) by auto with map typeclass_instances.
+ apply fold_right_equivlistA_restr2 with (R:=complement eqk)(eqA:=eqke).
+ auto with typeclass_instances.
+ 1, 10: auto.
+ 2: auto with crelations.
+ 4, 5: auto with map.
- intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; simpl in *; apply Comp; auto.
- unfold complement, eq_key, eq_key_elt; repeat red. intuition eauto.
- intros (k,e) (k',e') z z' h h'; unfold eq_key, uncurry;simpl; auto.
@@ -1136,8 +1142,8 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
auto.
- rewrite <- NoDupA_altdef; auto.
- intros (k,e).
- rewrite 2 InA_rev, <- 2 elements_mapsto_iff, 2 find_mapsto_iff, H;
- auto with *.
+ rewrite 2 InA_rev, <- 2 elements_mapsto_iff, 2 find_mapsto_iff, H.
+ auto with crelations.
Qed.
@@ -1149,18 +1155,22 @@ Module WProperties_fun (E:DecidableType)(M:WSfun E).
set (f':=uncurry f).
change (f k e (fold_right f' i (rev (elements m1))))
with (f' (k,e) (fold_right f' i (rev (elements m1)))).
- assert (NoDupA eqk (rev (elements m1))) by (auto with *).
- assert (NoDupA eqk (rev (elements m2))) by (auto with *).
+ assert (NoDupA eqk (rev (elements m1))) by auto with map typeclass_instances.
+ assert (NoDupA eqk (rev (elements m2))) by auto with map typeclass_instances.
apply fold_right_add_restr with
- (R:=complement eqk)(eqA:=eqke)(eqB:=eqA); auto with *.
+ (R:=complement eqk)(eqA:=eqke)(eqB:=eqA).
+ auto with typeclass_instances.
+ auto.
+ 2: auto with crelations.
+ 4, 5: auto with map.
intros (k1,e1) (k2,e2) (Hk,He) a a' Ha; unfold f'; simpl in *. apply Comp; auto.
unfold complement, eq_key_elt, eq_key; repeat red; intuition eauto.
unfold f'; intros (k1,e1) (k2,e2); unfold eq_key, uncurry; simpl; auto.
rewrite <- NoDupA_altdef; auto.
- rewrite InA_rev, <- elements_mapsto_iff by (auto with *). firstorder.
+ rewrite InA_rev, <- elements_mapsto_iff. firstorder.
intros (a,b).
rewrite InA_cons, 2 InA_rev, <- 2 elements_mapsto_iff,
- 2 find_mapsto_iff by (auto with *).
+ 2 find_mapsto_iff.
unfold eq_key_elt; simpl.
rewrite H0.
rewrite add_o.
@@ -1791,7 +1801,7 @@ Module OrdProperties (M:S).
Lemma sort_equivlistA_eqlistA : forall l l' : list (key*elt),
sort ltk l -> sort ltk l' -> equivlistA eqke l l' -> eqlistA eqke l l'.
Proof.
- apply SortA_equivlistA_eqlistA; eauto with *.
+ apply SortA_equivlistA_eqlistA; auto with typeclass_instances.
Qed.
Ltac clean_eauto := unfold O.eqke, O.ltk; simpl; intuition; eauto.
@@ -1842,7 +1852,9 @@ Module OrdProperties (M:S).
elements m = elements_lt p m ++ elements_ge p m.
Proof.
unfold elements_lt, elements_ge, leb; intros.
- apply filter_split with (eqA:=eqk) (ltA:=ltk); eauto with *.
+ apply filter_split with (eqA:=eqk) (ltA:=ltk).
+ 1-3: auto with typeclass_instances.
+ 2: auto with map.
intros; destruct x; destruct y; destruct p.
rewrite gtb_1 in H; unfold O.ltk in H; simpl in *.
assert (~ltk (t1,e0) (k,e1)).
@@ -1856,14 +1868,14 @@ Module OrdProperties (M:S).
(elements_lt (x,e) m ++ (x,e):: elements_ge (x,e) m).
Proof.
intros; unfold elements_lt, elements_ge.
- apply sort_equivlistA_eqlistA; auto with *.
- apply (@SortA_app _ eqke); auto with *.
- apply (@filter_sort _ eqke); auto with *; clean_eauto.
+ apply sort_equivlistA_eqlistA. auto with map.
+ apply (@SortA_app _ eqke). auto with typeclass_instances.
+ apply (@filter_sort _ eqke). 1-3: auto with typeclass_instances. auto with map.
constructor; auto with map.
- apply (@filter_sort _ eqke); auto with *; clean_eauto.
- rewrite (@InfA_alt _ eqke); auto with *; try (clean_eauto; fail).
+ apply (@filter_sort _ eqke). 1-3: auto with typeclass_instances. auto with map.
+ rewrite (@InfA_alt _ eqke). 2-4: auto with typeclass_instances.
intros.
- rewrite filter_InA in H1; auto with *; destruct H1.
+ rewrite filter_InA in H1 by auto with map. destruct H1.
rewrite leb_1 in H2.
destruct y; unfold O.ltk in *; simpl in *.
rewrite <- elements_mapsto_iff in H1.
@@ -1871,22 +1883,22 @@ Module OrdProperties (M:S).
contradict H.
exists e0; apply MapsTo_1 with t0; auto with ordered_type.
ME.order.
- apply (@filter_sort _ eqke); auto with *; clean_eauto.
+ apply (@filter_sort _ eqke). 1-3: auto with typeclass_instances. auto with map.
intros.
- rewrite filter_InA in H1; auto with *; destruct H1.
+ rewrite filter_InA in H1 by auto with map. destruct H1.
rewrite gtb_1 in H3.
destruct y; destruct x0; unfold O.ltk in *; simpl in *.
inversion_clear H2.
red in H4; simpl in *; destruct H4.
ME.order.
- rewrite filter_InA in H4; auto with *; destruct H4.
+ rewrite filter_InA in H4 by auto with map. destruct H4.
rewrite leb_1 in H4.
unfold O.ltk in *; simpl in *; ME.order.
red; intros a; destruct a.
rewrite InA_app_iff, InA_cons, 2 filter_InA,
<-2 elements_mapsto_iff, leb_1, gtb_1,
find_mapsto_iff, (H0 t0), <- find_mapsto_iff,
- add_mapsto_iff by (auto with *).
+ add_mapsto_iff by auto with map.
unfold O.eqke, O.ltk; simpl.
destruct (E.compare t0 x); intuition; try fold (~E.eq x t0); auto with ordered_type.
- elim H; exists e0; apply MapsTo_1 with t0; auto.
@@ -1898,8 +1910,8 @@ Module OrdProperties (M:S).
eqlistA eqke (elements m') (elements m ++ (x,e)::nil).
Proof.
intros.
- apply sort_equivlistA_eqlistA; auto with *.
- apply (@SortA_app _ eqke); auto with *.
+ apply sort_equivlistA_eqlistA. auto with map.
+ apply (@SortA_app _ eqke). auto with typeclass_instances. auto with map. auto.
intros.
inversion_clear H2.
destruct x0; destruct y.
@@ -1911,7 +1923,7 @@ Module OrdProperties (M:S).
red; intros a; destruct a.
rewrite InA_app_iff, InA_cons, InA_nil, <- 2 elements_mapsto_iff,
find_mapsto_iff, (H0 t0), <- find_mapsto_iff,
- add_mapsto_iff by (auto with *).
+ add_mapsto_iff.
unfold O.eqke; simpl. intuition.
destruct (E.eq_dec x t0) as [Heq|Hneq]; auto.
exfalso.
@@ -1926,9 +1938,9 @@ Module OrdProperties (M:S).
eqlistA eqke (elements m') ((x,e)::elements m).
Proof.
intros.
- apply sort_equivlistA_eqlistA; auto with *.
+ apply sort_equivlistA_eqlistA. auto with map.
change (sort ltk (((x,e)::nil) ++ elements m)).
- apply (@SortA_app _ eqke); auto with *.
+ apply (@SortA_app _ eqke). auto with typeclass_instances. auto. auto with map.
intros.
inversion_clear H1.
destruct y; destruct x0.
@@ -1940,7 +1952,7 @@ Module OrdProperties (M:S).
red; intros a; destruct a.
rewrite InA_cons, <- 2 elements_mapsto_iff,
find_mapsto_iff, (H0 t0), <- find_mapsto_iff,
- add_mapsto_iff by (auto with *).
+ add_mapsto_iff.
unfold O.eqke; simpl. intuition.
destruct (E.eq_dec x t0) as [Heq|Hneq]; auto.
exfalso.
@@ -1954,7 +1966,7 @@ Module OrdProperties (M:S).
Equal m m' -> eqlistA eqke (elements m) (elements m').
Proof.
intros.
- apply sort_equivlistA_eqlistA; auto with *.
+ apply sort_equivlistA_eqlistA. 1-2: auto with map.
red; intros.
destruct x; do 2 rewrite <- elements_mapsto_iff.
do 2 rewrite find_mapsto_iff; rewrite H; split; auto.
@@ -2056,7 +2068,7 @@ Module OrdProperties (M:S).
inversion_clear H1 as [? ? H2|? ? H2].
red in H2; destruct H2; simpl in *; ME.order.
inversion_clear H4. rename H1 into H3.
- rewrite (@InfA_alt _ eqke) in H3; eauto with *.
+ rewrite (@InfA_alt _ eqke) in H3 by auto with typeclass_instances.
apply (H3 (y,x0)); auto.
Qed.
diff --git a/theories/FSets/FMapPositive.v b/theories/FSets/FMapPositive.v
index e9cb0a6aa7..c3c6c96997 100644
--- a/theories/FSets/FMapPositive.v
+++ b/theories/FSets/FMapPositive.v
@@ -679,7 +679,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
simpl; auto.
destruct o; simpl; intros.
(* Some *)
- apply (SortA_app (eqA:=eq_key_elt)); auto with *.
+ apply (SortA_app (eqA:=eq_key_elt)). 1-2: auto with typeclass_instances.
constructor; auto.
apply In_InfA; intros.
destruct y0.
@@ -698,7 +698,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
eapply xelements_bits_lt_1; eauto.
eapply xelements_bits_lt_2; eauto.
(* None *)
- apply (SortA_app (eqA:=eq_key_elt)); auto with *.
+ apply (SortA_app (eqA:=eq_key_elt)). auto with typeclass_instances. 1-2: auto.
intros x0 y0.
do 2 rewrite InA_alt.
intros (y1,(Hy1,H)) (y2,(Hy2,H0)).
diff --git a/theories/FSets/FSetCompat.v b/theories/FSets/FSetCompat.v
index 4edeaee72e..bff999042b 100644
--- a/theories/FSets/FSetCompat.v
+++ b/theories/FSets/FSetCompat.v
@@ -381,22 +381,22 @@ Module Update_Sets
Instance lt_strorder : StrictOrder lt.
Proof.
split.
- intros x Hx. apply (M.lt_not_eq Hx); auto with *.
+ intros x Hx. apply (M.lt_not_eq Hx). auto with crelations.
exact M.lt_trans.
Qed.
Instance lt_compat : Proper (eq==>eq==>iff) lt.
Proof.
- apply proper_sym_impl_iff_2; auto with *.
+ apply proper_sym_impl_iff_2. 1-2: auto with crelations.
intros s s' Hs u u' Hu H.
assert (H0 : lt s' u).
destruct (M.compare s' u) as [H'|H'|H']; auto.
- elim (M.lt_not_eq H). transitivity s'; auto with *.
+ elim (M.lt_not_eq H). transitivity s'; auto.
elim (M.lt_not_eq (M.lt_trans H H')); auto.
destruct (M.compare s' u') as [H'|H'|H']; auto.
elim (M.lt_not_eq H).
- transitivity u'; auto with *. transitivity s'; auto with *.
- elim (M.lt_not_eq (M.lt_trans H' H0)); auto with *.
+ transitivity u'. 2: auto with crelations. transitivity s'; auto.
+ elim (M.lt_not_eq (M.lt_trans H' H0)); auto with crelations.
Qed.
Definition compare s s' :=
diff --git a/theories/FSets/FSetEqProperties.v b/theories/FSets/FSetEqProperties.v
index 1542660ab7..ac08351ad9 100644
--- a/theories/FSets/FSetEqProperties.v
+++ b/theories/FSets/FSetEqProperties.v
@@ -924,7 +924,7 @@ transitivity (f x (fold f s0 i)).
apply fold_add with (eqA:=eqA); auto with set.
transitivity (g x (fold f s0 i)); auto with set.
transitivity (g x (fold g s0 i)); auto with set.
-apply gc; auto with *.
+apply gc; auto with set.
symmetry; apply fold_add with (eqA:=eqA); auto.
do 2 rewrite fold_empty; reflexivity.
Qed.
diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v
index 88d12fc387..98b445580b 100644
--- a/theories/FSets/FSetProperties.v
+++ b/theories/FSets/FSetProperties.v
@@ -365,7 +365,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
assert (Pstep' : forall x a s' s'', InA x l -> ~In x s' -> Add x s' s'' ->
P s' a -> P s'' (f x a)).
intros; eapply Pstep; eauto.
- rewrite elements_iff, <- InA_rev; auto with *.
+ rewrite elements_iff, <- InA_rev; auto.
assert (Hdup : NoDup l) by
(unfold l; eauto using elements_3w, NoDupA_rev with *).
assert (Hsame : forall x, In x s <-> InA x l) by
@@ -435,7 +435,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
intros A B R f g i j s Rempty Rstep.
rewrite 2 fold_spec_right. set (l:=rev (elements s)).
assert (Rstep' : forall x a b, InA x l -> R a b -> R (f x a) (g x b)) by
- (intros; apply Rstep; auto; rewrite elements_iff, <- InA_rev; auto with *).
+ (intros; apply Rstep; auto; rewrite elements_iff, <- InA_rev; auto).
clearbody l; clear Rstep s.
induction l; simpl; auto.
Qed.
@@ -487,7 +487,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
fold f s i = fold_right f i l.
Proof.
intros; exists (rev (elements s)); split.
- apply NoDupA_rev; auto with *.
+ apply NoDupA_rev. auto with typeclass_instances. auto with set.
split; intros.
rewrite elements_iff; do 2 rewrite InA_alt.
split; destruct 1; generalize (In_rev (elements s) x0); exists x0; intuition.
@@ -521,7 +521,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
intros; destruct (fold_0 s i f) as (l,(Hl, (Hl1, Hl2)));
destruct (fold_0 s' i f) as (l',(Hl', (Hl'1, Hl'2))).
rewrite Hl2; rewrite Hl'2; clear Hl2 Hl'2.
- apply fold_right_add with (eqA:=E.eq)(eqB:=eqA); auto with *.
+ apply fold_right_add with (eqA:=E.eq)(eqB:=eqA). auto with typeclass_instances. 1-5: auto.
rewrite <- Hl1; auto.
intros a; rewrite InA_cons; rewrite <- Hl1; rewrite <- Hl'1;
rewrite (H2 a); intuition.
@@ -550,7 +550,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
intros.
apply fold_rel with (R:=fun u v => eqA u (f x v)); intros.
reflexivity.
- transitivity (f x0 (f x b)); auto. apply Comp; auto with *.
+ transitivity (f x0 (f x b)); auto. apply Comp; auto.
Qed.
(** ** Fold is a morphism *)
@@ -559,7 +559,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
eqA (fold f s i) (fold f s i').
Proof.
intros. apply fold_rel with (R:=eqA); auto.
- intros; apply Comp; auto with *.
+ intros; apply Comp; auto.
Qed.
Lemma fold_equal :
@@ -914,7 +914,7 @@ Module OrdProperties (M:S).
Lemma sort_equivlistA_eqlistA : forall l l' : list elt,
sort E.lt l -> sort E.lt l' -> equivlistA E.eq l l' -> eqlistA E.eq l l'.
Proof.
- apply SortA_equivlistA_eqlistA; eauto with *.
+ apply SortA_equivlistA_eqlistA; auto with typeclass_instances.
Qed.
Definition gtb x y := match E.compare x y with GT _ => true | _ => false end.
@@ -958,7 +958,7 @@ Module OrdProperties (M:S).
elements s = elements_lt x s ++ elements_ge x s.
Proof.
unfold elements_lt, elements_ge, leb; intros.
- eapply (@filter_split _ E.eq _ E.lt); auto with *.
+ eapply (@filter_split _ E.eq _ E.lt). 1-2: auto with typeclass_instances. 2: auto with set.
intros.
rewrite gtb_1 in H.
assert (~E.lt y x).
@@ -972,32 +972,32 @@ Module OrdProperties (M:S).
Proof.
intros; unfold elements_ge, elements_lt.
apply sort_equivlistA_eqlistA; auto with set.
- apply (@SortA_app _ E.eq); auto with *.
- apply (@filter_sort _ E.eq); auto with *.
+ apply (@SortA_app _ E.eq). auto with typeclass_instances.
+ apply (@filter_sort _ E.eq). 1-3: auto with typeclass_instances. auto with set.
constructor; auto.
- apply (@filter_sort _ E.eq); auto with *.
- rewrite ME.Inf_alt by (apply (@filter_sort _ E.eq); eauto with *).
+ apply (@filter_sort _ E.eq). 1-3: auto with typeclass_instances. auto with set.
+ rewrite ME.Inf_alt by (apply (@filter_sort _ E.eq); auto with set typeclass_instances).
intros.
- rewrite filter_InA in H1; auto with *; destruct H1.
+ rewrite filter_InA in H1 by auto with fset. destruct H1.
rewrite leb_1 in H2.
rewrite <- elements_iff in H1.
assert (~E.eq x y).
contradict H; rewrite H; auto.
ME.order.
intros.
- rewrite filter_InA in H1; auto with *; destruct H1.
+ rewrite filter_InA in H1 by auto with fset. destruct H1.
rewrite gtb_1 in H3.
inversion_clear H2.
ME.order.
- rewrite filter_InA in H4; auto with *; destruct H4.
+ rewrite filter_InA in H4 by auto with fset. destruct H4.
rewrite leb_1 in H4.
ME.order.
red; intros a.
rewrite InA_app_iff, InA_cons, !filter_InA, <-elements_iff,
- leb_1, gtb_1, (H0 a) by auto with *.
+ leb_1, gtb_1, (H0 a) by auto with fset.
intuition.
destruct (E.compare a x); intuition.
- fold (~E.lt a x); auto with *.
+ fold (~E.lt a x); auto with ordered_type set.
Qed.
Definition Above x s := forall y, In y s -> E.lt y x.
@@ -1008,15 +1008,15 @@ Module OrdProperties (M:S).
eqlistA E.eq (elements s') (elements s ++ x::nil).
Proof.
intros.
- apply sort_equivlistA_eqlistA; auto with *.
- apply (@SortA_app _ E.eq); auto with *.
+ apply sort_equivlistA_eqlistA. auto with set.
+ apply (@SortA_app _ E.eq). auto with typeclass_instances. auto with set. auto.
intros.
inversion_clear H2.
rewrite <- elements_iff in H1.
apply ME.lt_eq with x; auto with ordered_type.
inversion H3.
red; intros a.
- rewrite InA_app_iff, InA_cons, InA_nil by auto with *.
+ rewrite InA_app_iff, InA_cons, InA_nil.
do 2 rewrite <- elements_iff; rewrite (H0 a); intuition.
Qed.
@@ -1025,9 +1025,9 @@ Module OrdProperties (M:S).
eqlistA E.eq (elements s') (x::elements s).
Proof.
intros.
- apply sort_equivlistA_eqlistA; auto with *.
+ apply sort_equivlistA_eqlistA. auto with set.
change (sort E.lt ((x::nil) ++ elements s)).
- apply (@SortA_app _ E.eq); auto with *.
+ apply (@SortA_app _ E.eq). auto with typeclass_instances. auto. auto with set.
intros.
inversion_clear H1.
rewrite <- elements_iff in H2.