-
Notifications
You must be signed in to change notification settings - Fork 6
/
tutorial.v
333 lines (250 loc) · 9.27 KB
/
tutorial.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
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
(* Copy and paste this whole file into https://x80.org/collacoq/ *)
From Coq Require Import Arith Bool.
(* At its core, Coq is a *typed* programing language *)
(* There are a few (not many!) built-in types *)
(*
On the online version, you can move your cursor up and down with
Alt+p and Alt+n respectively.
The result of the command should show up in the lower right hand
side of the screen.
Remember to leave it time to load up at the very beginning!
*)
Check 3.
Check 3 + 4.
Check true.
Check (true && false).
(* And so it can evaluate programs *)
Eval compute in 3 + 4.
Eval compute in (true && false).
(* Every top-level command starts with a capital letter and ends with a period. *)
(* Grammar was very important to the original implementers... *)
(* No side effects though, so no printing, scanning or reading from a file! *)
Fail Check print.
(* Not to be confused with the *top level command* "Print", which asks
the Coq *system* to print a definition. *)
Print bool.
(*
NB: Every piece of data in a program is *constant* and *immutable*
There are good reasons for this, which we will go into later.
*)
(* We can define types. This is a simple sum type *)
Inductive week_day :=
| Monday : week_day
| Tuesday : week_day
| Wednesday : week_day
| Thursday : week_day
| Friday : week_day
| Saturday : week_day
| Sunday : week_day.
Check Thursday.
(* And we can define functions *)
Definition add_one (x : nat) : nat := x + 1.
Definition return_monday (w : week_day) : week_day := Monday.
Check add_one.
Check return_monday.
(* We can work by cases over sum types. *)
Definition is_monday (w : week_day) : bool :=
match w with
| Monday => true
| _ => false
end.
Check is_monday.
Eval compute in (is_monday Tuesday).
(* There is a powerful search feature, which takes types and returns
potentially useful functions with similar types *)
(* WARNING: this does not work on the online version! *)
(* Search (bool -> bool). *)
(* We can define types which may be recursive, e.g. the type of lists *)
Inductive week_day_list :=
| Nil : week_day_list
| Cons : week_day -> week_day_list -> week_day_list.
Check (Cons Monday (Cons Tuesday Nil)).
(* Ok, we have lisp-style lists. Can we define head? *)
Definition head (l : week_day_list) (default : week_day) : week_day :=
match l with
| Nil => default
| Cons w _ => w
end.
(* We need a default value! A little unfortunate... but the price of being pure! *)
(* How about tail? *)
(* Definition tail (l : week_day_list) : week_day_list := *)
(* ??? *)
(* List notation a la lisp is tedious. Thankfully, Coq's notation system is incredibly powerful! *)
Notation "[ ]" := Nil.
Notation "[ w ]" := (Cons w Nil).
Notation "[ w1 ; w2 ; .. ; wn ]" := (Cons w1 (Cons w2 .. (Cons wn Nil) .. )).
Definition work_week := [Monday; Tuesday; Wednesday; Thursday; Friday].
(* All right, let's define some funner things. *)
Definition eq_wd (w1 w2 : week_day) : bool :=
match (w1, w2) with
| (Monday,Monday) => true
| (Tuesday,Tuesday) => true
| (Wednesday,Wednesday) => true
| (Thursday,Thursday) => true
| (Friday,Friday) => true
| (Saturday,Saturday) => true
| (Sunday,Sunday) => true
| _ => false
end.
(* How about a length function? This requires a special syntax: *)
Fixpoint length (l : week_day_list) : nat :=
match l with
| [] => 0
| Cons w ws => (length ws) + 1
end.
Eval compute in (length work_week).
(* Now let's define a membership function *)
Fixpoint is_a_member (w : week_day) (l : week_day_list) : bool :=
match l with
| [] => false
| Cons w' ws => eq_wd w w' || is_a_member w ws
end.
(* Let's test it: *)
Eval compute in (is_a_member Monday work_week).
Eval compute in (is_a_member Sunday work_week).
(* Ok, we're ready for some specifications! *)
(* The simplest kind of specification is equality: *)
Check (Monday = Monday).
(* The Prop type is the type of *specifications*. It is *not* bool! *)
Check (forall w, w = Monday).
Check (exists w, w = Monday).
(* The basic propositional connectives: *)
(* True. Always provable. *)
(* False. Never provable. *)
(* P\/Q. Provable if one of P or Q (or both) is provable. *)
(* P/\Q. Provable if both P and Q are provable. *)
(* P -> Q. Provable if, when *assuming* P is provable, then so is Q. *)
(* ~P. Provable if P is *never* provable *)
(* forall x : A, P x. Provable if for an *arbitrary* x (of type A), P x is provable. *)
(* exists x : A, P x. Provable if there is some *specific* a such that P a is provable. *)
(* a = b. Provable if, in fact, a is equal to b. Not to be confused
with ":=" which is how we define functions and constants (and is
*not* a connective). *)
(* Obviously, we can state specifications about infinite types as well: *)
Check (forall l : week_day_list, l = [] \/ exists w l', l = Cons w l').
(* Some of these specifications are provable, some are not. *)
(* We give some basic tools to prove some of these specs: *)
Lemma test1 : forall x y : week_day, x = y -> x = y.
Proof.
Abort.
(* ---------------------------------------------------------
There are some simple cheats on how to prove various kinds of specifications:
roughly, one can try certain tactics based on the *shape* of the goal and the
hypotheses. The breakdown is like this:
| | in goal | in hypotheses |
|---------------+-------------+------------------ |
| A -> B | intros | apply |
| A /\ B | split | destruct |
| A \/ B | left/right | destruct |
| ~A | intro | apply |
| True | trivial | N/A |
| False | N/A | contradiction |
| forall x, P x | intros | apply |
| exists x, P x | exists t | destruct |
| t = u | reflexivity | rewrite/inversion |
but of course, these will not always suffice in all situations.
*)
(* Here's how one proves trivial equalities: *)
Lemma test2 : Monday = Monday.
Proof.
Abort.
(* It's nice to know that this doesn't always work: *)
Lemma test2_fail : Monday = Friday.
Proof.
Fail reflexivity.
Abort.
(* We can also perform computation steps in proofs, in order to
prove things about functions: *)
Lemma test3 : return_monday Friday = Monday.
Proof.
Abort.
Lemma test3' : length work_week = 5.
Proof.
Abort.
(* But really we're interested in behavior over *all inputs*. *)
Lemma test4 : forall x : week_day, return_monday x = Monday.
Proof.
Abort.
(* How about this? Hint: use the [case some_variable] tactic *)
Lemma test5 : forall x : week_day, eq_wd x x = true.
Proof.
Abort.
(* Let's play with some logical connectives *)
Lemma test6 : forall x y z : week_day, x = y -> y = z -> x = y /\ y = z.
Proof.
Abort.
Lemma test7 : forall x y z : week_day, x = y -> x = y \/ y = z.
Proof.
Abort.
Lemma test8 : forall x y z : week_day, x = y /\ y = z -> y = z.
Proof.
Abort.
(* This one is a little tougher! tactic order matters! *)
Lemma test9 : forall x y z : week_day, x = y \/ x = z -> x = z \/ x = y.
Proof.
Abort.
(* Proofs involving equalities *)
Lemma test10 : forall x y, x = Monday -> y = x -> y = Monday.
Proof.
Abort.
Lemma test11 : Monday = Tuesday -> False.
Proof.
Abort.
(* All right we're ready for some serious stuff *)
Lemma first_real_lemma : forall x y, x = y -> eq_wd x y = true.
Proof.
Abort.
(* The other direction is harder! *)
Lemma second_real_lemma : forall x y, eq_wd x y = true -> x = y.
Proof.
Abort.
(* Now we have a program which we have proven correct! We can use this to prove some theorems! *)
Lemma test12 : Monday = Monday.
Proof.
Abort.
(* Let's show correctness of our membership function! But first we need to *specify* membership. *)
(* To do this, we specify an *inductive predicate*, which describes all the ways an element can be in a list. *)
(* We can again use the keyword inductive. *)
Inductive Mem : forall (w : week_day) (l : week_day_list), Prop :=
| Mem_head : forall w l, Mem w (Cons w l)
| Mem_tail : forall w w' l, Mem w l -> Mem w (Cons w' l).
(* We can apply constructors of Mem like lemmas *)
Lemma test13 : Mem Monday [Tuesday; Monday; Thursday].
Proof.
Abort.
(* Here's some fun existential statements: *)
Lemma test14 : exists w, Mem w work_week.
Proof.
Abort.
(* This is harder! We'll be able to prove this easily once we have the
theorems.
*)
Lemma test15 : exists w, ~ (Mem w work_week).
Proof.
Abort.
(* The theorems we want are these: *)
Theorem is_a_member_correct : forall w l, is_a_member w l = true -> Mem w l.
Proof.
Abort.
Theorem is_a_member_complete : forall w l, Mem w l -> is_a_member w l = true.
Proof.
Abort.
(*
Finally, it's important to note that I purposefully hid a number of
powerful tactics from you, for pedagogical reasons. The easiest of
these is [auto], which tries to apply simple steps to finish a
goal. Exercise: how much easier are the lemmas to prove using
[auto]?
*)
(*
There is still a lot to learn! You can check out the resources on
https://coq.inria.fr/, most notably the tutorials on
https://coq.inria.fr/documentation.
You can try more online stuff on rhino-coq:
https://x80.org/rhino-coq/
I have a few things on my github:
https://github.com/codyroux?tab=repositories&q=&type=&language=coq
Or create an account on https://coq.zulipchat.com, and ask questions
there!
Go forth and prove!
*)