1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "logic/equality.ma".
16 include "nat/compare.ma".
17 include "list/list.ma".
20 (*** useful definitions and lemmas not really related to Fsub ***)
26 |false \Rightarrow m].
28 lemma le_n_max_m_n: \forall m,n:nat. n \le max m n.
32 [simplify.intros.apply le_n
33 |simplify.intros.autobatch
37 lemma le_m_max_m_n: \forall m,n:nat. m \le max m n.
41 [simplify.intro.assumption
42 |simplify.intros.autobatch
46 lemma max_case : \forall m,n.(max m n) = match (leb m n) with
48 | false \Rightarrow m ].
49 intros;elim m;simplify;reflexivity;
52 definition swap : nat → nat → nat → nat ≝
53 λu,v,x.match (eqb x u) with
55 |false ⇒ match (eqb x v) with
59 lemma swap_left : ∀x,y.(swap x y x) = y.
60 intros;unfold swap;rewrite > eqb_n_n;simplify;reflexivity;
63 lemma swap_right : ∀x,y.(swap x y y) = x.
64 intros;unfold swap;rewrite > eqb_n_n;apply (eqb_elim y x);intro;autobatch;
67 lemma swap_other : \forall x,y,z.(z \neq x) \to (z \neq y) \to (swap x y z) = z.
68 intros;unfold swap;apply (eqb_elim z x);intro;simplify
70 |rewrite > not_eq_to_eqb_false;autobatch]
73 lemma swap_inv : \forall u,v,x.(swap u v (swap u v x)) = x.
74 intros;unfold in match (swap u v x);apply (eqb_elim x u);intro;simplify
75 [autobatch paramodulation
76 |apply (eqb_elim x v);intro;simplify
77 [autobatch paramodulation
78 |apply swap_other;assumption]]
81 lemma swap_inj : ∀u,v,x,y.swap u v x = swap u v y → x = y.
82 intros 4;unfold swap;apply (eqb_elim x u);simplify;intro
83 [apply (eqb_elim y u);simplify;intro
84 [intro;autobatch paramodulation
85 |apply (eqb_elim y v);simplify;intros
86 [autobatch paramodulation
87 |elim H2;symmetry;assumption]]
88 |apply (eqb_elim y u);simplify;intro
89 [apply (eqb_elim x v);simplify;intros;
90 [autobatch paramodulation
92 |apply (eqb_elim x v);simplify;intro
93 [apply (eqb_elim y v);simplify;intros
94 [autobatch paramodulation
95 |elim H1;symmetry;assumption]
96 |apply (eqb_elim y v);simplify;intros
101 definition swap_list : nat \to nat \to (list nat) \to (list nat) ≝
102 \lambda u,v,l.(map ? ? (swap u v) l).
104 let rec posn l x on l ≝
107 | (cons (y:nat) l2) ⇒
110 | false ⇒ S (posn l2 x)]].
112 let rec lookup n l on l ≝
115 | cons (x:nat) l1 ⇒ match (eqb n x) with
117 |false ⇒ (lookup n l1)]].
119 let rec head (A:Type) l on l ≝
122 | (cons (x:A) l2) ⇒ Some A x].
124 definition max_list : (list nat) \to nat \def
125 \lambda l.let rec aux l0 m on l0 \def
128 | (cons n l2) ⇒ (aux l2 (max m n))]
131 let rec remove_nat (x:nat) (l:list nat) on l \def
134 | (cons y l2) ⇒ match (eqb x y) with
135 [ true ⇒ (remove_nat x l2)
136 | false ⇒ (y :: (remove_nat x l2)) ]].
138 lemma in_lookup : \forall x,l.x ∈ l \to (lookup x l = true).
139 intros;elim H;simplify
140 [rewrite > eqb_n_n;reflexivity
141 |rewrite > H2;elim (eqb a a1);reflexivity]
144 lemma lookup_in : \forall x,l.(lookup x l = true) \to x ∈ l.
146 [simplify in H;destruct H
147 |generalize in match H1;simplify;elim (decidable_eq_nat x a)
149 |apply in_list_cons;apply H;
150 rewrite > (not_eq_to_eqb_false ? ? H2) in H3;simplify in H3;assumption]]
153 lemma posn_length : \forall x,vars.x ∈ vars →
154 (posn vars x) < (length ? vars).
156 [elim (not_in_list_nil ? ? H)
157 |inversion H1;intros;destruct;simplify
158 [rewrite > eqb_n_n;simplify;
160 |elim (eqb x a);simplify
162 |apply le_S_S;apply H;assumption]]]
165 lemma in_remove : \forall a,b,l.(a ≠ b) \to a ∈ l \to
166 a ∈ (remove_nat b l).
168 [elim (not_in_list_nil ? ? H1)
169 |inversion H2;intros;destruct;simplify
170 [rewrite > not_eq_to_eqb_false
171 [simplify;apply in_list_head
172 |intro;apply H;symmetry;assumption]
173 |elim (eqb b a1);simplify
175 |apply in_list_cons;apply H1;assumption]]]
178 lemma swap_same : \forall x,y.swap x x y = y.
179 intros;elim (decidable_eq_nat y x)
180 [autobatch paramodulation
181 |rewrite > swap_other;autobatch]
184 lemma not_nat_in_to_lookup_false : ∀l,X. X ∉ l → lookup X l = false.
185 intros 2;elim l;simplify
187 |elim (decidable_eq_nat X a)
188 [elim H1;rewrite > H2;apply in_list_head
189 |rewrite > (not_eq_to_eqb_false ? ? H2);simplify;apply H;intro;apply H1;
190 apply (in_list_cons ? ? ? ? H3);]]
193 lemma lookup_swap : ∀a,b,x,vars.
194 lookup (swap a b x) (swap_list a b vars) =
196 intros 4;elim vars;simplify
198 |elim (decidable_eq_nat x a1)
199 [rewrite > H1;rewrite > eqb_n_n;rewrite > eqb_n_n;simplify;reflexivity
200 |rewrite > (not_eq_to_eqb_false ? ? H1);elim (decidable_eq_nat x a)
201 [rewrite > H;rewrite > H2;rewrite > swap_left;
202 elim (decidable_eq_nat b a1)
203 [rewrite < H3;rewrite > swap_right;
204 rewrite > (not_eq_to_eqb_false b a)
207 |rewrite > (swap_other a b a1)
208 [rewrite > (not_eq_to_eqb_false ? ? H3);simplify;reflexivity
210 |elim (decidable_eq_nat x b)
211 [rewrite > H;rewrite > H3;rewrite > swap_right;
212 elim (decidable_eq_nat a1 a)
213 [rewrite > H4;rewrite > swap_left;
214 rewrite > (not_eq_to_eqb_false a b)
217 |rewrite > (swap_other a b a1)
218 [rewrite > (not_eq_to_eqb_false a a1)
223 |rewrite > H;rewrite > (swap_other a b x)
224 [elim (decidable_eq_nat a a1)
225 [rewrite > H4;rewrite > swap_left;
226 rewrite > (not_eq_to_eqb_false ? ? H3);reflexivity
227 |elim (decidable_eq_nat b a1)
228 [rewrite > H5;rewrite > swap_right;
229 rewrite > (not_eq_to_eqb_false ? ? H2);reflexivity
230 |rewrite > (swap_other a b a1)
231 [rewrite > (not_eq_to_eqb_false ? ? H1);reflexivity
232 |*:intro;autobatch]]]
236 lemma posn_swap : ∀a,b,x,vars.
237 posn vars x = posn (swap_list a b vars) (swap a b x).
238 intros 4;elim vars;simplify
240 |rewrite < H;elim (decidable_eq_nat x a1)
241 [rewrite > H1;do 2 rewrite > eqb_n_n;reflexivity
242 |elim (decidable_eq_nat (swap a b x) (swap a b a1))
244 |rewrite > (not_eq_to_eqb_false ? ? H1);
245 rewrite > (not_eq_to_eqb_false ? ? H2);reflexivity]]]
248 lemma remove_inlist : \forall x,y,l.x ∈ (remove_nat y l) → x ∈ l \land x ≠ y.
249 intros 3;elim l 0;simplify;intro
250 [elim (not_in_list_nil ? ? H)
251 |elim (decidable_eq_nat y a)
252 [rewrite > H in H2;rewrite > eqb_n_n in H2;simplify in H2;
253 rewrite > H in H1;elim (H1 H2);rewrite > H;split
254 [apply in_list_cons;assumption
256 |rewrite > (not_eq_to_eqb_false ? ? H) in H2;simplify in H2;
257 inversion H2;intros;destruct;
259 |elim (H1 H3);split;autobatch]]]
262 lemma inlist_remove : ∀x,y,l.x ∈ l → x ≠ y → x ∈ (remove_nat y l).
264 [elim (not_in_list_nil ? ? H);
265 |simplify;elim (decidable_eq_nat y a)
266 [rewrite > (eq_to_eqb_true ? ? H3);simplify;apply H
267 [inversion H1;intros;destruct;
271 |rewrite > (not_eq_to_eqb_false ? ? H3);simplify;
272 inversion H1;intros;destruct;autobatch]]
275 lemma swap_case: ∀u,v,x.(swap u v x) = u ∨ (swap u v x) = v ∨ (swap u v x = x).
276 intros;unfold in match swap;simplify;elim (eqb x u);simplify
278 |elim (eqb x v);simplify;autobatch]
281 definition distinct_list ≝
282 λl.∀n,m.n < length ? l → m < length ? l → n ≠ m → nth ? l O n ≠ nth ? l O m.
284 lemma posn_nth: ∀l,n.distinct_list l → n < length ? l →
285 posn l (nth ? l O n) = n.
287 [simplify in H1;elim (not_le_Sn_O ? H1)
288 |simplify in H2;generalize in match H2;elim n
289 [simplify;rewrite > eqb_n_n;simplify;reflexivity
290 |simplify;cut (nth ? (a::l1) O (S n1) ≠ nth ? (a::l1) O O)
291 [simplify in Hcut;rewrite > (not_eq_to_eqb_false ? ? Hcut);simplify;
294 |unfold;intros;unfold in H1;lapply (H1 (S n2) (S m));
295 [simplify in Hletin;assumption
296 |2,3:simplify;autobatch
297 |intro;apply H7;destruct H8;reflexivity]
303 |(*** *:autobatch; *)
306 |intro;elim (not_eq_O_S n1);symmetry;assumption]]]]
309 lemma nth_in_list : ∀l,n. n < length ? l → (nth ? l O n) ∈ l.
311 [simplify in H;elim (not_le_Sn_O ? H)
312 |generalize in match H1;elim n;simplify
314 |apply in_list_cons;apply H;simplify in H3;apply (le_S_S_to_le ? ? H3)]]
317 lemma lookup_nth : \forall l,n.n < length ? l \to
318 lookup (nth nat l O n) l = true.
320 [elim (not_le_Sn_O ? H);
321 |generalize in match H1;elim n;simplify
322 [rewrite > eqb_n_n;reflexivity
323 |elim (eqb (nth nat l1 O n1) a);simplify;
325 |apply H;apply le_S_S_to_le;assumption]]]