Skip to content

Commit

Permalink
Adapt w.r.t. coq/coq#16004.
Browse files Browse the repository at this point in the history
  • Loading branch information
ppedrot committed Jun 29, 2022
1 parent 506d298 commit 39d315f
Show file tree
Hide file tree
Showing 14 changed files with 26 additions and 2 deletions.
2 changes: 1 addition & 1 deletion examples/StateTMonad.v
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,7 @@ Fail Definition foo : stateT unit option unit :=
ret tt.

(** Use [Existing Instance] to bring the Local [Monad_stateT] instance into context *)
Existing Instance Monad_stateT.
#[export] Existing Instance Monad_stateT.

(** Now the definition succeeds *)
Definition foo : stateT unit option unit :=
Expand Down
1 change: 1 addition & 0 deletions theories/Core/Any.v
Original file line number Diff line number Diff line change
Expand Up @@ -10,4 +10,5 @@ Polymorphic Definition RESOLVE (T : Type) : Type := T.

Existing Class RESOLVE.

#[global]
Hint Extern 0 (RESOLVE _) => unfold RESOLVE : typeclass_instances.
5 changes: 5 additions & 0 deletions theories/Data/Eq.v
Original file line number Diff line number Diff line change
Expand Up @@ -28,6 +28,7 @@ Lemma match_eq_sym_eq
Proof.
destruct pf. reflexivity.
Defined.
#[global]
Hint Rewrite match_eq_sym_eq : eq_rw.

Lemma match_eq_sym_eq'
Expand All @@ -40,6 +41,7 @@ Lemma match_eq_sym_eq'
Proof.
destruct pf. reflexivity.
Defined.
#[global]
Hint Rewrite match_eq_sym_eq' : eq_rw.


Expand Down Expand Up @@ -73,6 +75,7 @@ Lemma eq_Const_eq
Proof.
destruct pf. reflexivity.
Defined.
#[global]
Hint Rewrite eq_Const_eq : eq_rw.

Lemma eq_Arr_eq
Expand All @@ -88,11 +91,13 @@ Lemma eq_Arr_eq
Proof.
destruct pf. reflexivity.
Defined.
#[global]
Hint Rewrite eq_Arr_eq : eq_rw.

Lemma eq_sym_eq_sym : forall (T : Type) (a b : T) (pf : a = b),
eq_sym (eq_sym pf) = pf.
Proof. destruct pf. reflexivity. Defined.
#[global]
Hint Rewrite eq_sym_eq_sym : eq_rw.

Ltac autorewrite_eq_rw :=
Expand Down
2 changes: 2 additions & 0 deletions theories/Data/ListFirstnSkipn.v
Original file line number Diff line number Diff line change
Expand Up @@ -45,6 +45,7 @@ Proof.
simpl. replace (n - 0) with n; [ | lia ]. reflexivity.
Qed.

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

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

#[global]
Hint Rewrite skipn_app_L skipn_app_R skipn_0 skipn_all skipn_cons using lia : list_rw.
1 change: 1 addition & 0 deletions theories/Data/Option.v
Original file line number Diff line number Diff line change
Expand Up @@ -180,4 +180,5 @@ Proof.
destruct pf. destruct val; reflexivity.
Defined.

#[global]
Hint Rewrite eq_option_eq : eq_rw.
1 change: 1 addition & 0 deletions theories/Data/PList.v
Original file line number Diff line number Diff line change
Expand Up @@ -216,6 +216,7 @@ End pmap.

Polymorphic Definition Functor_plist@{i} : Functor@{i i} plist@{i} :=
{| fmap := @fmap_plist@{i i} |}.
#[global]
Existing Instance Functor_plist.


Expand Down
2 changes: 2 additions & 0 deletions theories/Data/POption.v
Original file line number Diff line number Diff line change
Expand Up @@ -80,9 +80,11 @@ End poption_map.

Definition Functor_poption@{i} : Functor@{i i} poption@{i} :=
{| fmap := @fmap_poption@{i i} |}.
#[global]
Existing Instance Functor_poption.

Definition Applicative_poption@{i} : Applicative@{i i} poption@{i} :=
{| pure := @pSome@{i}
; ap := @ap_poption |}.
#[global]
Existing Instance Applicative_poption.
1 change: 1 addition & 0 deletions theories/Data/SigT.v
Original file line number Diff line number Diff line change
Expand Up @@ -39,4 +39,5 @@ Lemma eq_sigT_rw
end.
Proof. destruct pf. destruct s; reflexivity. Qed.

#[global]
Hint Rewrite eq_sigT_rw : eq_rw.
2 changes: 2 additions & 0 deletions theories/Generic/Ind.v
Original file line number Diff line number Diff line change
Expand Up @@ -132,7 +132,9 @@ Ltac all_resolve :=
| |- _ => solve [ eauto with typeclass_instances ]
end.

#[global]
Hint Extern 0 (ProductResolve _ _) => all_resolve : typeclass_instances.
#[global]
Hint Extern 0 (VariantResolve _ _) => all_resolve : typeclass_instances.

Definition comma_before (b : bool) (s : showM) : showM :=
Expand Down
1 change: 1 addition & 0 deletions theories/Programming/Eqv.v
Original file line number Diff line number Diff line change
Expand Up @@ -9,6 +9,7 @@ Class EqvWF T :=
{ eqvWFEqv :> Eqv T
; eqvWFEquivalence :> Equivalence eqv
}.
#[global]
Instance EqvWF_Build {T} {E:Eqv T} {EV:Equivalence eqv} : EqvWF T :=
{ eqvWFEqv := E ; eqvWFEquivalence := EV }.

Expand Down
3 changes: 3 additions & 0 deletions theories/Programming/Injection.v
Original file line number Diff line number Diff line change
Expand Up @@ -9,10 +9,13 @@ Polymorphic Class Injection (x : Type) (t : Type) := inject : x -> t.
Class Projection x t := { project : t -> x ; pmodify : (x -> x) -> (t -> t) }.
*)

#[global]
Polymorphic Instance Injection_refl {T : Type} : Injection T T :=
{ inject := @id T }.

#[global]
Instance Injection_ascii_string : Injection ascii string :=
{ inject a := String a EmptyString }.
#[global]
Instance Injection_ascii_string_cons : Injection ascii (string -> string) :=
{ inject := String }.
3 changes: 3 additions & 0 deletions theories/Programming/Le.v
Original file line number Diff line number Diff line change
Expand Up @@ -8,9 +8,11 @@ Definition neg_lte {T} {L:Lte T} (x:T) (y:T) : Prop := not (lte x y).
Definition lt {T} {L:Lte T} x y := lte x y /\ ~lte y x.
Definition neg_lt {T} {L:Lte T} x y := not (lt x y).

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

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

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

Expand Down
3 changes: 2 additions & 1 deletion theories/Tactics/BoolTac.v
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Require Import Coq.Bool.Bool.
Set Implicit Arguments.
Set Strict Implicit.

#[global]
Hint Rewrite negb_orb negb_andb negb_involutive if_negb : bool_rw.

Lemma negb_true : forall a, negb a = true -> a = false.
Expand Down Expand Up @@ -56,4 +57,4 @@ Proof.
intros.
do_bool.
Abort.
*)
*)
1 change: 1 addition & 0 deletions theories/Tactics/Consider.v
Original file line number Diff line number Diff line change
Expand Up @@ -101,6 +101,7 @@ Section from_rel_dec.
Qed.
End from_rel_dec.

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

Expand Down

0 comments on commit 39d315f

Please sign in to comment.