2 ||M|| This file is part of HELM, an Hypertextual, Electronic
3 ||A|| Library of Mathematics, developed at the Computer Science
4 ||T|| Department of the University of Bologna, Italy.
8 \ / This file is distributed under the terms of the
9 \ / GNU General Public License Version 2
10 V_______________________________________________________________ *)
12 include "basics/types.ma".
13 include "arithmetics/nat.ma".
15 inductive list (A:Type[0]) : Type[0] :=
17 | cons: A -> list A -> list A.
19 notation "hvbox(hd break :: tl)"
20 right associative with precedence 47
23 notation "[ list0 x sep ; ]"
24 non associative with precedence 90
25 for ${fold right @'nil rec acc @{'cons $x $acc}}.
27 notation "hvbox(l1 break @ l2)"
28 right associative with precedence 47
29 for @{'append $l1 $l2 }.
31 interpretation "nil" 'nil = (nil ?).
32 interpretation "cons" 'cons hd tl = (cons ? hd tl).
34 definition not_nil: ∀A:Type[0].list A → Prop ≝
35 λA.λl.match l with [ nil ⇒ True | cons hd tl ⇒ False ].
38 ∀A:Type[0].∀l:list A.∀a:A. a::l ≠ [].
39 #A #l #a @nmk #Heq (change with (not_nil ? (a::l))) >Heq //
43 let rec id_list A (l: list A) on l :=
46 | (cons hd tl) => hd :: id_list A tl ]. *)
48 let rec append A (l1: list A) l2 on l1 ≝
51 | cons hd tl ⇒ hd :: append A tl l2 ].
53 definition hd ≝ λA.λl: list A.λd:A.
54 match l with [ nil ⇒ d | cons a _ ⇒ a].
56 definition tail ≝ λA.λl: list A.
57 match l with [ nil ⇒ [] | cons hd tl ⇒ tl].
59 interpretation "append" 'append l1 l2 = (append ? l1 l2).
61 theorem append_nil: ∀A.∀l:list A.l @ [] = l.
62 #A #l (elim l) normalize // qed.
64 theorem associative_append:
65 ∀A.associative (list A) (append A).
66 #A #l1 #l2 #l3 (elim l1) normalize // qed.
69 ntheorem cons_append_commute:
70 ∀A:Type.∀l1,l2:list A.∀a:A.
71 a :: (l1 @ l2) = (a :: l1) @ l2.
74 theorem append_cons:∀A.∀a:A.∀l,l1.l@(a::l1)=(l@[a])@l1.
77 theorem nil_append_elim: ∀A.∀l1,l2: list A.∀P:?→?→Prop.
78 l1@l2=[] → P (nil A) (nil A) → P l1 l2.
79 #A #l1 #l2 #P (cases l1) normalize //
83 theorem nil_to_nil: ∀A.∀l1,l2:list A.
84 l1@l2 = [] → l1 = [] ∧ l2 = [].
85 #A #l1 #l2 #isnil @(nil_append_elim A l1 l2) /2/
90 let rec map (A,B:Type[0]) (f: A → B) (l:list A) on l: list B ≝
91 match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map A B f tl)].
93 let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:list A) on l :B ≝
94 match l with [ nil ⇒ b | cons a l ⇒ f a (foldr A B f b l)].
98 foldr T (list T) (λx,l0.if_then_else ? (p x) (x::l0) l0) (nil T).
100 lemma filter_true : ∀A,l,a,p. p a = true →
101 filter A p (a::l) = a :: filter A p l.
102 #A #l #a #p #pa (elim l) normalize >pa normalize // qed.
104 lemma filter_false : ∀A,l,a,p. p a = false →
105 filter A p (a::l) = filter A p l.
106 #A #l #a #p #pa (elim l) normalize >pa normalize // qed.
108 theorem eq_map : ∀A,B,f,g,l. (∀x.f x = g x) → map A B f l = map A B g l.
109 #A #B #f #g #l #eqfg (elim l) normalize // qed.
111 let rec dprodl (A:Type[0]) (f:A→Type[0]) (l1:list A) (g:(∀a:A.list (f a))) on l1 ≝
114 | cons a tl ⇒ (map ??(dp ?? a) (g a)) @ dprodl A f tl g
117 (**************************** length ******************************)
119 let rec length (A:Type[0]) (l:list A) on l ≝
122 | cons a tl ⇒ S (length A tl)].
124 notation "|M|" non associative with precedence 60 for @{'norm $M}.
125 interpretation "norm" 'norm l = (length ? l).
127 let rec nth n (A:Type[0]) (l:list A) (d:A) ≝
130 |S m ⇒ nth m A (tail A l) d].
132 (**************************** fold *******************************)
134 let rec fold (A,B:Type[0]) (op:B → B → B) (b:B) (p:A→bool) (f:A→B) (l:list A) on l :B ≝
137 | cons a l ⇒ if_then_else ? (p a) (op (f a) (fold A B op b p f l))
138 (fold A B op b p f l)].
140 notation "\fold [ op , nil ]_{ ident i ∈ l | p} f"
142 for @{'fold $op $nil (λ${ident i}. $p) (λ${ident i}. $f) $l}.
144 notation "\fold [ op , nil ]_{ident i ∈ l } f"
146 for @{'fold $op $nil (λ${ident i}.true) (λ${ident i}. $f) $l}.
148 interpretation "\fold" 'fold op nil p f l = (fold ? ? op nil p f l).
151 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A→B. p a = true →
152 \fold[op,nil]_{i ∈ a::l| p i} (f i) =
153 op (f a) \fold[op,nil]_{i ∈ l| p i} (f i).
154 #A #B #a #l #p #op #nil #f #pa normalize >pa // qed.
157 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f.
158 p a = false → \fold[op,nil]_{i ∈ a::l| p i} (f i) =
159 \fold[op,nil]_{i ∈ l| p i} (f i).
160 #A #B #a #l #p #op #nil #f #pa normalize >pa // qed.
163 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A →B.
164 \fold[op,nil]_{i ∈ l| p i} (f i) =
165 \fold[op,nil]_{i ∈ (filter A p l)} (f i).
166 #A #B #a #l #p #op #nil #f elim l //
167 #a #tl #Hind cases(true_or_false (p a)) #pa
168 [ >filter_true // > fold_true // >fold_true //
169 | >filter_false // >fold_false // ]
172 record Aop (A:Type[0]) (nil:A) : Type[0] ≝
174 nill:∀a. op nil a = a;
175 nilr:∀a. op a nil = a;
176 assoc: ∀a,b,c.op a (op b c) = op (op a b) c
179 theorem fold_sum: ∀A,B. ∀I,J:list A.∀nil.∀op:Aop B nil.∀f.
180 op (\fold[op,nil]_{i∈I} (f i)) (\fold[op,nil]_{i∈J} (f i)) =
181 \fold[op,nil]_{i∈(I@J)} (f i).
182 #A #B #I #J #nil #op #f (elim I) normalize
183 [>nill //|#a #tl #Hind <assoc //]