5 include "tutorial/chapter4.ma".
6 include "arithmetics/nat.ma".
9 (******************************** Option Type *********************************)
11 (* Matita may only define total functions. However, not always we may return a
12 meaningful value: for instance, working on natural numbers, the predecessor of
13 0 is undefined. In these situations, we may either return a default value
14 (usually, pred 0 = 0), or decide to return an "option type" as a result. An
15 Option type is a polymorphic type that represents encapsulation of an optional
16 value. It consists of either an empty constructor (traditionally called None),
17 or a constructor encapsulating the original data type A (written Some A): *)
19 inductive option (A:Type[0]) : Type[0] ≝
21 | Some : A → option A.
23 let rec pred (n:nat) : option nat ≝
29 (* The type option A is isomorphic to the disjoint sum between A and unit. The
30 two bijections are simply defined as follows: *)
32 definition option_to_sum : ∀A.option A → unit + A ≝ λA.λx:option A.
35 | Some a ⇒ inr ?? a ].
37 definition sum_to_option :∀A. unit + A → option A ≝ λA.λx:unit+A.
42 (* The fact that going from the option type to the sum and back again we get the
43 original element follows from a straightforward case analysis: *)
45 lemma option_bij1: ∀A,x. sum_to_option A (option_to_sum A x) = x.
46 #A * [normalize // | #a normalize // ] qed.
48 (* The other direction is just slightly more complex, since we need to exploit
49 the fact that the unit type contains a single element: we could use lemma
50 eq_unit or proceed by cases on the unit element. *)
52 lemma option_bij2: ∀A,x. option_to_sum A (sum_to_option A x) = x.
53 #A * #x [normalize cases x // | //] qed.
55 (* We shall see more examples of functions involving the otion type in the
58 (********************************** Lists *************************************)
60 (* The option type, as well as the cartesian product or the sum are simple
61 examples of polymorphic types, that is types that depend in a uniform and
62 effective way on other types. This form of polymorphism (sometimes called
63 "parametric" polymporphism to distinguish it from the "ad hoc" polymorphism
64 typical of object oriented languages) is a major feature of modern functional
65 programming languages: in many intersting cases, it allows to write generic
66 functions operating on inputs without depending on their type. Typical examples
67 can be found on lists.
68 For instance, appending two lists is an operation that is essentially
69 independent from the type A of the elements in the list, and we would like to
70 write a single append function working parametrically on A.
72 The first step is to define a type for lists polimorphic over the type A of its
75 inductive list (A:Type[0]) : Type[0] :=
77 | cons: A -> list A -> list A.
79 (* The definition should be clear: a list is either empty (nil) or it is
80 obtained by concatenating (cons) an element of type A to a given list. In other
81 words, the type of lists is the smallest inductive type generated by the two
82 constructors nil and cons.
84 The first element of a list is called its "head". If the list is empty the head
85 is undefined: as discussed in the previous section, we should either return a
86 ``default'' element, or decide to return an option type.
87 We have an additional complication in this case, since the ``default'' element
88 should have type A, and we know nothing about A (it could even be an empty
89 type!). We have no way to guess a canonical element, and the only possibility
90 (along this approach) is to take it in input. These are the two options: *)
92 definition hd ≝ λA.λl: list A.λd:A.
93 match l with [ nil ⇒ d | cons a _ ⇒ a].
95 definition option_hd ≝ λA.λl:list A.
96 match l with [ nil ⇒ None ? | cons a _ ⇒ Some ? a ].
98 (* What remains of a list after removing its head is called the "tail" of the
99 list. Again, the tail of an empty list is undefined, but in this case the result
100 must be a list, and we have a canonical inhabitant of lists, that is the empty
101 list. So, it is natural to extend the definition of tail saying that the tail of
102 nil is nil (that is very similar to say that the predecessor of 0 is 0): *)
104 definition tail ≝ λA.λl: list A.
105 match l with [ nil ⇒ nil ? | cons hd tl ⇒ tl].
107 (* Using destruct, it is easy to prove that cons is injective in both
110 lemma cons_injective_l : ∀A.∀a1,a2:A.∀l1,l2.
111 cons ?a1 l1 = cons ? a2 l2 → a1 = a2.
112 #A #a1 #a2 #l1 #l2 #Heq destruct // qed.
114 lemma cons_injective_r : ∀A.∀a1,a2:A.∀l1,l2.
115 cons ? a1 l1 = cons ? a2 l2 → l1 = l2.
116 #A #a1 #a2 #l1 #l2 #Heq destruct // qed.
118 (* The append function is defined by recursion over the first list: *)
119 let rec append A (l1: list A) l2 on l1 ≝
122 | cons hd tl ⇒ cons ? hd (append A tl l2) ].
124 (******************************* list notation ********************************)
126 (* Before discussing more operations over lists, it is convenient to introduce a
127 bit of notation; in particular, we shall write a::l for (cons ? a l);
128 [a1;a2,…, an] for the list composed by the elements a1, a2,…, an and l1@l2 for
129 the concatenation of l1 and l2.
131 This can be obtained by the following declarations: *)
133 notation "hvbox(hd break :: tl)"
134 right associative with precedence 47 for @{'cons $hd $tl}.
136 notation "[ list0 term 19 x sep ; ]"
137 non associative with precedence 90
138 for ${fold right @'nil rec acc @{'cons $x $acc}}.
140 notation "hvbox(l1 break @ l2)"
141 right associative with precedence 47 for @{'append $l1 $l2}.
143 interpretation "nil" 'nil = (nil ?).
144 interpretation "cons" 'cons hd tl = (cons ? hd tl).
145 interpretation "append" 'append l1 l2 = (append ? l1 l2).
147 (* Note that [ ] is an alternative notation for the empty list, and [a] is a
148 list containing a singleton element a.
149 Let us check the notation: *)
151 example append_ex1: [true;false]@[true] = [true;false;true].
154 (* As a simple exerices, now that we have a comfortable notation, let to prove
155 that append is associative, and that nil is a neutral element for it. *)
157 lemma associative_append: ∀A.∀l1,l2,l3:list A. l1 @ (l2 @ l3) = (l1 @ l2) @ l3.
158 #A #l1 elim l1 [normalize //]
159 #a #tl #Hind #l2 #l3 normalize @eq_f @Hind qed.
161 lemma append_nil: ∀A.∀l:list A. l@[ ] = l.
162 #A #l elim l // #a #tl #Hind normalize >Hind // qed.
164 let rec length (A:Type[0]) (l:list A) on l ≝
167 | cons a tl ⇒ 1+(length A tl)
170 interpretation "list length" 'card l = (length ? l).
172 lemma length_tail: ∀A,l. length ? (tail A l) = pred (length ? l).
176 lemma length_append: ∀A.∀l1,l2:list A.
178 #A #l1 elim l1 // normalize /2/
181 let rec nth (A:Type[0]) (l:list A) n on n ≝
184 | S x ⇒ nth A (tail ? l) x
187 lemma nth_length : ∀A,l,n. n < length A l → nth A l n ≠ None ?.
189 [#n normalize #H @False_ind /2 by absurd/
190 |#a #tl #Hind #n normalize
191 #H lapply (le_S_S_to_le … H) -H cases n
192 [#_ normalize % #H1 destruct (H1)
193 |#m #H normalize @Hind @H
199 (*************************** More list operations *****************************)
201 (* We conclude this section discussing another important operation over lists,
202 namely the reverse funtion, retrurning a list with the same elements of the
203 original list but in reverse order.
205 One could define the revert operation as follows: *)
207 let rec rev A (l1: list A) on l1 ≝
210 | cons a tl ⇒ rev A tl @ [a]
213 (* However, this implementation is sligtly inefficient, since it has a quadratic
214 complexity. A better solution consists in exploiting an accumulator, passing
215 through the following auxilary function: *)
217 let rec rev_append A (l1,l2:list A) on l1 ≝
220 | cons a tl ⇒ rev_append A tl (a::l2)
223 definition reverse ≝λS.λl.rev_append S l [].
225 lemma reverse_single : ∀S,a. reverse S [a] = [a].
228 example example_rev1: reverse ? [true;false] = [false;true].
231 (* Let us prove the following results, as a simple exercise.
233 lemma reverse_cons : ∀A,a,l. reverse A (a::l) = (reverse A l)@[a].
235 If you try a brute force approach (e.g., induction on l); you will get in
236 trouble. The point is that reverse is defined as a particular case of
237 rev_append, and you should suitably generalize your result taking rev_append in
241 lemma rev_append_append: ∀A,l1,l2,l3.
242 rev_append A l1 (l2@l3) = rev_append A l1 l2 @l3.
244 #hd #tl normalize #Hind #l2 #l3 @(Hind (hd::l2)) qed.
246 (* Now we get reverse_cons as a trivial corollary, since
247 reverse A (a::l) = rev_append A (a::l) [ ] = rev_append l [a] =
248 rev_append l ([ ]@[a]) = (rev_append l [ ])@[a] = reverse A l @ [a]
251 lemma reverse_cons : ∀A,a,l. reverse A (a::l) = reverse A l@[a].
252 #A #a #l normalize <rev_append_append % qed.
254 (* More generally, we can prove the following property of rev_append
255 specyfing its general behaviour *)
257 lemma rev_append_def : ∀S,l1,l2.
258 rev_append S l1 l2 = (reverse S l1) @ l2 .
259 #S #l1 #l2 normalize <rev_append_append //
262 lemma reverse_append: ∀S,l1,l2.
263 reverse S (l1 @ l2) = (reverse S l2)@(reverse S l1).
264 #S #l1 elim l1 [normalize // | #a #tl #Hind #l2 >reverse_cons
265 >reverse_cons >Hind >associative_append //
268 lemma reverse_reverse : ∀S,l. reverse S (reverse S l) = l.
269 #S #l elim l // #a #tl #Hind >reverse_cons >reverse_append
270 >reverse_single >Hind // qed.
273 (***************************** List iterators *********************************)
275 (* In general, an iterator for some data type is an object that enables to
276 traverse its structure performing a given computation.
277 There is a canonical set of iterators over lists deriving from the programming
278 experience. In this section we shall review a few of them, proving some of their
281 A first, famous iterator is the map function, that applies a function f to each
282 element of a list [a1,…, an], building the list [f a1; … ; f an]. *)
284 let rec map (A,B:Type[0]) (f: A → B) (l:list A) on l: list B ≝
285 match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map A B f tl)].
287 (* The map function distributes over append, as can be simply proved by
288 induction over the first list: *)
290 lemma map_append : ∀A,B,f,l1,l2.
291 (map A B f l1) @ (map A B f l2) = map A B f (l1@l2).
293 [ #l2 @refl | #h #t #IH #l2 normalize >IH // ] qed.
295 lemma length_map: ∀A,B,l.∀f:A→B. length ? (map ?? f l) = length ? l.
296 #A #B #l #f elim l // #a #tl #Hind normalize //
299 (* The most important iterator is traditionally called "fold"; we shall only
300 consider the so called fold-right variant, that takes in input a function
301 f:A→B→B, a list [a1; …; an] of elements of type A, a base element b:B and
302 returns the value f a1 (f a2 (… (f an b) …)). *)
304 let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:list A) on l :B ≝
305 match l with [ nil ⇒ b | cons a l ⇒ f a (foldr A B f b l)].
307 (* Many other interesting functions can be defined in terms of foldr. A first
308 interesting example is the filter function, taking in input a boolean test p and
309 a list l and returning the sublist of all elements of l satisfying the test. *)
313 foldr T (list T) (λx,l0.if p x then x::l0 else l0) (nil T).
315 (* As another example, we can generalize the map function to an operation taking
316 in input a binary function f:A → B → C, two lists [a1;… ;an] and [b1;… ;bm] and
317 returning the list [f a1 b1; ˙… ;f an b1;… ;f a1 b_m; … f an bm]. *)
319 definition compose ≝ λA,B,C.λf:A→B→C.λl1,l2.
320 foldr ?? (λi,acc.(map ?? (f i) l2)@acc) [ ] l1.
322 (* The folded function
323 λi,acc.(map ?? (f i) l2)@acc
324 takes in input an element i and an accumulator, and add to the accumulator the
325 values obtained by mapping the partial instantiation (f i) on the elements of
326 l2. This function is iterated over all elements of the first list l1, starting
327 with an empty accumulator. *)
329 (************************************ Vectors *********************************)
331 (* A vector of length n of elements of type A is simply defined in Matita as a
332 record composed by a list of elements of type A, plus a proof that the list has
333 the expected length. Vectors are a paradigmatic example of dependent type, that
334 is of a type whose definition depends on a term. *)
336 record Vector (A:Type[0]) (n:nat): Type[0] ≝
338 len: length ? vec = n }.
340 (* Given a list l we may trivially turn it into a vector of length |l|; we just
341 need to prove that |l|=|l| that follows from the reflexivity of equality. *)
343 lemma Vector_of_list ≝ λA,l.mk_Vector A (|l|) l (refl ??).
345 (* Then we can prove that any vector of zero lenght has no elements: *)
347 lemma vector_nil: ∀A.∀v:Vector A 0.
348 v = Vector_of_list A [ ].
349 #A * * // #a #tl normalize #H destruct
352 (* It is convenient to re-define on vectors the hd, tail and cons operations
355 definition vec_tail ≝ λA.λn.λv:Vector A n.
356 mk_Vector A (pred n) (tail A v) ?.
357 >length_tail >(len A n v) //
360 definition vec_cons ≝ λA.λa.λn.λv:Vector A n.
361 mk_Vector A (S n) (cons A a v) ?.
362 normalize >(len A n v) //
365 definition vec_hd ≝ λA.λn.λv:Vector A (S n). hd A v ?.
366 elim v * [normalize #H destruct | #a #H #eq @a]
369 (* Any vector of lenght (S n) can be expanded as a concatenation of a head and a
370 tail vector of length n *)
372 lemma vec_expand: ∀A,n,v.
373 v = vec_cons A (vec_hd A n v) n (vec_tail A (S n) v).
374 #A #n * #l cases l [normalize in ⊢ (%→?); #H destruct |//]
377 (* Most functions operating on lists can be naturally extended to vectors:
378 interesting cases are vec_append, concatenating vectors, and vec_map, mapping a
379 function f on all elements of the input vector and returning a vector of the
380 same dimension of the input one. *)
382 definition vec_append ≝ λA.λn1,n2.λv1:Vector A n1.λv2: Vector A n2.
383 mk_Vector A (n1+n2) (v1@v2) ?.
384 >length_append >(len … v1) >(len … v2) // qed.
386 definition vec_map ≝ λA,B.λf:A→B.λn.λv:Vector A n.
387 mk_Vector B n (map ?? f v)
388 (trans_eq … (length_map …) (len A n v)).
390 (* Since a vector is automatically coerced, if needed, to the list of its
391 elements, we may simply use typical functions over lists (such as nth) to
392 extract/filter specific components.
394 The following function (change_vec A n v a i) replaces the content of the vector
395 v at position i with the value a. *)
397 let rec change_vec (A:Type[0]) (n:nat) on n ≝
398 match n return λn0.∀v:Vector A n0.A→nat→Vector A n0 with
402 [ O ⇒ vec_cons A a m (vec_tail … v)
403 | S j ⇒ vec_cons A (vec_hd A m v) m (change_vec A m (vec_tail … v) a j)
407 (* A frequent operation between vectors is their comparison. The following
408 lemma allow us to reduce the comparisons of vectors of a same lenght to the
409 comparison of their components. *)
411 lemma eq_vec: ∀A,n.∀v1,v2:Vector A n.
412 (∀i. i < n → nth A v1 i = nth A v2 i) → v1 = v2.
414 [#v1 #v2 #_ >(vector_nil A v1) >(vector_nil A v2) //
415 |#n #Hind #v1 #v2 #H >(vec_expand … v1) >(vec_expand … v2)
416 >(Hind (vec_tail … v1) (vec_tail … v2))
417 [cut (vec_hd A n v1 = vec_hd A n v2) // lapply (H 0 (lt_O_S ?)) -H
418 >(vec_expand … v1) in ⊢ (%→?); >(vec_expand … v2) in ⊢ (%→?);
419 change with (Some ?? = Some ?? → ?) #H destruct (H) @e0
420 |#i #ltin @(H (S i)) @le_S_S //