forked from uds-psl/base-library
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathEqDec.v
236 lines (185 loc) · 4.63 KB
/
EqDec.v
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
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
Require Import Prelim.
(** * Decidable predicates *)
Definition dec (X: Prop) : Type := {X} + {~ X}.
Coercion dec2bool P (d: dec P) := if d then true else false.
Existing Class dec.
Definition Dec (X: Prop) (d: dec X) : dec X := d.
Arguments Dec X {d}.
Lemma Dec_reflect (X: Prop) (d: dec X) :
Dec X <-> X.
Proof.
destruct d as [A|A]; cbn; tauto.
Qed.
Notation Decb X := (dec2bool (Dec X)).
Lemma Dec_reflect_eq (X Y: Prop) (d: dec X) (e: dec Y) :
Decb X = Decb Y <-> (X <-> Y).
Proof.
destruct d as [D|D], e as [E|E]; cbn; intuition congruence.
Qed.
Lemma Dec_auto (X: Prop) (d: dec X) :
X -> Dec X.
Proof.
destruct d as [A|A]; cbn; tauto.
Qed.
Lemma Dec_auto_not (X: Prop) (d: dec X) :
~ X -> ~ Dec X.
Proof.
destruct d as [A|A]; cbn; tauto.
Qed.
Hint Resolve Dec_auto Dec_auto_not.
Hint Extern 4 => (* Improves type class inference *)
match goal with
| [ |- dec ((fun _ => _) _) ] => cbn
end : typeclass_instances.
Tactic Notation "decide" constr(p) :=
destruct (Dec p).
Tactic Notation "decide" constr(p) "as" simple_intropattern(i) :=
destruct (Dec p) as i.
Tactic Notation "decide" "_" :=
destruct (Dec _).
Tactic Notation "have" constr(E) := let X := fresh "E" in decide E as [X|X]; subst; try congruence; try omega; clear X.
Lemma Dec_true P {H : dec P} : dec2bool (Dec P) = true -> P.
Proof.
decide P; cbv in *; firstorder.
Qed.
Lemma Dec_false P {H : dec P} : dec2bool (Dec P) = false -> ~P.
Proof.
decide P; cbv in *; firstorder.
Qed.
Lemma Dec_true' (P : Prop) (d : dec P) : P -> Decb P = true.
Proof. intros H. decide P; cbn; tauto. Qed.
Lemma Dec_false' (P : Prop) (d : dec P) : (~ P) -> Decb P = false.
Proof. intros H. decide P; cbn; tauto. Qed.
Hint Extern 4 =>
match goal with
[ H : dec2bool (Dec ?P) = true |- _ ] => apply Dec_true in H
| [ H : dec2bool (Dec ?P) = false |- _ ] => apply Dec_false in H
| [ |- dec2bool (Dec ?P) = true] => apply Dec_true'
| [ |- dec2bool (Dec ?P) = false] => apply Dec_false'
end.
(** Decided propositions behave classically *)
Lemma dec_DN X :
dec X -> ~~ X -> X.
Proof.
unfold dec; tauto.
Qed.
Lemma dec_DM_and X Y :
dec X -> dec Y -> ~ (X /\ Y) -> ~ X \/ ~ Y.
Proof.
unfold dec; tauto.
Qed.
Lemma dec_DM_impl X Y :
dec X -> dec Y -> ~ (X -> Y) -> X /\ ~ Y.
Proof.
unfold dec; tauto.
Qed.
(** Propagation rules for decisions *)
Fact dec_transfer P Q :
P <-> Q -> dec P -> dec Q.
Proof.
unfold dec. tauto.
Defined.
Instance bool_dec (b: bool) :
dec b.
Proof.
unfold dec. destruct b; cbn; auto.
Defined.
Instance True_dec :
dec True.
Proof.
unfold dec; tauto.
Defined.
Instance False_dec :
dec False.
Proof.
unfold dec; tauto.
Defined.
Instance impl_dec (X Y : Prop) :
dec X -> dec Y -> dec (X -> Y).
Proof.
unfold dec; tauto.
Defined.
Instance and_dec (X Y : Prop) :
dec X -> dec Y -> dec (X /\ Y).
Proof.
unfold dec; tauto.
Defined.
Instance or_dec (X Y : Prop) :
dec X -> dec Y -> dec (X \/ Y).
Proof.
unfold dec; tauto.
Defined.
(* Coq standard modules make "not" and "iff" opaque for type class inference,
can be seen with Print HintDb typeclass_instances. *)
Instance not_dec (X : Prop) :
dec X -> dec (~ X).
Proof.
unfold not. auto.
Defined.
Instance iff_dec (X Y : Prop) :
dec X -> dec Y -> dec (X <-> Y).
Proof.
unfold iff. auto.
Defined.
(** ** Discrete types *)
Notation "'eq_dec' X" := (forall x y : X, dec (x=y)) (at level 70).
Structure eqType :=
EqType {
eqType_X :> Type;
eqType_dec : eq_dec eqType_X
}.
Arguments EqType X {_} : rename.
Canonical Structure eqType_CS X (A: eq_dec X) := EqType X.
Existing Instance eqType_dec.
(** Print the base type of [eqType] in the Canonical Structure. *)
Arguments eqType_CS (X) {_}.
Instance unit_eq_dec :
eq_dec unit.
Proof.
unfold dec. decide equality.
Defined.
Instance bool_eq_dec :
eq_dec bool.
Proof.
unfold dec. decide equality.
Defined.
Instance nat_eq_dec :
eq_dec nat.
Proof.
unfold dec. decide equality.
Defined.
Instance prod_eq_dec X Y :
eq_dec X -> eq_dec Y -> eq_dec (X * Y).
Proof.
unfold dec. decide equality.
Defined.
Instance list_eq_dec X :
eq_dec X -> eq_dec (list X).
Proof.
unfold dec. decide equality.
Defined.
Instance sum_eq_dec X Y :
eq_dec X -> eq_dec Y -> eq_dec (X + Y).
Proof.
unfold dec. decide equality.
Defined.
Instance option_eq_dec X :
eq_dec X -> eq_dec (option X).
Proof.
unfold dec. decide equality.
Defined.
Instance Empty_set_eq_dec:
eq_dec Empty_set.
Proof.
unfold dec. decide equality.
Defined.
Instance True_eq_dec:
eq_dec True.
Proof.
intros x y. destruct x,y. now left.
Defined.
Instance False_eq_dec:
eq_dec False.
Proof.
intros [].
Defined.