-
Notifications
You must be signed in to change notification settings - Fork 0
/
Lists.v
1142 lines (866 loc) · 37 KB
/
Lists.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
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
698
699
700
701
702
703
704
705
706
707
708
709
710
711
712
713
714
715
716
717
718
719
720
721
722
723
724
725
726
727
728
729
730
731
732
733
734
735
736
737
738
739
740
741
742
743
744
745
746
747
748
749
750
751
752
753
754
755
756
757
758
759
760
761
762
763
764
765
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
793
794
795
796
797
798
799
800
801
802
803
804
805
806
807
808
809
810
811
812
813
814
815
816
817
818
819
820
821
822
823
824
825
826
827
828
829
830
831
832
833
834
835
836
837
838
839
840
841
842
843
844
845
846
847
848
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
874
875
876
877
878
879
880
881
882
883
884
885
886
887
888
889
890
891
892
893
894
895
896
897
898
899
900
901
902
903
904
905
906
907
908
909
910
911
912
913
914
915
916
917
918
919
920
921
922
923
924
925
926
927
928
929
930
931
932
933
934
935
936
937
938
939
940
941
942
943
944
945
946
947
948
949
950
951
952
953
954
955
956
957
958
959
960
961
962
963
964
965
966
967
968
969
970
971
972
973
974
975
976
977
978
979
980
981
982
983
984
985
986
987
988
989
990
991
992
993
994
995
996
997
998
999
1000
(** * Lists: Working with Structured Data *)
Require Export Induction.
Module NatList.
(* ################################################################# *)
(** * Pairs of Numbers *)
(** In an [Inductive] type definition, each constructor can take
any number of arguments -- none (as with [true] and [O]), one (as
with [S]), or more than one (as with [nybble], and here): *)
Inductive natprod : Type :=
| pair (n1 n2 : nat).
(** This declaration can be read: "The one and only way to
construct a pair of numbers is by applying the constructor [pair]
to two arguments of type [nat]." *)
Check (pair 3 5) : natprod.
(** Here are simple functions for extracting the first and
second components of a pair. *)
Definition fst (p : natprod) : nat :=
match p with
| pair x y => x
end.
Definition snd (p : natprod) : nat :=
match p with
| pair x y => y
end.
Compute (fst (pair 3 5)).
(* ===> 3 *)
(** Since pairs will be used heavily in what follows, it is nice
to be able to write them with the standard mathematical notation
[(x,y)] instead of [pair x y]. We can tell Coq to allow this with
a [Notation] declaration. *)
Notation "( x , y )" := (pair x y).
(** The new notation can be used both in expressions and in pattern
matches. *)
Compute (fst (3,5)).
Definition fst' (p : natprod) : nat :=
match p with
| (x,y) => x
end.
Definition snd' (p : natprod) : nat :=
match p with
| (x,y) => y
end.
Definition swap_pair (p : natprod) : natprod :=
match p with
| (x,y) => (y,x)
end.
(** Note that pattern-matching on a pair (with parentheses: [(x, y)])
is not to be confused with the "multiple pattern" syntax (with no
parentheses: [x, y]) that we have seen previously. The above
examples illustrate pattern matching on a pair with elements [x]
and [y], whereas, for example, the definition of [minus] in
[Basics] performs pattern matching on the values [n] and [m]:
Fixpoint minus (n m : nat) : nat :=
match n, m with
| O , _ => O
| S _ , O => n
| S n', S m' => minus n' m'
end.
The distinction is minor, but it is worth knowing that they
are not the same. For instance, the following definitions are
ill-formed:
(* Can't match on a pair with multiple patterns: *)
Definition bad_fst (p : natprod) : nat :=
match p with
| x, y => x
end.
(* Can't match on multiple values with pair patterns: *)
Definition bad_minus (n m : nat) : nat :=
match n, m with
| (O , _ ) => O
| (S _ , O ) => n
| (S n', S m') => bad_minus n' m'
end.
*)
(** Now let's try to prove a few simple facts about pairs.
If we state properties of pairs in a slightly peculiar way, we can
sometimes complete their proofs with just reflexivity (and its
built-in simplification): *)
Theorem surjective_pairing' : forall (n m : nat),
(n,m) = (fst (n,m), snd (n,m)).
Proof.
reflexivity. Qed.
(** But [reflexivity] is not enough if we state the lemma in a more
natural way: *)
Theorem surjective_pairing_stuck : forall (p : natprod),
p = (fst p, snd p).
Proof.
simpl. (* Doesn't reduce anything! *)
Abort.
(** Instead, we need to expose the structure of [p] so that
[simpl] can perform the pattern match in [fst] and [snd]. We can
do this with [destruct]. *)
Theorem surjective_pairing : forall (p : natprod),
p = (fst p, snd p).
Proof.
intros p. destruct p as [n m]. simpl. reflexivity. Qed.
(** Notice that, unlike its behavior with [nat]s, where it
generates two subgoals, [destruct] generates just one subgoal
here. That's because [natprod]s can only be constructed in one
way. *)
(** **** Exercise: 1 star, standard (snd_fst_is_swap) *)
Theorem snd_fst_is_swap : forall (p : natprod),
(snd p, fst p) = swap_pair p.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 1 star, standard, optional (fst_swap_is_snd) *)
Theorem fst_swap_is_snd : forall (p : natprod),
fst (swap_pair p) = snd p.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(* ################################################################# *)
(** * Lists of Numbers *)
(** Generalizing the definition of pairs, we can describe the
type of _lists_ of numbers like this: "A list is either the empty
list or else a pair of a number and another list." *)
Inductive natlist : Type :=
| nil
| cons (n : nat) (l : natlist).
(** For example, here is a three-element list: *)
Definition mylist := cons 1 (cons 2 (cons 3 nil)).
(** As with pairs, it is more convenient to write lists in
familiar programming notation. The following declarations
allow us to use [::] as an infix [cons] operator and square
brackets as an "outfix" notation for constructing lists. *)
Notation "x :: l" := (cons x l)
(at level 60, right associativity).
Notation "[ ]" := nil.
Notation "[ x ; .. ; y ]" := (cons x .. (cons y nil) ..).
(** It is not necessary to understand the details of these
declarations, but here is roughly what's going on in case you are
interested. The "[right associativity]" annotation tells Coq how to
parenthesize expressions involving multiple uses of [::] so that,
for example, the next three declarations mean exactly the same
thing: *)
Definition mylist1 := 1 :: (2 :: (3 :: nil)).
Definition mylist2 := 1 :: 2 :: 3 :: nil.
Definition mylist3 := [1;2;3].
(** The "[at level 60]" part tells Coq how to parenthesize
expressions that involve both [::] and some other infix operator.
For example, since we defined [+] as infix notation for the [plus]
function at level 50,
Notation "x + y" := (plus x y)
(at level 50, left associativity).
the [+] operator will bind tighter than [::], so [1 + 2 :: [3]]
will be parsed, as we'd expect, as [(1 + 2) :: [3]] rather than
[1 + (2 :: [3])].
(Expressions like "[1 + 2 :: [3]]" can be a little confusing when
you read them in a [.v] file. The inner brackets, around 3, indicate
a list, but the outer brackets, which are invisible in the HTML
rendering, are there to instruct the "coqdoc" tool that the bracketed
part should be displayed as Coq code rather than running text.)
The second and third [Notation] declarations above introduce the
standard square-bracket notation for lists; the right-hand side of
the third one illustrates Coq's syntax for declaring n-ary
notations and translating them to nested sequences of binary
constructors. *)
(* ----------------------------------------------------------------- *)
(** *** Repeat *)
(** Next let's look at several functions for constructing and
manipulating lists. First, the [repeat] function takes a number
[n] and a [count] and returns a list of length [count] in which
every element is [n]. *)
Fixpoint repeat (n count : nat) : natlist :=
match count with
| O => nil
| S count' => n :: (repeat n count')
end.
(* ----------------------------------------------------------------- *)
(** *** Length *)
(** The [length] function calculates the length of a list. *)
Fixpoint length (l:natlist) : nat :=
match l with
| nil => O
| h :: t => S (length t)
end.
(* ----------------------------------------------------------------- *)
(** *** Append *)
(** The [app] function concatenates (appends) two lists. *)
Fixpoint app (l1 l2 : natlist) : natlist :=
match l1 with
| nil => l2
| h :: t => h :: (app t l2)
end.
(** Since [app] will be used extensively, it is again convenient
to have an infix operator for it. *)
Notation "x ++ y" := (app x y)
(right associativity, at level 60).
Example test_app1: [1;2;3] ++ [4;5] = [1;2;3;4;5].
Proof. reflexivity. Qed.
Example test_app2: nil ++ [4;5] = [4;5].
Proof. reflexivity. Qed.
Example test_app3: [1;2;3] ++ nil = [1;2;3].
Proof. reflexivity. Qed.
(* ----------------------------------------------------------------- *)
(** *** Head and Tail *)
(** Here are two smaller examples of programming with lists.
The [hd] function returns the first element (the "head") of the
list, while [tl] returns everything but the first element (the
"tail"). Since the empty list has no first element, we pass
a default value to be returned in that case. *)
Definition hd (default : nat) (l : natlist) : nat :=
match l with
| nil => default
| h :: t => h
end.
Definition tl (l : natlist) : natlist :=
match l with
| nil => nil
| h :: t => t
end.
Example test_hd1: hd 0 [1;2;3] = 1.
Proof. reflexivity. Qed.
Example test_hd2: hd 0 [] = 0.
Proof. reflexivity. Qed.
Example test_tl: tl [1;2;3] = [2;3].
Proof. reflexivity. Qed.
(* ----------------------------------------------------------------- *)
(** *** Exercises *)
(** **** Exercise: 2 stars, standard, especially useful (list_funs)
Complete the definitions of [nonzeros], [oddmembers], and
[countoddmembers] below. Have a look at the tests to understand
what these functions should do. *)
Fixpoint nonzeros (l:natlist) : natlist
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_nonzeros:
nonzeros [0;1;0;2;3;0;0] = [1;2;3].
(* FILL IN HERE *) Admitted.
Fixpoint oddmembers (l:natlist) : natlist
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_oddmembers:
oddmembers [0;1;0;2;3;0;0] = [1;3].
(* FILL IN HERE *) Admitted.
Definition countoddmembers (l:natlist) : nat
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_countoddmembers1:
countoddmembers [1;0;3;1;4;5] = 4.
(* FILL IN HERE *) Admitted.
Example test_countoddmembers2:
countoddmembers [0;2;4] = 0.
(* FILL IN HERE *) Admitted.
Example test_countoddmembers3:
countoddmembers nil = 0.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 3 stars, advanced (alternate)
Complete the following definition of [alternate], which
interleaves two lists into one, alternating between elements taken
from the first list and elements from the second. See the tests
below for more specific examples.
(Note: one natural and elegant way of writing [alternate] will
fail to satisfy Coq's requirement that all [Fixpoint] definitions
be "obviously terminating." If you find yourself in this rut,
look for a slightly more verbose solution that considers elements
of both lists at the same time. One possible solution involves
defining a new kind of pairs, but this is not the only way.) *)
Fixpoint alternate (l1 l2 : natlist) : natlist
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_alternate1:
alternate [1;2;3] [4;5;6] = [1;4;2;5;3;6].
(* FILL IN HERE *) Admitted.
Example test_alternate2:
alternate [1] [4;5;6] = [1;4;5;6].
(* FILL IN HERE *) Admitted.
Example test_alternate3:
alternate [1;2;3] [4] = [1;4;2;3].
(* FILL IN HERE *) Admitted.
Example test_alternate4:
alternate [] [20;30] = [20;30].
(* FILL IN HERE *) Admitted.
(** [] *)
(* ----------------------------------------------------------------- *)
(** *** Bags via Lists *)
(** A [bag] (or [multiset]) is like a set, except that each element
can appear multiple times rather than just once. One possible
representation for a bag of numbers is as a list. *)
Definition bag := natlist.
(** **** Exercise: 3 stars, standard, especially useful (bag_functions)
Complete the following definitions for the functions
[count], [sum], [add], and [member] for bags. *)
Fixpoint count (v : nat) (s : bag) : nat
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(** All these proofs can be done just by [reflexivity]. *)
Example test_count1: count 1 [1;2;3;1;4;1] = 3.
(* FILL IN HERE *) Admitted.
Example test_count2: count 6 [1;2;3;1;4;1] = 0.
(* FILL IN HERE *) Admitted.
(** Multiset [sum] is similar to set [union]: [sum a b] contains all
the elements of [a] and of [b]. (Mathematicians usually define
[union] on multisets a little bit differently -- using max instead
of sum -- which is why we don't call this operation [union].) For
[sum], we're giving you a header that does not give explicit names
to the arguments. Moreover, it uses the keyword [Definition]
instead of [Fixpoint], so even if you had names for the arguments,
you wouldn't be able to process them recursively. The point of
stating the question this way is to encourage you to think about
whether [sum] can be implemented in another way -- perhaps by
using one or more functions that have already been defined. *)
Definition sum : bag -> bag -> bag
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_sum1: count 1 (sum [1;2;3] [1;4;1]) = 3.
(* FILL IN HERE *) Admitted.
Definition add (v : nat) (s : bag) : bag
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_add1: count 1 (add 1 [1;4;1]) = 3.
(* FILL IN HERE *) Admitted.
Example test_add2: count 5 (add 1 [1;4;1]) = 0.
(* FILL IN HERE *) Admitted.
Definition member (v : nat) (s : bag) : bool
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_member1: member 1 [1;4;1] = true.
(* FILL IN HERE *) Admitted.
Example test_member2: member 2 [1;4;1] = false.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 3 stars, standard, optional (bag_more_functions)
Here are some more [bag] functions for you to practice with. *)
(** When [remove_one] is applied to a bag without the number to
remove, it should return the same bag unchanged. (This exercise
is optional, but students following the advanced track will need
to fill in the definition of [remove_one] for a later
exercise.) *)
Fixpoint remove_one (v : nat) (s : bag) : bag
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_remove_one1:
count 5 (remove_one 5 [2;1;5;4;1]) = 0.
(* FILL IN HERE *) Admitted.
Example test_remove_one2:
count 5 (remove_one 5 [2;1;4;1]) = 0.
(* FILL IN HERE *) Admitted.
Example test_remove_one3:
count 4 (remove_one 5 [2;1;4;5;1;4]) = 2.
(* FILL IN HERE *) Admitted.
Example test_remove_one4:
count 5 (remove_one 5 [2;1;5;4;5;1;4]) = 1.
(* FILL IN HERE *) Admitted.
Fixpoint remove_all (v:nat) (s:bag) : bag
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_remove_all1: count 5 (remove_all 5 [2;1;5;4;1]) = 0.
(* FILL IN HERE *) Admitted.
Example test_remove_all2: count 5 (remove_all 5 [2;1;4;1]) = 0.
(* FILL IN HERE *) Admitted.
Example test_remove_all3: count 4 (remove_all 5 [2;1;4;5;1;4]) = 2.
(* FILL IN HERE *) Admitted.
Example test_remove_all4: count 5 (remove_all 5 [2;1;5;4;5;1;4;5;1;4]) = 0.
(* FILL IN HERE *) Admitted.
Fixpoint subset (s1 : bag) (s2 : bag) : bool
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_subset1: subset [1;2] [2;1;4;1] = true.
(* FILL IN HERE *) Admitted.
Example test_subset2: subset [1;2;2] [2;1;4;1] = false.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard, especially useful (add_inc_count)
Adding a value to a bag should increase the value's count by one.
State that as a theorem and prove it. *)
(*
Theorem bag_theorem : ...
Proof.
...
Qed.
*)
(* Do not modify the following line: *)
Definition manual_grade_for_add_inc_count : option (nat*string) := None.
(** [] *)
(* ################################################################# *)
(** * Reasoning About Lists *)
(** As with numbers, simple facts about list-processing
functions can sometimes be proved entirely by simplification. For
example, just the simplification performed by [reflexivity] is
enough for this theorem... *)
Theorem nil_app : forall l : natlist,
[] ++ l = l.
Proof. reflexivity. Qed.
(** ...because the [[]] is substituted into the
"scrutinee" (the expression whose value is being "scrutinized" by
the match) in the definition of [app], allowing the match itself
to be simplified. *)
(** Also, as with numbers, it is sometimes helpful to perform case
analysis on the possible shapes (empty or non-empty) of an unknown
list. *)
Theorem tl_length_pred : forall l:natlist,
pred (length l) = length (tl l).
Proof.
intros l. destruct l as [| n l'].
- (* l = nil *)
reflexivity.
- (* l = cons n l' *)
reflexivity. Qed.
(** Here, the [nil] case works because we've chosen to define
[tl nil = nil]. Notice that the [as] annotation on the [destruct]
tactic here introduces two names, [n] and [l'], corresponding to
the fact that the [cons] constructor for lists takes two
arguments (the head and tail of the list it is constructing). *)
(** Usually, though, interesting theorems about lists require
induction for their proofs. We'll see how to do this next. *)
(** (Micro-Sermon: As we get deeper into this material, simply
_reading_ proof scripts will not get you very far! It is
important to step through the details of each one using Coq and
think about what each step achieves. Otherwise it is more or less
guaranteed that the exercises will make no sense when you get to
them. 'Nuff said.) *)
(* ================================================================= *)
(** ** Induction on Lists *)
(** Proofs by induction over datatypes like [natlist] are a
little less familiar than standard natural number induction, but
the idea is equally simple. Each [Inductive] declaration defines
a set of data values that can be built up using the declared
constructors. For example, a boolean can be either [true] or
[false]; a number can be either [O] or [S] applied to another
number; and a list can be either [nil] or [cons] applied to a
number and a list. Moreover, applications of the declared
constructors to one another are the _only_ possible shapes
that elements of an inductively defined set can have.
This last fact directly gives rise to a way of reasoning about
inductively defined sets: a number is either [O] or else it is [S]
applied to some _smaller_ number; a list is either [nil] or else
it is [cons] applied to some number and some _smaller_ list;
etc. So, if we have in mind some proposition [P] that mentions a
list [l] and we want to argue that [P] holds for _all_ lists, we
can reason as follows:
- First, show that [P] is true of [l] when [l] is [nil].
- Then show that [P] is true of [l] when [l] is [cons n l'] for
some number [n] and some smaller list [l'], assuming that [P]
is true for [l'].
Since larger lists can always be broken down into smaller ones,
eventually reaching [nil], these two arguments together establish
the truth of [P] for all lists [l]. Here's a concrete example: *)
Theorem app_assoc : forall l1 l2 l3 : natlist,
(l1 ++ l2) ++ l3 = l1 ++ (l2 ++ l3).
Proof.
intros l1 l2 l3. induction l1 as [| n l1' IHl1'].
- (* l1 = nil *)
reflexivity.
- (* l1 = cons n l1' *)
simpl. rewrite -> IHl1'. reflexivity. Qed.
(** Notice that, as when doing induction on natural numbers, the
[as...] clause provided to the [induction] tactic gives a name to
the induction hypothesis corresponding to the smaller list [l1']
in the [cons] case.
Once again, this Coq proof is not especially illuminating as a
static document -- it is easy to see what's going on if you are
reading the proof in an interactive Coq session and you can see
the current goal and context at each point, but this state is not
visible in the written-down parts of the Coq proof. So a
natural-language proof -- one written for human readers -- will
need to include more explicit signposts; in particular, it will
help the reader stay oriented if we remind them exactly what the
induction hypothesis is in the second case. *)
(** For comparison, here is an informal proof of the same theorem. *)
(** _Theorem_: For all lists [l1], [l2], and [l3],
[(l1 ++ l2) ++ l3 = l1 ++ (l2 ++ l3)].
_Proof_: By induction on [l1].
- First, suppose [l1 = []]. We must show
([] ++ l2) ++ l3 = [] ++ (l2 ++ l3),
which follows directly from the definition of [++].
- Next, suppose [l1 = n::l1'], with
(l1' ++ l2) ++ l3 = l1' ++ (l2 ++ l3)
(the induction hypothesis). We must show
((n :: l1') ++ l2) ++ l3 = (n :: l1') ++ (l2 ++ l3).
By the definition of [++], this follows from
n :: ((l1' ++ l2) ++ l3) = n :: (l1' ++ (l2 ++ l3)),
which is immediate from the induction hypothesis. [] *)
(* ----------------------------------------------------------------- *)
(** *** Reversing a List *)
(** For a slightly more involved example of inductive proof over
lists, suppose we use [app] to define a list-reversing
function [rev]: *)
Fixpoint rev (l:natlist) : natlist :=
match l with
| nil => nil
| h :: t => rev t ++ [h]
end.
Example test_rev1: rev [1;2;3] = [3;2;1].
Proof. reflexivity. Qed.
Example test_rev2: rev nil = nil.
Proof. reflexivity. Qed.
(** For something a bit more challenging than the proofs
we've seen so far, let's prove that reversing a list does not
change its length. Our first attempt gets stuck in the successor
case... *)
Theorem rev_length_firsttry : forall l : natlist,
length (rev l) = length l.
Proof.
intros l. induction l as [| n l' IHl'].
- (* l = nil *)
reflexivity.
- (* l = n :: l' *)
(* This is the tricky case. Let's begin as usual
by simplifying. *)
simpl.
(* Now we seem to be stuck: the goal is an equality
involving [++], but we don't have any useful equations
in either the immediate context or in the global
environment! We can make a little progress by using
the IH to rewrite the goal... *)
rewrite <- IHl'.
(* ... but now we can't go any further. *)
Abort.
(** So let's take the equation relating [++] and [length] that
would have enabled us to make progress at the point where we got
stuck and state it as a separate lemma. *)
Theorem app_length : forall l1 l2 : natlist,
length (l1 ++ l2) = (length l1) + (length l2).
Proof.
(* WORKED IN CLASS *)
intros l1 l2. induction l1 as [| n l1' IHl1'].
- (* l1 = nil *)
reflexivity.
- (* l1 = cons *)
simpl. rewrite -> IHl1'. reflexivity. Qed.
(** Note that, to make the lemma as general as possible, we
quantify over _all_ [natlist]s, not just those that result from an
application of [rev]. This should seem natural, because the truth
of the goal clearly doesn't depend on the list having been
reversed. Moreover, it is easier to prove the more general
property. *)
(** Now we can complete the original proof. *)
Theorem rev_length : forall l : natlist,
length (rev l) = length l.
Proof.
intros l. induction l as [| n l' IHl'].
- (* l = nil *)
reflexivity.
- (* l = cons *)
simpl. rewrite -> app_length.
simpl. rewrite -> IHl'. rewrite plus_comm.
reflexivity.
Qed.
(** For comparison, here are informal proofs of these two theorems:
_Theorem_: For all lists [l1] and [l2],
[length (l1 ++ l2) = length l1 + length l2].
_Proof_: By induction on [l1].
- First, suppose [l1 = []]. We must show
length ([] ++ l2) = length [] + length l2,
which follows directly from the definitions of
[length], [++], and [plus].
- Next, suppose [l1 = n::l1'], with
length (l1' ++ l2) = length l1' + length l2.
We must show
length ((n::l1') ++ l2) = length (n::l1') + length l2.
This follows directly from the definitions of [length] and [++]
together with the induction hypothesis. [] *)
(** _Theorem_: For all lists [l], [length (rev l) = length l].
_Proof_: By induction on [l].
- First, suppose [l = []]. We must show
length (rev []) = length [],
which follows directly from the definitions of [length]
and [rev].
- Next, suppose [l = n::l'], with
length (rev l') = length l'.
We must show
length (rev (n :: l')) = length (n :: l').
By the definition of [rev], this follows from
length ((rev l') ++ [n]) = S (length l')
which, by the previous lemma, is the same as
length (rev l') + length [n] = S (length l').
This follows directly from the induction hypothesis and the
definition of [length]. [] *)
(** The style of these proofs is rather longwinded and pedantic.
After reading a couple like this, we might find it easier to
follow proofs that give fewer details (which we can easily work
out in our own minds or on scratch paper if necessary) and just
highlight the non-obvious steps. In this more compressed style,
the above proof might look like this: *)
(** _Theorem_: For all lists [l], [length (rev l) = length l].
_Proof_: First, observe that [length (l ++ [n]) = S (length l)]
for any [l], by a straightforward induction on [l]. The main
property again follows by induction on [l], using the observation
together with the induction hypothesis in the case where [l =
n'::l']. [] *)
(** Which style is preferable in a given situation depends on
the sophistication of the expected audience and how similar the
proof at hand is to ones that they will already be familiar with.
The more pedantic style is a good default for our present
purposes. *)
(* ================================================================= *)
(** ** [Search] *)
(** We've seen that proofs can make use of other theorems we've
already proved, e.g., using [rewrite]. But in order to refer to a
theorem, we need to know its name! Indeed, it is often hard even
to remember what theorems have been proven, much less what they
are called.
Coq's [Search] command is quite helpful with this. Let's say
you've forgotten the name of a theorem about [rev]. The command
[Search rev] will cause Coq to display a list of all theorems
involving [rev]. *)
Search rev.
(** Or say you've forgotten the name of the theorem showing that plus
is commutative. You can use a pattern to search for all theorems
involving the equality of two additions. *)
Search (_ + _ = _ + _).
(** You'll see a lot of results there, nearly all of them from the
standard library. To restrict the results, you can search inside
a particular module: *)
Search (_ + _ = _ + _) inside Induction.
(** You can also make the search more precise by using variables in
the search pattern instead of wildcards: *)
Search (?x + ?y = ?y + ?x).
(** The question mark in front of the variable is needed to indicate
that it is a variable in the search pattern, rather than a
variable that is expected to be in scope currently. *)
(** Keep [Search] in mind as you do the following exercises and
throughout the rest of the book; it can save you a lot of time!
Your IDE likely has its own functionality to help with searching.
For example, in ProofGeneral, you can run [Search] with [C-c C-a
C-a], and paste its response into your buffer with [C-c C-;]. *)
(* ================================================================= *)
(** ** List Exercises, Part 1 *)
(** **** Exercise: 3 stars, standard (list_exercises)
More practice with lists: *)
Theorem app_nil_r : forall l : natlist,
l ++ [] = l.
Proof.
(* FILL IN HERE *) Admitted.
Theorem rev_app_distr: forall l1 l2 : natlist,
rev (l1 ++ l2) = rev l2 ++ rev l1.
Proof.
(* FILL IN HERE *) Admitted.
Theorem rev_involutive : forall l : natlist,
rev (rev l) = l.
Proof.
(* FILL IN HERE *) Admitted.
(** There is a short solution to the next one. If you find yourself
getting tangled up, step back and try to look for a simpler
way. *)
Theorem app_assoc4 : forall l1 l2 l3 l4 : natlist,
l1 ++ (l2 ++ (l3 ++ l4)) = ((l1 ++ l2) ++ l3) ++ l4.
Proof.
(* FILL IN HERE *) Admitted.
(** An exercise about your implementation of [nonzeros]: *)
Lemma nonzeros_app : forall l1 l2 : natlist,
nonzeros (l1 ++ l2) = (nonzeros l1) ++ (nonzeros l2).
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 2 stars, standard (eqblist)
Fill in the definition of [eqblist], which compares
lists of numbers for equality. Prove that [eqblist l l]
yields [true] for every list [l]. *)
Fixpoint eqblist (l1 l2 : natlist) : bool
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example test_eqblist1 :
(eqblist nil nil = true).
(* FILL IN HERE *) Admitted.
Example test_eqblist2 :
eqblist [1;2;3] [1;2;3] = true.
(* FILL IN HERE *) Admitted.
Example test_eqblist3 :
eqblist [1;2;3] [1;2;4] = false.
(* FILL IN HERE *) Admitted.
Theorem eqblist_refl : forall l:natlist,
true = eqblist l l.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(* ================================================================= *)
(** ** List Exercises, Part 2 *)
(** Here are a couple of little theorems to prove about your
definitions about bags above. *)
(** **** Exercise: 1 star, standard (count_member_nonzero) *)
Theorem count_member_nonzero : forall (s : bag),
1 <=? (count 1 (1 :: s)) = true.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** The following lemma about [leb] might help you in the next exercise. *)
Theorem leb_n_Sn : forall n,
n <=? (S n) = true.
Proof.
intros n. induction n as [| n' IHn'].
- (* 0 *)
simpl. reflexivity.
- (* S n' *)
simpl. rewrite IHn'. reflexivity. Qed.
(** Before doing the next exercise, make sure you've filled in the
definition of [remove_one] above. *)
(** **** Exercise: 3 stars, advanced (remove_does_not_increase_count) *)
Theorem remove_does_not_increase_count: forall (s : bag),
(count 0 (remove_one 0 s)) <=? (count 0 s) = true.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(** **** Exercise: 3 stars, standard, optional (bag_count_sum)
Write down an interesting theorem [bag_count_sum] about bags
involving the functions [count] and [sum], and prove it using
Coq. (You may find that the difficulty of the proof depends on
how you defined [count]!) *)
(* FILL IN HERE
[] *)
(** **** Exercise: 4 stars, advanced (rev_injective)
Prove that the [rev] function is injective. There is a hard way
and an easy way to do this. *)
Theorem rev_injective : forall (l1 l2 : natlist),
rev l1 = rev l2 -> l1 = l2.
Proof.
(* FILL IN HERE *) Admitted.
(** [] *)
(* ################################################################# *)
(** * Options *)
(** Suppose we want to write a function that returns the [n]th
element of some list. If we give it type [nat -> natlist -> nat],
then we'll have to choose some number to return when the list is
too short... *)
Fixpoint nth_bad (l:natlist) (n:nat) : nat :=
match l with
| nil => 42
| a :: l' => match n with
| 0 => a
| S n' => nth_bad l' n'
end
end.
(** This solution is not so good: If [nth_bad] returns [42], we
can't tell whether that value actually appears on the input
without further processing. A better alternative is to change the
return type of [nth_bad] to include an error value as a possible
outcome. We call this type [natoption]. *)
Inductive natoption : Type :=
| Some (n : nat)
| None.
(** We can then change the above definition of [nth_bad] to
return [None] when the list is too short and [Some a] when the
list has enough members and [a] appears at position [n]. We call
this new function [nth_error] to indicate that it may result in an
error. As we see here, constructors of inductive definitions can
be capitalized. *)
Fixpoint nth_error (l:natlist) (n:nat) : natoption :=
match l with
| nil => None
| a :: l' => match n with
| O => Some a
| S n' => nth_error l' n'
end
end.
Example test_nth_error1 : nth_error [4;5;6;7] 0 = Some 4.
Proof. reflexivity. Qed.
Example test_nth_error2 : nth_error [4;5;6;7] 3 = Some 7.
Proof. reflexivity. Qed.
Example test_nth_error3 : nth_error [4;5;6;7] 9 = None.
Proof. reflexivity. Qed.
(** (In the HTML version, the boilerplate proofs of these
examples are elided. Click on a box if you want to see one.)
This example is also an opportunity to introduce one more small
feature of Coq's programming language: conditional
expressions... *)
Fixpoint nth_error' (l:natlist) (n:nat) : natoption :=
match l with
| nil => None
| a :: l' => if n =? O then Some a
else nth_error' l' (pred n)
end.
(** Coq's conditionals are exactly like those found in any other
language, with one small generalization. Since the [bool] type
is not built in, Coq actually supports conditional expressions over
_any_ inductively defined type with exactly two constructors. The
guard is considered true if it evaluates to the first constructor
in the [Inductive] definition and false if it evaluates to the
second. *)
(** The function below pulls the [nat] out of a [natoption], returning
a supplied default in the [None] case. *)