Skip to content

Commit 39d315f

Browse files
committedJun 29, 2022
Adapt w.r.t. coq/coq#16004.
1 parent 506d298 commit 39d315f

File tree

14 files changed

+26
-2
lines changed

14 files changed

+26
-2
lines changed
 

‎examples/StateTMonad.v

+1-1
Original file line numberDiff line numberDiff line change
@@ -8,7 +8,7 @@ Fail Definition foo : stateT unit option unit :=
88
ret tt.
99

1010
(** Use [Existing Instance] to bring the Local [Monad_stateT] instance into context *)
11-
Existing Instance Monad_stateT.
11+
#[export] Existing Instance Monad_stateT.
1212

1313
(** Now the definition succeeds *)
1414
Definition foo : stateT unit option unit :=

‎theories/Core/Any.v

+1
Original file line numberDiff line numberDiff line change
@@ -10,4 +10,5 @@ Polymorphic Definition RESOLVE (T : Type) : Type := T.
1010

1111
Existing Class RESOLVE.
1212

13+
#[global]
1314
Hint Extern 0 (RESOLVE _) => unfold RESOLVE : typeclass_instances.

‎theories/Data/Eq.v

+5
Original file line numberDiff line numberDiff line change
@@ -28,6 +28,7 @@ Lemma match_eq_sym_eq
2828
Proof.
2929
destruct pf. reflexivity.
3030
Defined.
31+
#[global]
3132
Hint Rewrite match_eq_sym_eq : eq_rw.
3233

3334
Lemma match_eq_sym_eq'
@@ -40,6 +41,7 @@ Lemma match_eq_sym_eq'
4041
Proof.
4142
destruct pf. reflexivity.
4243
Defined.
44+
#[global]
4345
Hint Rewrite match_eq_sym_eq' : eq_rw.
4446

4547

@@ -73,6 +75,7 @@ Lemma eq_Const_eq
7375
Proof.
7476
destruct pf. reflexivity.
7577
Defined.
78+
#[global]
7679
Hint Rewrite eq_Const_eq : eq_rw.
7780

7881
Lemma eq_Arr_eq
@@ -88,11 +91,13 @@ Lemma eq_Arr_eq
8891
Proof.
8992
destruct pf. reflexivity.
9093
Defined.
94+
#[global]
9195
Hint Rewrite eq_Arr_eq : eq_rw.
9296

9397
Lemma eq_sym_eq_sym : forall (T : Type) (a b : T) (pf : a = b),
9498
eq_sym (eq_sym pf) = pf.
9599
Proof. destruct pf. reflexivity. Defined.
100+
#[global]
96101
Hint Rewrite eq_sym_eq_sym : eq_rw.
97102

98103
Ltac autorewrite_eq_rw :=

‎theories/Data/ListFirstnSkipn.v

+2
Original file line numberDiff line numberDiff line change
@@ -45,6 +45,7 @@ Proof.
4545
simpl. replace (n - 0) with n; [ | lia ]. reflexivity.
4646
Qed.
4747

48+
#[global]
4849
Hint Rewrite firstn_app_L firstn_app_R firstn_all firstn_0 firstn_cons using lia : list_rw.
4950

5051
Lemma skipn_app_R : forall T n (a b : list T),
@@ -90,4 +91,5 @@ Proof.
9091
simpl. replace (n - 0) with n; [ | lia ]. reflexivity.
9192
Qed.
9293

94+
#[global]
9395
Hint Rewrite skipn_app_L skipn_app_R skipn_0 skipn_all skipn_cons using lia : list_rw.

‎theories/Data/Option.v

+1
Original file line numberDiff line numberDiff line change
@@ -180,4 +180,5 @@ Proof.
180180
destruct pf. destruct val; reflexivity.
181181
Defined.
182182

183+
#[global]
183184
Hint Rewrite eq_option_eq : eq_rw.

‎theories/Data/PList.v

+1
Original file line numberDiff line numberDiff line change
@@ -216,6 +216,7 @@ End pmap.
216216

217217
Polymorphic Definition Functor_plist@{i} : Functor@{i i} plist@{i} :=
218218
{| fmap := @fmap_plist@{i i} |}.
219+
#[global]
219220
Existing Instance Functor_plist.
220221

221222

‎theories/Data/POption.v

+2
Original file line numberDiff line numberDiff line change
@@ -80,9 +80,11 @@ End poption_map.
8080

8181
Definition Functor_poption@{i} : Functor@{i i} poption@{i} :=
8282
{| fmap := @fmap_poption@{i i} |}.
83+
#[global]
8384
Existing Instance Functor_poption.
8485

8586
Definition Applicative_poption@{i} : Applicative@{i i} poption@{i} :=
8687
{| pure := @pSome@{i}
8788
; ap := @ap_poption |}.
89+
#[global]
8890
Existing Instance Applicative_poption.

‎theories/Data/SigT.v

+1
Original file line numberDiff line numberDiff line change
@@ -39,4 +39,5 @@ Lemma eq_sigT_rw
3939
end.
4040
Proof. destruct pf. destruct s; reflexivity. Qed.
4141

42+
#[global]
4243
Hint Rewrite eq_sigT_rw : eq_rw.

‎theories/Generic/Ind.v

+2
Original file line numberDiff line numberDiff line change
@@ -132,7 +132,9 @@ Ltac all_resolve :=
132132
| |- _ => solve [ eauto with typeclass_instances ]
133133
end.
134134

135+
#[global]
135136
Hint Extern 0 (ProductResolve _ _) => all_resolve : typeclass_instances.
137+
#[global]
136138
Hint Extern 0 (VariantResolve _ _) => all_resolve : typeclass_instances.
137139

138140
Definition comma_before (b : bool) (s : showM) : showM :=

‎theories/Programming/Eqv.v

+1
Original file line numberDiff line numberDiff line change
@@ -9,6 +9,7 @@ Class EqvWF T :=
99
{ eqvWFEqv :> Eqv T
1010
; eqvWFEquivalence :> Equivalence eqv
1111
}.
12+
#[global]
1213
Instance EqvWF_Build {T} {E:Eqv T} {EV:Equivalence eqv} : EqvWF T :=
1314
{ eqvWFEqv := E ; eqvWFEquivalence := EV }.
1415

‎theories/Programming/Injection.v

+3
Original file line numberDiff line numberDiff line change
@@ -9,10 +9,13 @@ Polymorphic Class Injection (x : Type) (t : Type) := inject : x -> t.
99
Class Projection x t := { project : t -> x ; pmodify : (x -> x) -> (t -> t) }.
1010
*)
1111

12+
#[global]
1213
Polymorphic Instance Injection_refl {T : Type} : Injection T T :=
1314
{ inject := @id T }.
1415

16+
#[global]
1517
Instance Injection_ascii_string : Injection ascii string :=
1618
{ inject a := String a EmptyString }.
19+
#[global]
1720
Instance Injection_ascii_string_cons : Injection ascii (string -> string) :=
1821
{ inject := String }.

‎theories/Programming/Le.v

+3
Original file line numberDiff line numberDiff line change
@@ -8,9 +8,11 @@ Definition neg_lte {T} {L:Lte T} (x:T) (y:T) : Prop := not (lte x y).
88
Definition lt {T} {L:Lte T} x y := lte x y /\ ~lte y x.
99
Definition neg_lt {T} {L:Lte T} x y := not (lt x y).
1010

11+
#[global]
1112
Instance lt_RelDec {T} {L:Lte T} {RD:RelDec lte} : RelDec lt :=
1213
{ rel_dec x y := (rel_dec x y && negb (rel_dec y x))%bool }.
1314

15+
#[global]
1416
Instance lt_RelDecCorrect {T} {L:Lte T} {RD:RelDec lte} {RDC:RelDec_Correct RD}
1517
: RelDec_Correct lt_RelDec.
1618
Proof. constructor.
@@ -29,6 +31,7 @@ Class LteWF T :=
2931
; lteWFPreOrder :> PreOrder lte
3032
}.
3133

34+
#[global]
3235
Instance LteWF_Build {T} {L:Lte T} {PO:PreOrder lte} : LteWF T :=
3336
{ lteWFLte := L ; lteWFPreOrder := PO }.
3437

‎theories/Tactics/BoolTac.v

+2-1
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Require Import Coq.Bool.Bool.
33
Set Implicit Arguments.
44
Set Strict Implicit.
55

6+
#[global]
67
Hint Rewrite negb_orb negb_andb negb_involutive if_negb : bool_rw.
78

89
Lemma negb_true : forall a, negb a = true -> a = false.
@@ -56,4 +57,4 @@ Proof.
5657
intros.
5758
do_bool.
5859
Abort.
59-
*)
60+
*)

‎theories/Tactics/Consider.v

+1
Original file line numberDiff line numberDiff line change
@@ -101,6 +101,7 @@ Section from_rel_dec.
101101
Qed.
102102
End from_rel_dec.
103103

104+
#[global]
104105
Hint Extern 10 (@Reflect (?f ?a ?b) _ _) =>
105106
eapply (@Reflect_RelDecCorrect _ _ (@Build_RelDec _ _ f) _) : typeclass_instances.
106107

0 commit comments

Comments
 (0)