File tree 14 files changed +26
-2
lines changed
14 files changed +26
-2
lines changed Original file line number Diff line number Diff line change @@ -8,7 +8,7 @@ Fail Definition foo : stateT unit option unit :=
8
8
ret tt.
9
9
10
10
(** Use [Existing Instance] to bring the Local [Monad_stateT] instance into context *)
11
- Existing Instance Monad_stateT.
11
+ #[export] Existing Instance Monad_stateT.
12
12
13
13
(** Now the definition succeeds *)
14
14
Definition foo : stateT unit option unit :=
Original file line number Diff line number Diff line change @@ -10,4 +10,5 @@ Polymorphic Definition RESOLVE (T : Type) : Type := T.
10
10
11
11
Existing Class RESOLVE.
12
12
13
+ #[global]
13
14
Hint Extern 0 (RESOLVE _) => unfold RESOLVE : typeclass_instances.
Original file line number Diff line number Diff line change @@ -28,6 +28,7 @@ Lemma match_eq_sym_eq
28
28
Proof .
29
29
destruct pf. reflexivity.
30
30
Defined .
31
+ #[global]
31
32
Hint Rewrite match_eq_sym_eq : eq_rw.
32
33
33
34
Lemma match_eq_sym_eq'
@@ -40,6 +41,7 @@ Lemma match_eq_sym_eq'
40
41
Proof .
41
42
destruct pf. reflexivity.
42
43
Defined .
44
+ #[global]
43
45
Hint Rewrite match_eq_sym_eq' : eq_rw.
44
46
45
47
@@ -73,6 +75,7 @@ Lemma eq_Const_eq
73
75
Proof .
74
76
destruct pf. reflexivity.
75
77
Defined .
78
+ #[global]
76
79
Hint Rewrite eq_Const_eq : eq_rw.
77
80
78
81
Lemma eq_Arr_eq
@@ -88,11 +91,13 @@ Lemma eq_Arr_eq
88
91
Proof .
89
92
destruct pf. reflexivity.
90
93
Defined .
94
+ #[global]
91
95
Hint Rewrite eq_Arr_eq : eq_rw.
92
96
93
97
Lemma eq_sym_eq_sym : forall (T : Type) (a b : T) (pf : a = b),
94
98
eq_sym (eq_sym pf) = pf.
95
99
Proof . destruct pf. reflexivity. Defined .
100
+ #[global]
96
101
Hint Rewrite eq_sym_eq_sym : eq_rw.
97
102
98
103
Ltac autorewrite_eq_rw :=
Original file line number Diff line number Diff line change 45
45
simpl. replace (n - 0) with n; [ | lia ]. reflexivity.
46
46
Qed .
47
47
48
+ #[global]
48
49
Hint Rewrite firstn_app_L firstn_app_R firstn_all firstn_0 firstn_cons using lia : list_rw.
49
50
50
51
Lemma skipn_app_R : forall T n (a b : list T),
90
91
simpl. replace (n - 0) with n; [ | lia ]. reflexivity.
91
92
Qed .
92
93
94
+ #[global]
93
95
Hint Rewrite skipn_app_L skipn_app_R skipn_0 skipn_all skipn_cons using lia : list_rw.
Original file line number Diff line number Diff line change @@ -180,4 +180,5 @@ Proof.
180
180
destruct pf. destruct val; reflexivity.
181
181
Defined .
182
182
183
+ #[global]
183
184
Hint Rewrite eq_option_eq : eq_rw.
Original file line number Diff line number Diff line change @@ -216,6 +216,7 @@ End pmap.
216
216
217
217
Polymorphic Definition Functor_plist@{i} : Functor@{i i} plist@{i} :=
218
218
{| fmap := @fmap_plist@{i i} |}.
219
+ #[global]
219
220
Existing Instance Functor_plist.
220
221
221
222
Original file line number Diff line number Diff line change @@ -80,9 +80,11 @@ End poption_map.
80
80
81
81
Definition Functor_poption@{i} : Functor@{i i} poption@{i} :=
82
82
{| fmap := @fmap_poption@{i i} |}.
83
+ #[global]
83
84
Existing Instance Functor_poption.
84
85
85
86
Definition Applicative_poption@{i} : Applicative@{i i} poption@{i} :=
86
87
{| pure := @pSome@{i}
87
88
; ap := @ap_poption |}.
89
+ #[global]
88
90
Existing Instance Applicative_poption.
Original file line number Diff line number Diff line change @@ -39,4 +39,5 @@ Lemma eq_sigT_rw
39
39
end .
40
40
Proof . destruct pf. destruct s; reflexivity. Qed .
41
41
42
+ #[global]
42
43
Hint Rewrite eq_sigT_rw : eq_rw.
Original file line number Diff line number Diff line change @@ -132,7 +132,9 @@ Ltac all_resolve :=
132
132
| |- _ => solve [ eauto with typeclass_instances ]
133
133
end .
134
134
135
+ #[global]
135
136
Hint Extern 0 (ProductResolve _ _) => all_resolve : typeclass_instances.
137
+ #[global]
136
138
Hint Extern 0 (VariantResolve _ _) => all_resolve : typeclass_instances.
137
139
138
140
Definition comma_before (b : bool) (s : showM) : showM :=
Original file line number Diff line number Diff line change @@ -9,6 +9,7 @@ Class EqvWF T :=
9
9
{ eqvWFEqv :> Eqv T
10
10
; eqvWFEquivalence :> Equivalence eqv
11
11
}.
12
+ #[global]
12
13
Instance EqvWF_Build {T} {E:Eqv T} {EV:Equivalence eqv} : EqvWF T :=
13
14
{ eqvWFEqv := E ; eqvWFEquivalence := EV }.
14
15
Original file line number Diff line number Diff line change @@ -9,10 +9,13 @@ Polymorphic Class Injection (x : Type) (t : Type) := inject : x -> t.
9
9
Class Projection x t := { project : t -> x ; pmodify : (x -> x) -> (t -> t) }.
10
10
*)
11
11
12
+ #[global]
12
13
Polymorphic Instance Injection_refl {T : Type } : Injection T T :=
13
14
{ inject := @id T }.
14
15
16
+ #[global]
15
17
Instance Injection_ascii_string : Injection ascii string :=
16
18
{ inject a := String a EmptyString }.
19
+ #[global]
17
20
Instance Injection_ascii_string_cons : Injection ascii (string -> string) :=
18
21
{ inject := String }.
Original file line number Diff line number Diff line change @@ -8,9 +8,11 @@ Definition neg_lte {T} {L:Lte T} (x:T) (y:T) : Prop := not (lte x y).
8
8
Definition lt {T} {L:Lte T} x y := lte x y /\ ~lte y x.
9
9
Definition neg_lt {T} {L:Lte T} x y := not (lt x y).
10
10
11
+ #[global]
11
12
Instance lt_RelDec {T} {L:Lte T} {RD:RelDec lte} : RelDec lt :=
12
13
{ rel_dec x y := (rel_dec x y && negb (rel_dec y x))%bool }.
13
14
15
+ #[global]
14
16
Instance lt_RelDecCorrect {T} {L:Lte T} {RD:RelDec lte} {RDC:RelDec_Correct RD}
15
17
: RelDec_Correct lt_RelDec.
16
18
Proof . constructor.
@@ -29,6 +31,7 @@ Class LteWF T :=
29
31
; lteWFPreOrder :> PreOrder lte
30
32
}.
31
33
34
+ #[global]
32
35
Instance LteWF_Build {T} {L:Lte T} {PO:PreOrder lte} : LteWF T :=
33
36
{ lteWFLte := L ; lteWFPreOrder := PO }.
34
37
Original file line number Diff line number Diff line change @@ -3,6 +3,7 @@ Require Import Coq.Bool.Bool.
3
3
Set Implicit Arguments .
4
4
Set Strict Implicit .
5
5
6
+ #[global]
6
7
Hint Rewrite negb_orb negb_andb negb_involutive if_negb : bool_rw.
7
8
8
9
Lemma negb_true : forall a, negb a = true -> a = false.
56
57
intros.
57
58
do_bool.
58
59
Abort.
59
- *)
60
+ *)
Original file line number Diff line number Diff line change @@ -101,6 +101,7 @@ Section from_rel_dec.
101
101
Qed .
102
102
End from_rel_dec.
103
103
104
+ #[global]
104
105
Hint Extern 10 (@Reflect (?f ?a ?b) _ _) =>
105
106
eapply (@Reflect_RelDecCorrect _ _ (@Build_RelDec _ _ f) _) : typeclass_instances.
106
107
You can’t perform that action at this time.
0 commit comments