aboutsummaryrefslogtreecommitdiff
path: root/src/Util/Decidable.v
blob: 20ec7f0c9b7508062344cad94bf1c95a435399b6 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
(** Typeclass for decidable propositions *)

Require Import Coq.Logic.Eqdep_dec.
Require Import Crypto.Util.Sigma.
Require Import Crypto.Util.HProp.
Require Import Crypto.Util.Equality.

Local Open Scope type_scope.

Class Decidable (P : Prop) := dec : {P} + {~P}.

Notation DecidableRel R := (forall x y, Decidable (R x y)).

Global Instance hprop_eq_dec {A} `{DecidableRel (@eq A)} : IsHPropRel (@eq A) | 10.
Proof. repeat intro; apply UIP_dec; trivial with nocore. Qed.

Global Instance eq_dec_hprop {A} {x y : A} `{hp : IsHProp A} : Decidable (@eq A x y) | 5.
Proof. left; apply hp. Qed.

Ltac no_equalities_about x0 y0 :=
  lazymatch goal with
  | [ H' : x0 = y0 |- _ ] => fail
  | [ H' : y0 = x0 |- _ ] => fail
  | [ H' : x0 <> y0 |- _ ] => fail
  | [ H' : y0 <> x0 |- _ ] => fail
  | _ => idtac
  end.

Ltac destruct_decidable_step :=
  match goal with
  | [ H : Decidable _ |- _ ] => destruct H
  | [ H : forall x y : ?A, Decidable (x = y), x0 : ?A, y0 : ?A  |- _ ]
    => no_equalities_about x0 y0; destruct (H x0 y0)
  | [ H : forall a0 (x y : _), Decidable (x = y), x0 : ?A, y0 : ?A  |- _ ]
    => no_equalities_about x0 y0; destruct (H _ x0 y0)
  end.
Ltac destruct_decidable := repeat destruct_decidable_step.

Ltac pre_decide_destruct_sigma_step :=
  match goal with
  | [ H : sigT _ |- _ ] => destruct H
  | [ H : sig _ |- _ ] => destruct H
  | [ H : ex _ |- _ ] => destruct H
  end.
Ltac pre_decide_destruct_sigma := repeat pre_decide_destruct_sigma_step.

Ltac pre_decide :=
  repeat (intros
          || subst
          || destruct_decidable
          || split
          || pre_decide_destruct_sigma
          || unfold Decidable in *
          || hnf
          || pre_hprop).

Ltac solve_decidable_transparent_with tac :=
  pre_decide;
  try solve [ left; abstract tac
            | right; abstract tac
            | decide equality; eauto with nocore ].

Ltac solve_decidable_transparent := solve_decidable_transparent_with firstorder.

Local Hint Extern 0 => solve [ solve_decidable_transparent ] : typeclass_instances.
Local Hint Extern 1 => progress inversion_sigma : core.

Global Instance dec_True : Decidable True | 10 := left I.
Global Instance dec_False : Decidable False | 10 := right (fun x => x).
Global Instance dec_or {A B} `{Decidable A, Decidable B} : Decidable (A \/ B) | 10. exact _. Defined.
Global Instance dec_and {A B} `{Decidable A, Decidable B} : Decidable (A /\ B) | 10. exact _. Defined.
Global Instance dec_impl {A B} `{Decidable (B \/ ~A)} : Decidable (A -> B) | 10. exact _. Defined.
Global Instance dec_impl_simple {A B} `{Decidable A, Decidable B} : Decidable (A -> B) | 10. exact _. Defined.
Global Instance dec_iff {A B} `{Decidable A, Decidable B} : Decidable (A <-> B) | 10. exact _. Defined.
Lemma dec_not {A} `{Decidable A} : Decidable (~A).
Proof. solve_decidable_transparent. Defined.
(** Disallow infinite loops of dec_not *)
Hint Extern 0 (Decidable (~?A)) => apply (@dec_not A) : typeclass_instances.

Global Instance dec_eq_unit : DecidableRel (@eq unit) | 10. exact _. Defined.
Global Instance dec_eq_bool : DecidableRel (@eq bool) | 10. exact _. Defined.
Global Instance dec_eq_Empty_set : DecidableRel (@eq Empty_set) | 10. exact _. Defined.
Global Instance dec_eq_nat : DecidableRel (@eq nat) | 10. exact _. Defined.
Global Instance dec_eq_prod {A B} `{DecidableRel (@eq A), DecidableRel (@eq B)} : DecidableRel (@eq (A * B)) | 10. exact _. Defined.
Global Instance dec_eq_sum {A B} `{DecidableRel (@eq A), DecidableRel (@eq B)} : DecidableRel (@eq (A + B)) | 10. exact _. Defined.
Global Instance dec_eq_sigT_hprop {A P} `{DecidableRel (@eq A), forall a : A, IsHProp (P a)} : DecidableRel (@eq (@sigT A P)) | 10. exact _. Defined.
Global Instance dec_eq_sig_hprop {A} {P : A -> Prop} `{DecidableRel (@eq A), forall a : A, IsHProp (P a)} : DecidableRel (@eq (@sig A P)) | 10. exact _. Defined.

Lemma Decidable_respects_iff A B (H : A <-> B) : (Decidable A -> Decidable B) * (Decidable B -> Decidable A).
Proof. solve_decidable_transparent. Defined.

Lemma Decidable_iff_to_impl A B (H : A <-> B) : Decidable A -> Decidable B.
Proof. solve_decidable_transparent. Defined.

Lemma Decidable_iff_to_flip_impl A B (H : A <-> B) : Decidable B -> Decidable A.
Proof. solve_decidable_transparent. Defined.

(** For dubious compatibility with [eauto using]. *)
Hint Extern 2 (Decidable _) => progress unfold Decidable : typeclass_instances core.