1 | (* Various small homeless results. *) |
---|
2 | |
---|
3 | include "Clight/TypeComparison.ma". |
---|
4 | include "Clight/Csem.ma". |
---|
5 | include "common/Pointers.ma". |
---|
6 | include "basics/jmeq.ma". |
---|
7 | include "basics/star.ma". (* well-founded relations *) |
---|
8 | include "common/IOMonad.ma". |
---|
9 | include "common/IO.ma". |
---|
10 | include "basics/lists/listb.ma". |
---|
11 | include "basics/lists/list.ma". |
---|
12 | |
---|
13 | |
---|
14 | (* --------------------------------------------------------------------------- *) |
---|
15 | (* [cthulhu] plays the same role as daemon. It should be droppable. *) |
---|
16 | (* --------------------------------------------------------------------------- *) |
---|
17 | |
---|
18 | axiom cthulhu : ∀A:Prop. A. (* Because of the nightmares. *) |
---|
19 | |
---|
20 | (* --------------------------------------------------------------------------- *) |
---|
21 | (* Misc. *) |
---|
22 | (* --------------------------------------------------------------------------- *) |
---|
23 | |
---|
24 | lemma eq_intsize_identity : ∀id. eq_intsize id id = true. |
---|
25 | * normalize // |
---|
26 | qed. |
---|
27 | |
---|
28 | lemma sz_eq_identity : ∀s. sz_eq_dec s s = inl ?? (refl ? s). |
---|
29 | * normalize // |
---|
30 | qed. |
---|
31 | |
---|
32 | lemma type_eq_identity : ∀t. type_eq t t = true. |
---|
33 | #t normalize elim (type_eq_dec t t) |
---|
34 | [ 1: #Heq normalize // |
---|
35 | | 2: #H destruct elim H #Hcontr elim (Hcontr (refl ? t)) ] qed. |
---|
36 | |
---|
37 | lemma type_neq_not_identity : ∀t1, t2. t1 ≠ t2 → type_eq t1 t2 = false. |
---|
38 | #t1 #t2 #Hneq normalize elim (type_eq_dec t1 t2) |
---|
39 | [ 1: #Heq destruct elim Hneq #Hcontr elim (Hcontr (refl ? t2)) |
---|
40 | | 2: #Hneq' normalize // ] qed. |
---|
41 | |
---|
42 | lemma le_S_O_absurd : ∀x:nat. S x ≤ 0 → False. /2 by absurd/ qed. |
---|
43 | |
---|
44 | lemma some_inj : ∀A : Type[0]. ∀a,b : A. Some ? a = Some ? b → a = b. #A #a #b #H destruct (H) @refl qed. |
---|
45 | |
---|
46 | lemma prod_eq_lproj : ∀A,B : Type[0]. ∀a : A. ∀b : B. ∀c : A × B. 〈a,b〉 = c → a = \fst c. |
---|
47 | #A #B #a #b * #a' #b' #H destruct @refl |
---|
48 | qed. |
---|
49 | |
---|
50 | lemma prod_eq_rproj : ∀A,B : Type[0]. ∀a : A. ∀b : B. ∀c : A × B. 〈a,b〉 = c → b = \snd c. |
---|
51 | #A #B #a #b * #a' #b' #H destruct @refl |
---|
52 | qed. |
---|
53 | |
---|
54 | lemma bindIO_Error : ∀err,f. bindIO io_out io_in (val×trace) (trace×state) (Error … err) f = Wrong io_out io_in (trace×state) err. |
---|
55 | // qed. |
---|
56 | |
---|
57 | lemma bindIO_Value : ∀v,f. bindIO io_out io_in (val×trace) (trace×state) (Value … v) f = (f v). |
---|
58 | // qed. |
---|
59 | |
---|
60 | lemma bindIO_elim : |
---|
61 | ∀A. |
---|
62 | ∀P : (IO io_out io_in A) → Prop. |
---|
63 | ∀e : res A. (*IO io_out io_in A.*) |
---|
64 | ∀f. |
---|
65 | (∀v. (e:IO io_out io_in A) = OK … v → P (f v)) → |
---|
66 | (∀err. (e:IO io_out io_in A) = Error … err → P (Wrong ??? err)) → |
---|
67 | P (bindIO io_out io_in ? A (e:IO io_out io_in A) f). |
---|
68 | #A #P * try /2/ qed. |
---|
69 | |
---|
70 | lemma opt_to_io_Value : |
---|
71 | ∀A,B,C,err,x,res. opt_to_io A B C err x = return res → x = Some ? res. |
---|
72 | #A #B #C #err #x cases x normalize |
---|
73 | [ 1: #res #Habsurd destruct |
---|
74 | | 2: #c #res #Heq destruct @refl ] |
---|
75 | qed. |
---|
76 | |
---|
77 | lemma some_inversion : |
---|
78 | ∀A,B:Type[0]. |
---|
79 | ∀e:option A. |
---|
80 | ∀res:B. |
---|
81 | ∀f:A → option B. |
---|
82 | match e with |
---|
83 | [ None ⇒ None ? |
---|
84 | | Some x ⇒ f x ] = Some ? res → |
---|
85 | ∃x. e = Some ? x ∧ f x = Some ? res. |
---|
86 | #A #B #e #res #f cases e normalize nodelta |
---|
87 | [ 1: #Habsurd destruct (Habsurd) |
---|
88 | | 2: #a #Heq %{a} @conj >Heq @refl ] |
---|
89 | qed. |
---|
90 | |
---|
91 | lemma res_inversion : |
---|
92 | ∀A,B:Type[0]. |
---|
93 | ∀e:option A. |
---|
94 | ∀errmsg. |
---|
95 | ∀result:B. |
---|
96 | ∀f:A → res B. |
---|
97 | match e with |
---|
98 | [ None ⇒ Error ? errmsg |
---|
99 | | Some x ⇒ f x ] = OK ? result → |
---|
100 | ∃x. e = Some ? x ∧ f x = OK ? result. |
---|
101 | #A #B #e #errmsg #result #f cases e normalize nodelta |
---|
102 | [ 1: #Habsurd destruct (Habsurd) |
---|
103 | | 2: #a #Heq %{a} @conj >Heq @refl ] |
---|
104 | qed. |
---|
105 | |
---|
106 | lemma cons_inversion : |
---|
107 | ∀A,B:Type[0]. |
---|
108 | ∀e:list A. |
---|
109 | ∀res:B. |
---|
110 | ∀f:A → list A → option B. |
---|
111 | match e with |
---|
112 | [ nil ⇒ None ? |
---|
113 | | cons hd tl ⇒ f hd tl ] = Some ? res → |
---|
114 | ∃hd,tl. e = hd::tl ∧ f hd tl = Some ? res. |
---|
115 | #A #B #e #res #f cases e normalize nodelta |
---|
116 | [ 1: #Habsurd destruct (Habsurd) |
---|
117 | | 2: #hd #tl #Heq %{hd} %{tl} @conj >Heq @refl ] |
---|
118 | qed. |
---|
119 | |
---|
120 | lemma if_opt_inversion : |
---|
121 | ∀A:Type[0]. |
---|
122 | ∀x. |
---|
123 | ∀y:A. |
---|
124 | ∀e:bool. |
---|
125 | (if e then |
---|
126 | x |
---|
127 | else |
---|
128 | None ?) = Some ? y → |
---|
129 | e = true ∧ x = Some ? y. |
---|
130 | #A #x #y * normalize |
---|
131 | #H destruct @conj @refl |
---|
132 | qed. |
---|
133 | |
---|
134 | lemma opt_to_res_inversion : |
---|
135 | ∀A:Type[0]. ∀errmsg. ∀opt. ∀val. opt_to_res A errmsg opt = OK ? val → |
---|
136 | opt = Some ? val. |
---|
137 | #A #errmsg * |
---|
138 | [ 1: #val normalize #Habsurd destruct |
---|
139 | | 2: #res #val normalize #Heq destruct @refl ] |
---|
140 | qed. |
---|
141 | |
---|
142 | lemma andb_inversion : |
---|
143 | ∀a,b : bool. andb a b = true → a = true ∧ b = true. |
---|
144 | * * normalize /2 by conj, refl/ qed. |
---|
145 | |
---|
146 | lemma identifier_eq_i_i : ∀tag,i. ∃pf. identifier_eq tag i i = inl … pf. |
---|
147 | #tag #i cases (identifier_eq ? i i) |
---|
148 | [ 1: #H %{H} @refl |
---|
149 | | 2: * #Habsurd @(False_ind … (Habsurd … (refl ? i))) ] |
---|
150 | qed. |
---|
151 | |
---|
152 | (* --------------------------------------------------------------------------- *) |
---|
153 | (* Useful facts on various boolean operations. *) |
---|
154 | (* --------------------------------------------------------------------------- *) |
---|
155 | |
---|
156 | lemma andb_lsimpl_true : ∀x. andb true x = x. // qed. |
---|
157 | lemma andb_lsimpl_false : ∀x. andb false x = false. normalize // qed. |
---|
158 | lemma andb_comm : ∀x,y. andb x y = andb y x. // qed. |
---|
159 | lemma notb_true : notb true = false. // qed. |
---|
160 | lemma notb_false : notb false = true. % #H destruct qed. |
---|
161 | lemma notb_fold : ∀x. if x then false else true = (¬x). // qed. |
---|
162 | |
---|
163 | (* --------------------------------------------------------------------------- *) |
---|
164 | (* Useful facts on Z. *) |
---|
165 | (* --------------------------------------------------------------------------- *) |
---|
166 | |
---|
167 | lemma Zltb_to_Zleb_true : ∀a,b. Zltb a b = true → Zleb a b = true. |
---|
168 | #a #b #Hltb lapply (Zltb_true_to_Zlt … Hltb) #Hlt @Zle_to_Zleb_true |
---|
169 | /3 by Zlt_to_Zle, transitive_Zle/ qed. |
---|
170 | |
---|
171 | lemma Zle_to_Zle_to_eq : ∀a,b. Zle a b → Zle b a → a = b. |
---|
172 | #a #b elim b |
---|
173 | [ 1: elim a // #a' normalize [ 1: @False_ind | 2: #_ @False_ind ] ] |
---|
174 | #b' elim a normalize |
---|
175 | [ 1: #_ @False_ind |
---|
176 | | 2: #a' #H1 #H2 >(le_to_le_to_eq … H1 H2) @refl |
---|
177 | | 3: #a' #_ @False_ind |
---|
178 | | 4: @False_ind |
---|
179 | | 5: #a' @False_ind |
---|
180 | | 6: #a' #H2 #H1 >(le_to_le_to_eq … H1 H2) @refl |
---|
181 | ] qed. |
---|
182 | |
---|
183 | lemma Zleb_true_to_Zleb_true_to_eq : ∀a,b. (Zleb a b = true) → (Zleb b a = true) → a = b. |
---|
184 | #a #b #H1 #H2 |
---|
185 | /3 by Zle_to_Zle_to_eq, Zleb_true_to_Zle/ |
---|
186 | qed. |
---|
187 | |
---|
188 | lemma Zltb_dec : ∀a,b. (Zltb a b = true) ∨ (Zltb a b = false ∧ Zleb b a = true). |
---|
189 | #a #b |
---|
190 | lapply (Z_compare_to_Prop … a b) |
---|
191 | cases a |
---|
192 | [ 1: | 2,3: #a' ] |
---|
193 | cases b |
---|
194 | whd in match (Z_compare OZ OZ); normalize nodelta |
---|
195 | [ 2,3,5,6,8,9: #b' ] |
---|
196 | whd in match (Zleb ? ?); |
---|
197 | try /3 by or_introl, or_intror, conj, I, refl/ |
---|
198 | whd in match (Zltb ??); |
---|
199 | whd in match (Zleb ??); #_ |
---|
200 | [ 1: cases (decidable_le (succ a') b') |
---|
201 | [ 1: #H lapply (le_to_leb_true … H) #H %1 assumption |
---|
202 | | 2: #Hnotle lapply (not_le_to_lt … Hnotle) #Hlt %2 @conj try @le_to_leb_true |
---|
203 | /2 by monotonic_pred/ @(not_le_to_leb_false … Hnotle) ] |
---|
204 | | 2: cases (decidable_le (succ b') a') |
---|
205 | [ 1: #H lapply (le_to_leb_true … H) #H %1 assumption |
---|
206 | | 2: #Hnotle lapply (not_le_to_lt … Hnotle) #Hlt %2 @conj try @le_to_leb_true |
---|
207 | /2 by monotonic_pred/ @(not_le_to_leb_false … Hnotle) ] |
---|
208 | ] qed. |
---|
209 | |
---|
210 | lemma Zleb_unsigned_OZ : ∀bv. Zleb OZ (Z_of_unsigned_bitvector 16 bv) = true. |
---|
211 | #bv elim bv try // #n' * #tl normalize /2/ qed. |
---|
212 | |
---|
213 | lemma Zltb_unsigned_OZ : ∀bv. Zltb (Z_of_unsigned_bitvector 16 bv) OZ = false. |
---|
214 | #bv elim bv try // #n' * #tl normalize /2/ qed. |
---|
215 | |
---|
216 | lemma Z_of_unsigned_not_neg : ∀bv. |
---|
217 | (Z_of_unsigned_bitvector 16 bv = OZ) ∨ (∃p. Z_of_unsigned_bitvector 16 bv = pos p). |
---|
218 | #bv elim bv |
---|
219 | [ 1: normalize %1 @refl |
---|
220 | | 2: #n #hd #tl #Hind cases hd |
---|
221 | [ 1: normalize %2 /2 by ex_intro/ |
---|
222 | | 2: cases Hind normalize [ 1: #H %1 @H | 2: * #p #H >H %2 %{p} @refl ] |
---|
223 | ] |
---|
224 | ] qed. |
---|
225 | |
---|
226 | lemma free_not_in_bounds : ∀x. if Zleb (pos one) x |
---|
227 | then Zltb x OZ |
---|
228 | else false = false. |
---|
229 | #x lapply (Zltb_to_Zleb_true x OZ) |
---|
230 | elim (Zltb_dec … x OZ) |
---|
231 | [ 1: #Hlt >Hlt #H lapply (H (refl ??)) elim x |
---|
232 | [ 2,3: #x' ] normalize in ⊢ (% → ?); |
---|
233 | [ 1: #Habsurd destruct (Habsurd) |
---|
234 | | 2,3: #_ @refl ] |
---|
235 | | 2: * #Hlt #Hle #_ >Hlt cases (Zleb (pos one) x) @refl ] |
---|
236 | qed. |
---|
237 | |
---|
238 | lemma free_not_valid : ∀x. if Zleb (pos one) x |
---|
239 | then Zltb x OZ |
---|
240 | else false = false. |
---|
241 | #x |
---|
242 | cut (Zle (pos one) x ∧ Zlt x OZ → False) |
---|
243 | [ * #H1 #H2 lapply (transitive_Zle … (monotonic_Zle_Zsucc … H1) (Zlt_to_Zle … H2)) #H @H ] #Hguard |
---|
244 | cut (Zleb (pos one) x = true ∧ Zltb x OZ = true → False) |
---|
245 | [ * #H1 #H2 @Hguard @conj /2 by Zleb_true_to_Zle/ @Zltb_true_to_Zlt assumption ] |
---|
246 | cases (Zleb (pos one) x) cases (Zltb x OZ) |
---|
247 | /2 by refl/ #H @(False_ind … (H (conj … (refl ??) (refl ??)))) |
---|
248 | qed. |
---|
249 | |
---|
250 | (* follows from (0 ≤ a < b → mod a b = a) *) |
---|
251 | axiom Z_of_unsigned_bitvector_of_small_Z : |
---|
252 | ∀z. OZ ≤ z → z < Z_two_power_nat 16 → Z_of_unsigned_bitvector 16 (bitvector_of_Z 16 z) = z. |
---|
253 | |
---|
254 | theorem Zle_to_Zlt_to_Zlt: ∀n,m,p:Z. n ≤ m → m < p → n < p. |
---|
255 | #n #m #p #Hle #Hlt /4 by monotonic_Zle_Zplus_r, Zle_to_Zlt, Zlt_to_Zle, transitive_Zle/ |
---|
256 | qed. |
---|
257 | |
---|
258 | (* --------------------------------------------------------------------------- *) |
---|
259 | (* Useful facts on blocks. *) |
---|
260 | (* --------------------------------------------------------------------------- *) |
---|
261 | |
---|
262 | lemma not_eq_block : ∀b1,b2. b1 ≠ b2 → eq_block b1 b2 = false. |
---|
263 | #b1 #b2 #Hneq |
---|
264 | @(eq_block_elim … b1 b2) |
---|
265 | [ 1: #Heq destruct elim Hneq #H @(False_ind … (H (refl ??))) |
---|
266 | | 2: #_ @refl ] qed. |
---|
267 | |
---|
268 | lemma not_eq_block_rev : ∀b1,b2. b2 ≠ b1 → eq_block b1 b2 = false. |
---|
269 | #b1 #b2 #Hneq |
---|
270 | @(eq_block_elim … b1 b2) |
---|
271 | [ 1: #Heq destruct elim Hneq #H @(False_ind … (H (refl ??))) |
---|
272 | | 2: #_ @refl ] qed. |
---|
273 | |
---|
274 | definition block_DeqSet : DeqSet ≝ mk_DeqSet block eq_block ?. |
---|
275 | * #r1 #id1 * #r2 #id2 @(eqZb_elim … id1 id2) |
---|
276 | [ 1: #Heq >Heq cases r1 cases r2 normalize |
---|
277 | >eqZb_z_z normalize try // @conj #H destruct (H) |
---|
278 | try @refl |
---|
279 | | 2: #Hneq cases r1 cases r2 normalize |
---|
280 | >(eqZb_false … Hneq) normalize @conj |
---|
281 | #H destruct (H) elim Hneq #H @(False_ind … (H (refl ??))) |
---|
282 | ] qed. |
---|
283 | |
---|
284 | (* --------------------------------------------------------------------------- *) |
---|
285 | (* General results on lists. *) |
---|
286 | (* --------------------------------------------------------------------------- *) |
---|
287 | |
---|
288 | (* If mem succeeds in finding an element, then the list can be partitioned around this element. *) |
---|
289 | lemma list_mem_split : ∀A. ∀l : list A. ∀x : A. mem … x l → ∃l1,l2. l = l1 @ [x] @ l2. |
---|
290 | #A #l elim l |
---|
291 | [ 1: normalize #x @False_ind |
---|
292 | | 2: #hd #tl #Hind #x whd in ⊢ (% → ?); * |
---|
293 | [ 1: #Heq %{(nil ?)} %{tl} destruct @refl |
---|
294 | | 2: #Hmem lapply (Hind … Hmem) * #l1 * #l2 #Heq_tl >Heq_tl |
---|
295 | %{(hd :: l1)} %{l2} @refl |
---|
296 | ] |
---|
297 | ] qed. |
---|
298 | |
---|
299 | lemma cons_to_append : ∀A. ∀hd : A. ∀l : list A. hd :: l = [hd] @ l. #A #hd #l @refl qed. |
---|
300 | |
---|
301 | lemma fold_append : |
---|
302 | ∀A,B:Type[0]. ∀l1, l2 : list A. ∀f:A → B → B. ∀seed. foldr ?? f seed (l1 @ l2) = foldr ?? f (foldr ?? f seed l2) l1. |
---|
303 | #A #B #l1 elim l1 // |
---|
304 | #hd #tl #Hind #l2 #f #seed normalize >Hind @refl |
---|
305 | qed. |
---|
306 | |
---|
307 | lemma filter_append : ∀A:Type[0]. ∀l1,l2 : list A. ∀f. filter ? f (l1 @ l2) = (filter ? f l1) @ (filter ? f l2). |
---|
308 | #A #l1 elim l1 // |
---|
309 | #hd #tl #Hind #l2 #f |
---|
310 | >cons_to_append >associative_append |
---|
311 | normalize cases (f hd) normalize |
---|
312 | <Hind @refl |
---|
313 | qed. |
---|
314 | |
---|
315 | lemma filter_cons_unfold : ∀A:Type[0]. ∀f. ∀hd,tl. |
---|
316 | filter ? f (hd :: tl) = |
---|
317 | if f hd then |
---|
318 | (hd :: filter A f tl) |
---|
319 | else |
---|
320 | (filter A f tl). |
---|
321 | #A #f #hd #tl elim tl // qed. |
---|
322 | |
---|
323 | |
---|
324 | lemma filter_elt_empty : ∀A:DeqSet. ∀elt,l. filter (carr A) (λx.¬(x==elt)) l = [ ] → All (carr A) (λx. x = elt) l. |
---|
325 | #A #elt #l elim l |
---|
326 | [ 1: // |
---|
327 | | 2: #hd #tl #Hind >filter_cons_unfold |
---|
328 | lapply (eqb_true A hd elt) |
---|
329 | cases (hd==elt) normalize nodelta |
---|
330 | [ 2: #_ #Habsurd destruct |
---|
331 | | 1: * #H1 #H2 #Heq lapply (Hind Heq) #Hall whd @conj // |
---|
332 | @H1 @refl |
---|
333 | ] |
---|
334 | ] qed. |
---|
335 | |
---|
336 | lemma nil_append : ∀A. ∀(l : list A). [ ] @ l = l. // qed. |
---|
337 | |
---|
338 | alias id "meml" = "cic:/matita/basics/lists/list/mem.fix(0,2,1)". |
---|
339 | |
---|
340 | lemma mem_append : ∀A:Type[0]. ∀elt : A. ∀l1,l2. mem … elt (l1 @ l2) ↔ (mem … elt l1) ∨ (mem … elt l2). |
---|
341 | #A #elt #l1 elim l1 |
---|
342 | [ 1: #l2 normalize @conj [ 1: #H %2 @H | 2: * [ 1: @False_ind | 2: #H @H ] ] |
---|
343 | | 2: #hd #tl #Hind #l2 @conj |
---|
344 | [ 1: whd in match (meml ???); * |
---|
345 | [ 1: #Heq >Heq %1 normalize %1 @refl |
---|
346 | | 2: #H1 lapply (Hind l2) * #HA #HB normalize |
---|
347 | elim (HA H1) |
---|
348 | [ 1: #H %1 %2 assumption | 2: #H %2 assumption ] |
---|
349 | ] |
---|
350 | | 2: normalize * |
---|
351 | [ 1: * /2 by or_introl, or_intror/ |
---|
352 | #H %2 elim (Hind l2) #_ #H' @H' %1 @H |
---|
353 | | 2: #H %2 elim (Hind l2) #_ #H' @H' %2 @H ] |
---|
354 | ] |
---|
355 | ] qed. |
---|
356 | |
---|
357 | lemma mem_append_forward : ∀A:Type[0]. ∀elt : A. ∀l1,l2. mem … elt (l1 @ l2) → (mem … elt l1) ∨ (mem … elt l2). |
---|
358 | #A #elt #l1 #l2 #H elim (mem_append … elt l1 l2) #H' #_ @H' @H qed. |
---|
359 | |
---|
360 | lemma mem_append_backwards : ∀A:Type[0]. ∀elt : A. ∀l1,l2. (mem … elt l1) ∨ (mem … elt l2) → mem … elt (l1 @ l2) . |
---|
361 | #A #elt #l1 #l2 #H elim (mem_append … elt l1 l2) #_ #H' @H' @H qed. |
---|
362 | |
---|
363 | (* "Observational" equivalence on list implies concrete equivalence. Useful to |
---|
364 | * prove equality from some reasoning on indexings. Needs a particular induction |
---|
365 | * principle. *) |
---|
366 | |
---|
367 | let rec double_list_ind |
---|
368 | (A : Type[0]) |
---|
369 | (P : list A → list A → Prop) |
---|
370 | (base_nil : P [ ] [ ]) |
---|
371 | (base_l1 : ∀hd1,l1. P (hd1::l1) [ ]) |
---|
372 | (base_l2 : ∀hd2,l2. P [ ] (hd2::l2)) |
---|
373 | (ind : ∀hd1,hd2,tl1,tl2. P tl1 tl2 → P (hd1 :: tl1) (hd2 :: tl2)) |
---|
374 | (l1, l2 : list A) on l1 ≝ |
---|
375 | match l1 with |
---|
376 | [ nil ⇒ |
---|
377 | match l2 with |
---|
378 | [ nil ⇒ base_nil |
---|
379 | | cons hd2 tl2 ⇒ base_l2 hd2 tl2 ] |
---|
380 | | cons hd1 tl1 ⇒ |
---|
381 | match l2 with |
---|
382 | [ nil ⇒ base_l1 hd1 tl1 |
---|
383 | | cons hd2 tl2 ⇒ |
---|
384 | ind hd1 hd2 tl1 tl2 (double_list_ind A P base_nil base_l1 base_l2 ind tl1 tl2) |
---|
385 | ] |
---|
386 | ]. |
---|
387 | |
---|
388 | lemma nth_eq_tl : |
---|
389 | ∀A:Type[0]. ∀l1,l2:list A. ∀hd1,hd2. |
---|
390 | (∀i. nth_opt A i (hd1::l1) = nth_opt A i (hd2::l2)) → |
---|
391 | (∀i. nth_opt A i l1 = nth_opt A i l2). |
---|
392 | #A #l1 #l2 @(double_list_ind … l1 l2) |
---|
393 | [ 1: #hd1 #hd2 #_ #i elim i try /2/ |
---|
394 | | 2: #hd1' #l1' #hd1 #hd2 #H lapply (H 1) normalize #Habsurd destruct |
---|
395 | | 3: #hd2' #l2' #hd1 #hd2 #H lapply (H 1) normalize #Habsurd destruct |
---|
396 | | 4: #hd1' #hd2' #tl1' #tl2' #H #hd1 #hd2 |
---|
397 | #Hind |
---|
398 | @(λi. Hind (S i)) |
---|
399 | ] qed. |
---|
400 | |
---|
401 | |
---|
402 | lemma nth_eq_to_eq : |
---|
403 | ∀A:Type[0]. ∀l1,l2:list A. (∀i. nth_opt A i l1 = nth_opt A i l2) → l1 = l2. |
---|
404 | #A #l1 elim l1 |
---|
405 | [ 1: #l2 #H lapply (H 0) normalize |
---|
406 | cases l2 |
---|
407 | [ 1: // |
---|
408 | | 2: #hd2 #tl2 normalize #Habsurd destruct ] |
---|
409 | | 2: #hd1 #tl1 #Hind * |
---|
410 | [ 1: #H lapply (H 0) normalize #Habsurd destruct |
---|
411 | | 2: #hd2 #tl2 #H lapply (H 0) normalize #Heq destruct (Heq) |
---|
412 | >(Hind tl2) try @refl @(nth_eq_tl … H) |
---|
413 | ] |
---|
414 | ] qed. |
---|
415 | |
---|
416 | (* --------------------------------------------------------------------------- *) |
---|
417 | (* General results on vectors. *) |
---|
418 | (* --------------------------------------------------------------------------- *) |
---|
419 | |
---|
420 | (* copied from AssemblyProof, TODO get rid of the redundant stuff. *) |
---|
421 | lemma Vector_O: ∀A. ∀v: Vector A 0. v ≃ VEmpty A. |
---|
422 | #A #v generalize in match (refl … 0); cases v in ⊢ (??%? → ?%%??); // |
---|
423 | #n #hd #tl #abs @⊥ destruct (abs) |
---|
424 | qed. |
---|
425 | |
---|
426 | lemma Vector_Sn: ∀A. ∀n.∀v:Vector A (S n). |
---|
427 | ∃hd.∃tl.v ≃ VCons A n hd tl. |
---|
428 | #A #n #v generalize in match (refl … (S n)); cases v in ⊢ (??%? → ??(λ_.??(λ_.?%%??))); |
---|
429 | [ #abs @⊥ destruct (abs) |
---|
430 | | #m #hd #tl #EQ <(injective_S … EQ) %[@hd] %[@tl] // ] |
---|
431 | qed. |
---|
432 | |
---|
433 | lemma vector_append_zero: |
---|
434 | ∀A,m. |
---|
435 | ∀v: Vector A m. |
---|
436 | ∀q: Vector A 0. |
---|
437 | v = q@@v. |
---|
438 | #A #m #v #q |
---|
439 | >(Vector_O A q) % |
---|
440 | qed. |
---|
441 | |
---|
442 | corollary prod_vector_zero_eq_left: |
---|
443 | ∀A, n. |
---|
444 | ∀q: Vector A O. |
---|
445 | ∀r: Vector A n. |
---|
446 | 〈q, r〉 = 〈[[ ]], r〉. |
---|
447 | #A #n #q #r |
---|
448 | generalize in match (Vector_O A q …); |
---|
449 | #hyp |
---|
450 | >hyp in ⊢ (??%?); |
---|
451 | % |
---|
452 | qed. |
---|
453 | |
---|
454 | lemma vsplit_eq : ∀A. ∀m,n. ∀v : Vector A ((S m) + n). ∃v1:Vector A (S m). ∃v2:Vector A n. v = v1 @@ v2. |
---|
455 | # A #m #n elim m |
---|
456 | [ 1: normalize #v |
---|
457 | elim (Vector_Sn ?? v) #hd * #tl #Eq |
---|
458 | @(ex_intro … (hd ::: [[]])) @(ex_intro … tl) |
---|
459 | >Eq normalize // |
---|
460 | | 2: #n' #Hind #v |
---|
461 | elim (Vector_Sn ?? v) #hd * #tl #Eq |
---|
462 | elim (Hind tl) |
---|
463 | #tl1 * #tl2 #Eq_tl |
---|
464 | @(ex_intro … (hd ::: tl1)) |
---|
465 | @(ex_intro … tl2) |
---|
466 | destruct normalize // |
---|
467 | ] qed. |
---|
468 | |
---|
469 | lemma vsplit_zero: |
---|
470 | ∀A,m. |
---|
471 | ∀v: Vector A m. |
---|
472 | 〈[[]], v〉 = vsplit A 0 m v. |
---|
473 | #A #m #v |
---|
474 | elim v |
---|
475 | [ % |
---|
476 | | #n #hd #tl #ih |
---|
477 | normalize in ⊢ (???%); % |
---|
478 | ] |
---|
479 | qed. |
---|
480 | |
---|
481 | lemma vsplit_zero2: |
---|
482 | ∀A,m. |
---|
483 | ∀v: Vector A m. |
---|
484 | 〈[[]], v〉 = vsplit' A 0 m v. |
---|
485 | #A #m #v |
---|
486 | elim v |
---|
487 | [ % |
---|
488 | | #n #hd #tl #ih |
---|
489 | normalize in ⊢ (???%); % |
---|
490 | ] |
---|
491 | qed. |
---|
492 | |
---|
493 | lemma vsplit_eq2 : ∀A. ∀m,n : nat. ∀v : Vector A (m + n). ∃v1:Vector A m. ∃v2:Vector A n. v = v1 @@ v2. |
---|
494 | # A #m #n elim m |
---|
495 | [ 1: normalize #v @(ex_intro … (VEmpty ?)) @(ex_intro … v) normalize // |
---|
496 | | 2: #n' #Hind #v |
---|
497 | elim (Vector_Sn ?? v) #hd * #tl #Eq |
---|
498 | elim (Hind tl) |
---|
499 | #tl1 * #tl2 #Eq_tl |
---|
500 | @(ex_intro … (hd ::: tl1)) |
---|
501 | @(ex_intro … tl2) |
---|
502 | destruct normalize // |
---|
503 | ] qed. |
---|
504 | |
---|
505 | (* This is not very nice. Note that this axiom was copied verbatim from ASM/AssemblyProof.ma. |
---|
506 | * TODO sync with AssemblyProof.ma, in a better world we shouldn't have to copy all of this. *) |
---|
507 | axiom vsplit_succ: |
---|
508 | ∀A, m, n. |
---|
509 | ∀l: Vector A m. |
---|
510 | ∀r: Vector A n. |
---|
511 | ∀v: Vector A (m + n). |
---|
512 | ∀hd. |
---|
513 | v = l@@r → (〈l, r〉 = vsplit ? m n v → 〈hd:::l, r〉 = vsplit ? (S m) n (hd:::v)). |
---|
514 | |
---|
515 | axiom vsplit_succ2: |
---|
516 | ∀A, m, n. |
---|
517 | ∀l: Vector A m. |
---|
518 | ∀r: Vector A n. |
---|
519 | ∀v: Vector A (m + n). |
---|
520 | ∀hd. |
---|
521 | v = l@@r → (〈l, r〉 = vsplit' ? m n v → 〈hd:::l, r〉 = vsplit' ? (S m) n (hd:::v)). |
---|
522 | |
---|
523 | lemma vsplit_prod2: |
---|
524 | ∀A,m,n. |
---|
525 | ∀p: Vector A (m + n). |
---|
526 | ∀v: Vector A m. |
---|
527 | ∀q: Vector A n. |
---|
528 | p = v@@q → 〈v, q〉 = vsplit' A m n p. |
---|
529 | #A #m |
---|
530 | elim m |
---|
531 | [ #n #p #v #q #hyp |
---|
532 | >hyp <(vector_append_zero A n q v) |
---|
533 | >(prod_vector_zero_eq_left A …) |
---|
534 | @vsplit_zero2 |
---|
535 | | #r #ih #n #p #v #q #hyp |
---|
536 | >hyp |
---|
537 | cases (Vector_Sn A r v) |
---|
538 | #hd #exists |
---|
539 | cases exists |
---|
540 | #tl #jmeq >jmeq |
---|
541 | @vsplit_succ2 [1: % |2: @ih % ] |
---|
542 | ] |
---|
543 | qed. |
---|
544 | |
---|
545 | lemma vsplit_prod: |
---|
546 | ∀A,m,n. |
---|
547 | ∀p: Vector A (m + n). |
---|
548 | ∀v: Vector A m. |
---|
549 | ∀q: Vector A n. |
---|
550 | p = v@@q → 〈v, q〉 = vsplit A m n p. |
---|
551 | #A #m |
---|
552 | elim m |
---|
553 | [ #n #p #v #q #hyp |
---|
554 | >hyp <(vector_append_zero A n q v) |
---|
555 | >(prod_vector_zero_eq_left A …) |
---|
556 | @vsplit_zero |
---|
557 | | #r #ih #n #p #v #q #hyp |
---|
558 | >hyp |
---|
559 | cases (Vector_Sn A r v) |
---|
560 | #hd #exists |
---|
561 | cases exists |
---|
562 | #tl #jmeq >jmeq |
---|
563 | @vsplit_succ [1: % |2: @ih % ] |
---|
564 | ] |
---|
565 | qed. |
---|
566 | |
---|
567 | |
---|
568 | (* --------------------------------------------------------------------------- *) |
---|
569 | (* Generic properties of equivalence relations *) |
---|
570 | (* --------------------------------------------------------------------------- *) |
---|
571 | |
---|
572 | lemma triangle_diagram : |
---|
573 | ∀acctype : Type[0]. |
---|
574 | ∀eqrel : acctype → acctype → Prop. |
---|
575 | ∀refl_eqrel : reflexive ? eqrel. |
---|
576 | ∀trans_eqrel : transitive ? eqrel. |
---|
577 | ∀sym_eqrel : symmetric ? eqrel. |
---|
578 | ∀a,b,c. eqrel a b → eqrel a c → eqrel b c. |
---|
579 | #acctype #eqrel #R #T #S #a #b #c |
---|
580 | #H1 #H2 @(T … (S … H1) H2) |
---|
581 | qed. |
---|
582 | |
---|
583 | lemma cotriangle_diagram : |
---|
584 | ∀acctype : Type[0]. |
---|
585 | ∀eqrel : acctype → acctype → Prop. |
---|
586 | ∀refl_eqrel : reflexive ? eqrel. |
---|
587 | ∀trans_eqrel : transitive ? eqrel. |
---|
588 | ∀sym_eqrel : symmetric ? eqrel. |
---|
589 | ∀a,b,c. eqrel b a → eqrel c a → eqrel b c. |
---|
590 | #acctype #eqrel #R #T #S #a #b #c |
---|
591 | #H1 #H2 @S @(T … H2 (S … H1)) |
---|
592 | qed. |
---|
593 | |
---|
594 | (* --------------------------------------------------------------------------- *) |
---|
595 | (* Quick and dirty implementation of finite sets relying on lists. The |
---|
596 | * implementation is split in two: an abstract equivalence defined using inclusion |
---|
597 | * of lists, and a concrete one where equivalence is defined as the closure of |
---|
598 | * duplication, contraction and transposition of elements. We rely on the latter |
---|
599 | * to prove stuff on folds over sets. *) |
---|
600 | (* --------------------------------------------------------------------------- *) |
---|
601 | |
---|
602 | definition lset ≝ λA:Type[0]. list A. |
---|
603 | |
---|
604 | (* The empty set. *) |
---|
605 | definition empty_lset ≝ λA:Type[0]. nil A. |
---|
606 | |
---|
607 | (* Standard operations. *) |
---|
608 | definition lset_union ≝ λA:Type[0].λl1,l2 : lset A. l1 @ l2. |
---|
609 | |
---|
610 | definition lset_remove ≝ λA:DeqSet.λl:lset (carr A).λelt:carr A. (filter … (λx.¬x==elt) l). |
---|
611 | |
---|
612 | definition lset_difference ≝ λA:DeqSet.λl1,l2:lset (carr A). (filter … (λx.¬ (memb ? x l2)) l1). |
---|
613 | |
---|
614 | (* Standard predicates on sets *) |
---|
615 | definition lset_in ≝ λA:Type[0].λx : A. λl : lset A. mem … x l. |
---|
616 | |
---|
617 | definition lset_disjoint ≝ λA:Type[0].λl1, l2 : lset A. |
---|
618 | ∀x,y. mem … x l1 → mem … y l2 → x ≠ y. |
---|
619 | |
---|
620 | definition lset_inclusion ≝ λA:Type[0].λl1,l2. |
---|
621 | All A (λx1. mem … x1 l2) l1. |
---|
622 | |
---|
623 | (* Definition of abstract set equivalence. *) |
---|
624 | definition lset_eq ≝ λA:Type[0].λl1,l2. |
---|
625 | lset_inclusion A l1 l2 ∧ |
---|
626 | lset_inclusion A l2 l1. |
---|
627 | |
---|
628 | (* Properties of inclusion. *) |
---|
629 | lemma reflexive_lset_inclusion : ∀A. ∀l. lset_inclusion A l l. |
---|
630 | #A #l elim l try // |
---|
631 | #hd #tl #Hind whd @conj |
---|
632 | [ 1: %1 @refl |
---|
633 | | 2: whd in Hind; @(All_mp … Hind) |
---|
634 | #a #H whd %2 @H |
---|
635 | ] qed. |
---|
636 | |
---|
637 | lemma transitive_lset_inclusion : ∀A. ∀l1,l2,l3. lset_inclusion A l1 l2 → lset_inclusion A l2 l3 → lset_inclusion A l1 l3 . |
---|
638 | #A #l1 #l2 #l3 |
---|
639 | #Hincl1 #Hincl2 @(All_mp … Hincl1) |
---|
640 | whd in Hincl2; |
---|
641 | #a elim l2 in Hincl2; |
---|
642 | [ 1: normalize #_ @False_ind |
---|
643 | | 2: #hd #tl #Hind whd in ⊢ (% → ?); |
---|
644 | * #Hmem #Hincl_tl whd in ⊢ (% → ?); |
---|
645 | * [ 1: #Heq destruct @Hmem |
---|
646 | | 2: #Hmem_tl @Hind assumption ] |
---|
647 | ] qed. |
---|
648 | |
---|
649 | lemma cons_preserves_inclusion : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀x. lset_inclusion A l1 (x::l2). |
---|
650 | #A #l1 #l2 #Hincl #x @(All_mp … Hincl) #a #Hmem whd %2 @Hmem qed. |
---|
651 | |
---|
652 | lemma cons_monotonic_inclusion : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀x. lset_inclusion A (x::l1) (x::l2). |
---|
653 | #A #l1 #l2 #Hincl #x whd @conj |
---|
654 | [ 1: /2 by or_introl/ |
---|
655 | | 2: @(All_mp … Hincl) #a #Hmem whd %2 @Hmem ] qed. |
---|
656 | |
---|
657 | lemma lset_inclusion_concat : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀l3. lset_inclusion A l1 (l3 @ l2). |
---|
658 | #A #l1 #l2 #Hincl #l3 elim l3 |
---|
659 | try /2 by cons_preserves_inclusion/ |
---|
660 | qed. |
---|
661 | |
---|
662 | lemma lset_inclusion_concat_monotonic : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀l3. lset_inclusion A (l3 @ l1) (l3 @ l2). |
---|
663 | #A #l1 #l2 #Hincl #l3 elim l3 |
---|
664 | try @Hincl #hd #tl #Hind @cons_monotonic_inclusion @Hind |
---|
665 | qed. |
---|
666 | |
---|
667 | (* lset_eq is an equivalence relation. *) |
---|
668 | lemma reflexive_lset_eq : ∀A. ∀l. lset_eq A l l. /2 by conj/ qed. |
---|
669 | |
---|
670 | lemma transitive_lset_eq : ∀A. ∀l1,l2,l3. lset_eq A l1 l2 → lset_eq A l2 l3 → lset_eq A l1 l3. |
---|
671 | #A #l1 #l2 #l3 * #Hincl1 #Hincl2 * #Hincl3 #Hincl4 |
---|
672 | @conj @(transitive_lset_inclusion ??l2) assumption |
---|
673 | qed. |
---|
674 | |
---|
675 | lemma symmetric_lset_eq : ∀A. ∀l1,l2. lset_eq A l1 l2 → lset_eq A l2 l1. |
---|
676 | #A #l1 #l2 * #Hincl1 #Hincl2 @conj assumption |
---|
677 | qed. |
---|
678 | |
---|
679 | (* Properties of inclusion vs union and equality. *) |
---|
680 | lemma lset_union_inclusion : ∀A:Type[0]. ∀a,b,c. |
---|
681 | lset_inclusion A a c → lset_inclusion ? b c → lset_inclusion ? (lset_union ? a b) c. |
---|
682 | #A #a #b #c #H1 #H2 whd whd in match (lset_union ???); |
---|
683 | @All_append assumption qed. |
---|
684 | |
---|
685 | lemma lset_inclusion_union : ∀A:Type[0]. ∀a,b,c. |
---|
686 | lset_inclusion A a b ∨ lset_inclusion A a c → lset_inclusion ? a (lset_union ? b c). |
---|
687 | #A #a #b #c * |
---|
688 | [ 1: @All_mp #elt #Hmem @mem_append_backwards %1 @Hmem |
---|
689 | | 2: @All_mp #elt #Hmem @mem_append_backwards %2 @Hmem |
---|
690 | ] qed. |
---|
691 | |
---|
692 | lemma lset_inclusion_eq : ∀A : Type[0]. ∀a,b,c : lset A. |
---|
693 | lset_eq ? a b → lset_inclusion ? b c → lset_inclusion ? a c. |
---|
694 | #A #a #b #c * #H1 #H2 #H3 @(transitive_lset_inclusion … H1 H3) |
---|
695 | qed. |
---|
696 | |
---|
697 | lemma lset_inclusion_eq2 : ∀A : Type[0]. ∀a,b,c : lset A. |
---|
698 | lset_eq ? b c → lset_inclusion ? a b → lset_inclusion ? a c. |
---|
699 | #A #a #b #c * #H1 #H2 #H3 @(transitive_lset_inclusion … H3 H1) |
---|
700 | qed. |
---|
701 | |
---|
702 | (* Properties of lset_eq and mem *) |
---|
703 | lemma lset_eq_mem : |
---|
704 | ∀A:Type[0]. |
---|
705 | ∀s1,s2 : lset A. |
---|
706 | lset_eq ? s1 s2 → |
---|
707 | ∀b.mem ? b s1 → mem ? b s2. |
---|
708 | #A #s1 #s2 * #Hincl12 #_ #b |
---|
709 | whd in Hincl12; elim s1 in Hincl12; |
---|
710 | [ 1: normalize #_ * |
---|
711 | | 2: #hd #tl #Hind whd in ⊢ (% → % → ?); * #Hmem' #Hall * #Heq |
---|
712 | [ 1: destruct (Heq) assumption |
---|
713 | | 2: @Hind assumption ] |
---|
714 | ] qed. |
---|
715 | |
---|
716 | lemma lset_eq_memb : |
---|
717 | ∀A : DeqSet. |
---|
718 | ∀s1,s2 : lset (carr A). |
---|
719 | lset_eq ? s1 s2 → |
---|
720 | ∀b.memb ? b s1 = true → memb ? b s2 = true. |
---|
721 | #A #s1 #s2 #Heq #b |
---|
722 | lapply (memb_to_mem A s1 b) #H1 #H2 |
---|
723 | lapply (H1 H2) #Hmem lapply (lset_eq_mem … Heq ? Hmem) @mem_to_memb |
---|
724 | qed. |
---|
725 | |
---|
726 | lemma lset_eq_memb_eq : |
---|
727 | ∀A : DeqSet. |
---|
728 | ∀s1,s2 : lset (carr A). |
---|
729 | lset_eq ? s1 s2 → |
---|
730 | ∀b.memb ? b s1 = memb ? b s2. |
---|
731 | #A #s1 #s2 #Hlset_eq #b |
---|
732 | lapply (lset_eq_memb … Hlset_eq b) |
---|
733 | lapply (lset_eq_memb … (symmetric_lset_eq … Hlset_eq) b) |
---|
734 | cases (b∈s1) |
---|
735 | [ 1: #_ #H lapply (H (refl ??)) #Hmem >H @refl |
---|
736 | | 2: cases (b∈s2) // #H lapply (H (refl ??)) #Habsurd destruct |
---|
737 | ] qed. |
---|
738 | |
---|
739 | lemma lset_eq_filter_eq : |
---|
740 | ∀A : DeqSet. |
---|
741 | ∀s1,s2 : lset (carr A). |
---|
742 | lset_eq ? s1 s2 → |
---|
743 | ∀l. |
---|
744 | (filter ? (λwb.¬wb∈s1) l) = |
---|
745 | (filter ? (λwb.¬wb∈s2) l). |
---|
746 | #A #s1 #s2 #Heq #l elim l |
---|
747 | [ 1: @refl |
---|
748 | | 2: #hd #tl #Hind >filter_cons_unfold >filter_cons_unfold |
---|
749 | >(lset_eq_memb_eq … Heq) cases (hd∈s2) |
---|
750 | normalize in match (notb ?); normalize nodelta |
---|
751 | try @Hind >Hind @refl |
---|
752 | ] qed. |
---|
753 | |
---|
754 | lemma cons_monotonic_eq : ∀A. ∀l1,l2 : lset A. lset_eq A l1 l2 → ∀x. lset_eq A (x::l1) (x::l2). |
---|
755 | #A #l1 #l2 #Heq #x cases Heq #Hincl1 #Hincl2 |
---|
756 | @conj |
---|
757 | [ 1: @cons_monotonic_inclusion |
---|
758 | | 2: @cons_monotonic_inclusion ] |
---|
759 | assumption |
---|
760 | qed. |
---|
761 | |
---|
762 | (* Properties of difference and remove *) |
---|
763 | lemma lset_difference_unfold : |
---|
764 | ∀A : DeqSet. |
---|
765 | ∀s1, s2 : lset (carr A). |
---|
766 | ∀hd. lset_difference A (hd :: s1) s2 = |
---|
767 | if hd∈s2 then |
---|
768 | lset_difference A s1 s2 |
---|
769 | else |
---|
770 | hd :: (lset_difference A s1 s2). |
---|
771 | #A #s1 #s2 #hd normalize |
---|
772 | cases (hd∈s2) @refl |
---|
773 | qed. |
---|
774 | |
---|
775 | lemma lset_difference_unfold2 : |
---|
776 | ∀A : DeqSet. |
---|
777 | ∀s1, s2 : lset (carr A). |
---|
778 | ∀hd. lset_difference A s1 (hd :: s2) = |
---|
779 | lset_difference A (lset_remove ? s1 hd) s2. |
---|
780 | #A #s1 |
---|
781 | elim s1 |
---|
782 | [ 1: // |
---|
783 | | 2: #hd1 #tl1 #Hind #s2 #hd |
---|
784 | whd in match (lset_remove ???); |
---|
785 | whd in match (lset_difference A ??); |
---|
786 | whd in match (memb ???); |
---|
787 | lapply (eqb_true … hd1 hd) |
---|
788 | cases (hd1==hd) normalize nodelta |
---|
789 | [ 1: * #H #_ lapply (H (refl ??)) -H #H |
---|
790 | @Hind |
---|
791 | | 2: * #_ #Hguard >lset_difference_unfold |
---|
792 | cases (hd1∈s2) normalize in match (notb ?); normalize nodelta |
---|
793 | <Hind @refl ] |
---|
794 | ] qed. |
---|
795 | |
---|
796 | lemma lset_difference_disjoint : |
---|
797 | ∀A : DeqSet. |
---|
798 | ∀s1,s2 : lset (carr A). |
---|
799 | lset_disjoint A s1 (lset_difference A s2 s1). |
---|
800 | #A #s1 elim s1 |
---|
801 | [ 1: #s2 normalize #x #y * |
---|
802 | | 2: #hd1 #tl1 #Hind #s2 >lset_difference_unfold2 #x #y |
---|
803 | whd in ⊢ (% → ?); * |
---|
804 | [ 2: @Hind |
---|
805 | | 1: #Heq >Heq elim s2 |
---|
806 | [ 1: normalize * |
---|
807 | | 2: #hd2 #tl2 #Hind2 whd in match (lset_remove ???); |
---|
808 | lapply (eqb_true … hd2 hd1) |
---|
809 | cases (hd2 == hd1) normalize in match (notb ?); normalize nodelta * #H1 #H2 |
---|
810 | [ 1: @Hind2 |
---|
811 | | 2: >lset_difference_unfold cases (hd2 ∈ tl1) normalize nodelta try @Hind2 |
---|
812 | whd in ⊢ (% → ?); * |
---|
813 | [ 1: #Hyhd2 destruct % #Heq lapply (H2 … (sym_eq … Heq)) #Habsurd destruct |
---|
814 | | 2: @Hind2 ] |
---|
815 | ] |
---|
816 | ] |
---|
817 | ] |
---|
818 | ] qed. |
---|
819 | |
---|
820 | |
---|
821 | lemma lset_remove_split : ∀A : DeqSet. ∀l1,l2 : lset A. ∀elt. lset_remove A (l1 @ l2) elt = (lset_remove … l1 elt) @ (lset_remove … l2 elt). |
---|
822 | #A #l1 #l2 #elt /2 by filter_append/ qed. |
---|
823 | |
---|
824 | lemma lset_inclusion_remove : |
---|
825 | ∀A : DeqSet. |
---|
826 | ∀s1, s2 : lset A. |
---|
827 | lset_inclusion ? s1 s2 → |
---|
828 | ∀elt. lset_inclusion ? (lset_remove ? s1 elt) (lset_remove ? s2 elt). |
---|
829 | #A #s1 elim s1 |
---|
830 | [ 1: normalize // |
---|
831 | | 2: #hd1 #tl1 #Hind1 #s2 * #Hmem #Hincl |
---|
832 | elim (list_mem_split ??? Hmem) #s2A * #s2B #Hs2_eq destruct #elt |
---|
833 | whd in match (lset_remove ???); |
---|
834 | @(match (hd1 == elt) |
---|
835 | return λx. (hd1 == elt = x) → ? |
---|
836 | with |
---|
837 | [ true ⇒ λH. ? |
---|
838 | | false ⇒ λH. ? ] (refl ? (hd1 == elt))) normalize in match (notb ?); |
---|
839 | normalize nodelta |
---|
840 | [ 1: @Hind1 @Hincl |
---|
841 | | 2: whd @conj |
---|
842 | [ 2: @(Hind1 … Hincl) |
---|
843 | | 1: >lset_remove_split >lset_remove_split |
---|
844 | normalize in match (lset_remove A [hd1] elt); |
---|
845 | >H normalize nodelta @mem_append_backwards %2 |
---|
846 | @mem_append_backwards %1 normalize %1 @refl ] |
---|
847 | ] |
---|
848 | ] qed. |
---|
849 | |
---|
850 | lemma lset_difference_lset_eq : |
---|
851 | ∀A : DeqSet. ∀a,b,c. |
---|
852 | lset_eq A b c → |
---|
853 | lset_eq A (lset_difference A a b) (lset_difference A a c). |
---|
854 | #A #a #b #c #Heq |
---|
855 | whd in match (lset_difference ???) in ⊢ (??%%); |
---|
856 | elim a |
---|
857 | [ 1: normalize @conj @I |
---|
858 | | 2: #hda #tla #Hind whd in match (foldr ?????) in ⊢ (??%%); |
---|
859 | >(lset_eq_memb_eq … Heq hda) cases (hda∈c) |
---|
860 | normalize in match (notb ?); normalize nodelta |
---|
861 | try @Hind @cons_monotonic_eq @Hind |
---|
862 | ] qed. |
---|
863 | |
---|
864 | lemma lset_difference_lset_remove_commute : |
---|
865 | ∀A:DeqSet. |
---|
866 | ∀elt,s1,s2. |
---|
867 | (lset_difference A (lset_remove ? s1 elt) s2) = |
---|
868 | (lset_remove A (lset_difference ? s1 s2) elt). |
---|
869 | #A #elt #s1 #s2 |
---|
870 | elim s1 try // |
---|
871 | #hd #tl #Hind |
---|
872 | >lset_difference_unfold |
---|
873 | whd in match (lset_remove ???); |
---|
874 | @(match (hd==elt) return λx. (hd==elt) = x → ? |
---|
875 | with |
---|
876 | [ true ⇒ λHhd. ? |
---|
877 | | false ⇒ λHhd. ? |
---|
878 | ] (refl ? (hd==elt))) |
---|
879 | @(match (hd∈s2) return λx. (hd∈s2) = x → ? |
---|
880 | with |
---|
881 | [ true ⇒ λHmem. ? |
---|
882 | | false ⇒ λHmem. ? |
---|
883 | ] (refl ? (hd∈s2))) |
---|
884 | >notb_true >notb_false normalize nodelta try // |
---|
885 | try @Hind |
---|
886 | [ 1: whd in match (lset_remove ???) in ⊢ (???%); >Hhd |
---|
887 | elim (eqb_true ? elt elt) #_ #H >(H (refl ??)) |
---|
888 | normalize in match (notb ?); normalize nodelta @Hind |
---|
889 | | 2: >lset_difference_unfold >Hmem @Hind |
---|
890 | | 3: whd in match (lset_remove ???) in ⊢ (???%); |
---|
891 | >lset_difference_unfold >Hhd >Hmem |
---|
892 | normalize in match (notb ?); |
---|
893 | normalize nodelta >Hind @refl |
---|
894 | ] qed. |
---|
895 | |
---|
896 | (* Inversion lemma on emptyness *) |
---|
897 | lemma lset_eq_empty_inv : ∀A. ∀l. lset_eq A l [ ] → l = [ ]. |
---|
898 | #A #l elim l // |
---|
899 | #hd' #tl' normalize #Hind * * @False_ind |
---|
900 | qed. |
---|
901 | |
---|
902 | (* Inversion lemma on singletons *) |
---|
903 | lemma lset_eq_singleton_inv : ∀A,hd,l. lset_eq A [hd] (hd::l) → All ? (λx.x=hd) l. |
---|
904 | #A #hd #l |
---|
905 | * #Hincl1 whd in ⊢ (% → ?); * #_ @All_mp |
---|
906 | normalize #a * [ 1: #H @H | 2: @False_ind ] |
---|
907 | qed. |
---|
908 | |
---|
909 | (* Permutation of two elements on top of the list is ok. *) |
---|
910 | lemma lset_eq_permute : ∀A,l,x1,x2. lset_eq A (x1 :: x2 :: l) (x2 :: x1 :: l). |
---|
911 | #A #l #x1 #x2 @conj normalize |
---|
912 | [ 1: /5 by cons_preserves_inclusion, All_mp, or_introl, or_intror, conj/ |
---|
913 | | 2: /5 by cons_preserves_inclusion, All_mp, or_introl, or_intror, conj/ |
---|
914 | ] qed. |
---|
915 | |
---|
916 | (* "contraction" of an element. *) |
---|
917 | lemma lset_eq_contract : ∀A,l,x. lset_eq A (x :: x :: l) (x :: l). |
---|
918 | #A #l #x @conj |
---|
919 | [ 1: /5 by or_introl, conj, transitive_lset_inclusion/ |
---|
920 | | 2: /5 by cons_preserves_inclusion, cons_monotonic_inclusion/ ] |
---|
921 | qed. |
---|
922 | |
---|
923 | (* We don't need more than one instance of each element. *) |
---|
924 | lemma lset_eq_filter : ∀A:DeqSet.∀tl.∀hd. |
---|
925 | lset_eq A (hd :: (filter ? (λy.notb (eqb A y hd)) tl)) (hd :: tl). |
---|
926 | #A #tl elim tl |
---|
927 | [ 1: #hd normalize /4 by or_introl, conj, I/ |
---|
928 | | 2: #hd' #tl' #Hind #hd >filter_cons_unfold |
---|
929 | lapply (eqb_true A hd' hd) cases (hd'==hd) |
---|
930 | [ 2: #_ normalize in match (notb ?); normalize nodelta |
---|
931 | lapply (cons_monotonic_eq … (Hind hd) hd') #H |
---|
932 | lapply (lset_eq_permute ? tl' hd' hd) #H' |
---|
933 | @(transitive_lset_eq ? ? (hd'::hd::tl') ? … H') |
---|
934 | @(transitive_lset_eq ? ?? (hd'::hd::tl') … H) |
---|
935 | @lset_eq_permute |
---|
936 | | 1: * #Heq #_ >(Heq (refl ??)) normalize in match (notb ?); normalize nodelta |
---|
937 | lapply (Hind hd) #H |
---|
938 | @(transitive_lset_eq ? ? (hd::tl') (hd::hd::tl') H) |
---|
939 | @conj |
---|
940 | [ 1: whd @conj /2 by or_introl/ @cons_preserves_inclusion @cons_preserves_inclusion |
---|
941 | @reflexive_lset_inclusion |
---|
942 | | 2: whd @conj /2 by or_introl/ ] |
---|
943 | ] |
---|
944 | ] qed. |
---|
945 | |
---|
946 | lemma lset_inclusion_filter_self : |
---|
947 | ∀A:DeqSet.∀l,pred. |
---|
948 | lset_inclusion A (filter ? pred l) l. |
---|
949 | #A #l #pred elim l |
---|
950 | [ 1: normalize @I |
---|
951 | | 2: #hd #tl #Hind whd in match (filter ???); |
---|
952 | cases (pred hd) normalize nodelta |
---|
953 | [ 1: @cons_monotonic_inclusion @Hind |
---|
954 | | 2: @cons_preserves_inclusion @Hind ] |
---|
955 | ] qed. |
---|
956 | |
---|
957 | lemma lset_inclusion_filter_monotonic : |
---|
958 | ∀A:DeqSet.∀l1,l2. lset_inclusion (carr A) l1 l2 → |
---|
959 | ∀elt. lset_inclusion A (filter ? (λx.¬(x==elt)) l1) (filter ? (λx.¬(x==elt)) l2). |
---|
960 | #A #l1 elim l1 |
---|
961 | [ 1: #l2 normalize // |
---|
962 | | 2: #hd1 #tl1 #Hind #l2 whd in ⊢ (% → ?); * #Hmem1 #Htl1_incl #elt |
---|
963 | whd >filter_cons_unfold |
---|
964 | lapply (eqb_true A hd1 elt) cases (hd1==elt) |
---|
965 | [ 1: * #H1 #_ lapply (H1 (refl ??)) #H1_eq >H1_eq in Hmem1; #Hmem |
---|
966 | normalize in match (notb ?); normalize nodelta @Hind assumption |
---|
967 | | 2: * #_ #Hneq normalize in match (notb ?); normalize nodelta |
---|
968 | whd @conj |
---|
969 | [ 1: elim l2 in Hmem1; try // |
---|
970 | #hd2 #tl2 #Hincl whd in ⊢ (% → ?); * |
---|
971 | [ 1: #Heq >Heq in Hneq; normalize |
---|
972 | lapply (eqb_true A hd2 elt) cases (hd2==elt) |
---|
973 | [ 1: * #H #_ #H2 lapply (H2 (H (refl ??))) #Habsurd destruct (Habsurd) |
---|
974 | | 2: #_ normalize nodelta #_ /2 by or_introl/ ] |
---|
975 | | 2: #H lapply (Hincl H) #Hok |
---|
976 | normalize cases (hd2==elt) normalize nodelta |
---|
977 | [ 1: @Hok |
---|
978 | | 2: %2 @Hok ] ] |
---|
979 | | 2: @Hind assumption ] ] ] |
---|
980 | qed. |
---|
981 | |
---|
982 | (* removing an element of two equivalent sets conserves equivalence. *) |
---|
983 | lemma lset_eq_filter_monotonic : |
---|
984 | ∀A:DeqSet.∀l1,l2. lset_eq (carr A) l1 l2 → |
---|
985 | ∀elt. lset_eq A (filter ? (λx.¬(x==elt)) l1) (filter ? (λx.¬(x==elt)) l2). |
---|
986 | #A #l1 #l2 * #Hincl1 #Hincl2 #elt @conj |
---|
987 | /2 by lset_inclusion_filter_monotonic/ |
---|
988 | qed. |
---|
989 | |
---|
990 | (* ---------------- Concrete implementation of sets --------------------- *) |
---|
991 | |
---|
992 | (* We can give an explicit characterization of equivalent sets: they are permutations with repetitions, i,e. |
---|
993 | a composition of transpositions and duplications. We restrict ourselves to DeqSets. *) |
---|
994 | inductive iso_step_lset (A : DeqSet) : lset A → lset A → Prop ≝ |
---|
995 | | lset_transpose : ∀a,x,b,y,c. iso_step_lset A (a @ [x] @ b @ [y] @ c) (a @ [y] @ b @ [x] @ c) |
---|
996 | | lset_weaken : ∀a,x,b. iso_step_lset A (a @ [x] @ b) (a @ [x] @ [x] @ b) |
---|
997 | | lset_contract : ∀a,x,b. iso_step_lset A (a @ [x] @ [x] @ b) (a @ [x] @ b). |
---|
998 | |
---|
999 | (* The equivalence is the reflexive, transitive and symmetric closure of iso_step_lset *) |
---|
1000 | inductive lset_eq_concrete (A : DeqSet) : lset A → lset A → Prop ≝ |
---|
1001 | | lset_trans : ∀a,b,c. iso_step_lset A a b → lset_eq_concrete A b c → lset_eq_concrete A a c |
---|
1002 | | lset_refl : ∀a. lset_eq_concrete A a a. |
---|
1003 | |
---|
1004 | (* lset_eq_concrete is symmetric and transitive *) |
---|
1005 | lemma transitive_lset_eq_concrete : ∀A,l1,l2,l3. lset_eq_concrete A l1 l2 → lset_eq_concrete A l2 l3 → lset_eq_concrete A l1 l3. |
---|
1006 | #A #l1 #l2 #l3 #Hequiv |
---|
1007 | elim Hequiv // |
---|
1008 | #a #b #c #Hstep #Hequiv #Hind #Hcl3 lapply (Hind Hcl3) #Hbl3 |
---|
1009 | @(lset_trans ???? Hstep Hbl3) |
---|
1010 | qed. |
---|
1011 | |
---|
1012 | lemma symmetric_step : ∀A,l1,l2. iso_step_lset A l1 l2 → iso_step_lset A l2 l1. |
---|
1013 | #A #l1 #l2 * /2/ qed. |
---|
1014 | |
---|
1015 | lemma symmetric_lset_eq_concrete : ∀A,l1,l2. lset_eq_concrete A l1 l2 → lset_eq_concrete A l2 l1. |
---|
1016 | #A #l1 #l2 #H elim H // |
---|
1017 | #a #b #c #Hab #Hbc #Hcb |
---|
1018 | @(transitive_lset_eq_concrete ???? Hcb ?) |
---|
1019 | @(lset_trans … (symmetric_step ??? Hab) (lset_refl …)) |
---|
1020 | qed. |
---|
1021 | |
---|
1022 | (* lset_eq_concrete is conserved by cons. *) |
---|
1023 | lemma equivalent_step_cons : ∀A,l1,l2. iso_step_lset A l1 l2 → ∀x. iso_step_lset A (x :: l1) (x :: l2). |
---|
1024 | #A #l1 #l2 * // qed. (* That // was impressive. *) |
---|
1025 | |
---|
1026 | lemma lset_eq_concrete_cons : ∀A,l1,l2. lset_eq_concrete A l1 l2 → ∀x. lset_eq_concrete A (x :: l1) (x :: l2). |
---|
1027 | #A #l1 #l2 #Hequiv elim Hequiv try // |
---|
1028 | #a #b #c #Hab #Hbc #Hind #x %1{(equivalent_step_cons ??? Hab x) (Hind x)} |
---|
1029 | qed. |
---|
1030 | |
---|
1031 | lemma absurd_list_eq_append : ∀A.∀x.∀l1,l2:list A. [ ] = l1 @ [x] @ l2 → False. |
---|
1032 | #A #x #l1 #l2 elim l1 normalize |
---|
1033 | [ 1: #Habsurd destruct |
---|
1034 | | 2: #hd #tl #_ #Habsurd destruct |
---|
1035 | ] qed. |
---|
1036 | |
---|
1037 | (* Inversion lemma for emptyness, step case *) |
---|
1038 | lemma equivalent_iso_step_empty_inv : ∀A,l. iso_step_lset A l [] → l = [ ]. |
---|
1039 | #A #l elim l // |
---|
1040 | #hd #tl #Hind #H inversion H |
---|
1041 | [ 1: #a #x #b #y #c #_ #Habsurd |
---|
1042 | @(False_ind … (absurd_list_eq_append ? y … Habsurd)) |
---|
1043 | | 2: #a #x #b #_ #Habsurd |
---|
1044 | @(False_ind … (absurd_list_eq_append ? x … Habsurd)) |
---|
1045 | | 3: #a #x #b #_ #Habsurd |
---|
1046 | @(False_ind … (absurd_list_eq_append ? x … Habsurd)) |
---|
1047 | ] qed. |
---|
1048 | |
---|
1049 | (* Same thing for non-emptyness *) |
---|
1050 | lemma equivalent_iso_step_cons_inv : ∀A,l1,l2. l1 ≠ [ ] → iso_step_lset A l1 l2 → l2 ≠ [ ]. |
---|
1051 | #A #l1 elim l1 |
---|
1052 | [ 1: #l2 * #H @(False_ind … (H (refl ??))) |
---|
1053 | | 2: #hd #tl #H #l2 #_ #Hstep % #Hl2 >Hl2 in Hstep; #Hstep |
---|
1054 | lapply (equivalent_iso_step_empty_inv … Hstep) #Habsurd destruct |
---|
1055 | ] qed. |
---|
1056 | |
---|
1057 | lemma lset_eq_concrete_cons_inv : ∀A,l1,l2. l1 ≠ [ ] → lset_eq_concrete A l1 l2 → l2 ≠ [ ]. |
---|
1058 | #A #l1 #l2 #Hl1 #H lapply Hl1 elim H |
---|
1059 | [ 2: #a #H @H |
---|
1060 | | 1: #a #b #c #Hab #Hbc #H #Ha lapply (equivalent_iso_step_cons_inv … Ha Hab) #Hb @H @Hb |
---|
1061 | ] qed. |
---|
1062 | |
---|
1063 | lemma lset_eq_concrete_empty_inv : ∀A,l1,l2. l1 = [ ] → lset_eq_concrete A l1 l2 → l2 = [ ]. |
---|
1064 | #A #l1 #l2 #Hl1 #H lapply Hl1 elim H // |
---|
1065 | #a #b #c #Hab #Hbc #Hbc_eq #Ha >Ha in Hab; #H_b lapply (equivalent_iso_step_empty_inv … ?? (symmetric_step … H_b)) |
---|
1066 | #Hb @Hbc_eq @Hb |
---|
1067 | qed. |
---|
1068 | |
---|
1069 | (* Square equivalence diagram *) |
---|
1070 | lemma square_lset_eq_concrete : |
---|
1071 | ∀A. ∀a,b,a',b'. lset_eq_concrete A a b → lset_eq_concrete A a a' → lset_eq_concrete A b b' → lset_eq_concrete A a' b'. |
---|
1072 | #A #a #b #a' #b' #H1 #H2 #H3 |
---|
1073 | @(transitive_lset_eq_concrete ???? (symmetric_lset_eq_concrete … H2)) |
---|
1074 | @(transitive_lset_eq_concrete ???? H1) |
---|
1075 | @H3 |
---|
1076 | qed. |
---|
1077 | |
---|
1078 | (* Make the transposition of elements visible at top-level *) |
---|
1079 | lemma transpose_lset_eq_concrete : |
---|
1080 | ∀A. ∀x,y,a,b,c,a',b',c'. |
---|
1081 | lset_eq_concrete A (a @ [x] @ b @ [y] @ c) (a' @ [x] @ b' @ [y] @ c') → |
---|
1082 | lset_eq_concrete A (a @ [y] @ b @ [x] @ c) (a' @ [y] @ b' @ [x] @ c'). |
---|
1083 | #A #x #y #a #b #c #a' #b' #c |
---|
1084 | #H @(square_lset_eq_concrete ????? H) /2 by lset_trans, lset_refl, lset_transpose/ |
---|
1085 | qed. |
---|
1086 | |
---|
1087 | lemma switch_lset_eq_concrete : |
---|
1088 | ∀A. ∀a,b,c. lset_eq_concrete A (a@[b]@c) ([b]@a@c). |
---|
1089 | #A #a elim a // |
---|
1090 | #hda #tla #Hind #b #c lapply (Hind hda c) #H |
---|
1091 | lapply (lset_eq_concrete_cons … H b) |
---|
1092 | #H' normalize in H' ⊢ %; @symmetric_lset_eq_concrete |
---|
1093 | /3 by lset_transpose, lset_trans, symmetric_lset_eq_concrete/ |
---|
1094 | qed. |
---|
1095 | |
---|
1096 | (* Folding a commutative and idempotent function on equivalent sets yields the same result. *) |
---|
1097 | lemma lset_eq_concrete_fold : |
---|
1098 | ∀A : DeqSet. |
---|
1099 | ∀acctype : Type[0]. |
---|
1100 | ∀l1,l2 : list (carr A). |
---|
1101 | lset_eq_concrete A l1 l2 → |
---|
1102 | ∀f:carr A → acctype → acctype. |
---|
1103 | (∀x1,x2. ∀acc. f x1 (f x2 acc) = f x2 (f x1 acc)) → |
---|
1104 | (∀x.∀acc. f x (f x acc) = f x acc) → |
---|
1105 | ∀acc. foldr ?? f acc l1 = foldr ?? f acc l2. |
---|
1106 | #A #acctype #l1 #l2 #Heq #f #Hcomm #Hidem |
---|
1107 | elim Heq |
---|
1108 | try // |
---|
1109 | #a' #b' #c' #Hstep #Hbc #H #acc <H -H |
---|
1110 | elim Hstep |
---|
1111 | [ 1: #a #x #b #y #c |
---|
1112 | >fold_append >fold_append >fold_append >fold_append |
---|
1113 | >fold_append >fold_append >fold_append >fold_append |
---|
1114 | normalize |
---|
1115 | cut (f x (foldr A acctype f (f y (foldr A acctype f acc c)) b) = |
---|
1116 | f y (foldr A acctype f (f x (foldr A acctype f acc c)) b)) [ |
---|
1117 | elim c |
---|
1118 | [ 1: normalize elim b |
---|
1119 | [ 1: normalize >(Hcomm x y) @refl |
---|
1120 | | 2: #hdb #tlb #Hind normalize |
---|
1121 | >(Hcomm x hdb) >(Hcomm y hdb) >Hind @refl ] |
---|
1122 | | 2: #hdc #tlc #Hind normalize elim b |
---|
1123 | [ 1: normalize >(Hcomm x y) @refl |
---|
1124 | | 2: #hdb #tlb #Hind normalize |
---|
1125 | >(Hcomm x hdb) >(Hcomm y hdb) >Hind @refl ] ] |
---|
1126 | ] |
---|
1127 | #Hind >Hind @refl |
---|
1128 | | 2: #a #x #b |
---|
1129 | >fold_append >fold_append >fold_append >(fold_append ?? ([x]@[x])) |
---|
1130 | normalize >Hidem @refl |
---|
1131 | | 3: #a #x #b |
---|
1132 | >fold_append >(fold_append ?? ([x]@[x])) >fold_append >fold_append |
---|
1133 | normalize >Hidem @refl |
---|
1134 | ] qed. |
---|
1135 | |
---|
1136 | (* Folding over equivalent sets yields equivalent results, for any equivalence. *) |
---|
1137 | lemma inj_to_fold_inj : |
---|
1138 | ∀A,acctype : Type[0]. |
---|
1139 | ∀eqrel : acctype → acctype → Prop. |
---|
1140 | ∀refl_eqrel : reflexive ? eqrel. |
---|
1141 | ∀trans_eqrel : transitive ? eqrel. |
---|
1142 | ∀sym_eqrel : symmetric ? eqrel. |
---|
1143 | ∀f : A → acctype → acctype. |
---|
1144 | (∀x:A.∀acc1:acctype.∀acc2:acctype.eqrel acc1 acc2→eqrel (f x acc1) (f x acc2)) → |
---|
1145 | ∀l : list A. ∀acc1, acc2 : acctype. eqrel acc1 acc2 → eqrel (foldr … f acc1 l) (foldr … f acc2 l). |
---|
1146 | #A #acctype #eqrel #R #T #S #f #Hinj #l elim l |
---|
1147 | // |
---|
1148 | #hd #tl #Hind #acc1 #acc2 #Heq normalize @Hinj @Hind @Heq |
---|
1149 | qed. |
---|
1150 | |
---|
1151 | (* We need to extend the above proof to arbitrary equivalence relation instead of |
---|
1152 | just standard equality. *) |
---|
1153 | lemma lset_eq_concrete_fold_ext : |
---|
1154 | ∀A : DeqSet. |
---|
1155 | ∀acctype : Type[0]. |
---|
1156 | ∀eqrel : acctype → acctype → Prop. |
---|
1157 | ∀refl_eqrel : reflexive ? eqrel. |
---|
1158 | ∀trans_eqrel : transitive ? eqrel. |
---|
1159 | ∀sym_eqrel : symmetric ? eqrel. |
---|
1160 | ∀f:carr A → acctype → acctype. |
---|
1161 | (∀x,acc1,acc2. eqrel acc1 acc2 → eqrel (f x acc1) (f x acc2)) → |
---|
1162 | (∀x1,x2. ∀acc. eqrel (f x1 (f x2 acc)) (f x2 (f x1 acc))) → |
---|
1163 | (∀x.∀acc. eqrel (f x (f x acc)) (f x acc)) → |
---|
1164 | ∀l1,l2 : list (carr A). |
---|
1165 | lset_eq_concrete A l1 l2 → |
---|
1166 | ∀acc. eqrel (foldr ?? f acc l1) (foldr ?? f acc l2). |
---|
1167 | #A #acctype #eqrel #R #T #S #f #Hinj #Hcomm #Hidem #l1 #l2 #Heq |
---|
1168 | elim Heq |
---|
1169 | try // |
---|
1170 | #a' #b' #c' #Hstep #Hbc #H inversion Hstep |
---|
1171 | [ 1: #a #x #b #y #c #HlA #HlB #_ #acc |
---|
1172 | >HlB in H; #H @(T … ? (H acc)) |
---|
1173 | >fold_append >fold_append >fold_append >fold_append |
---|
1174 | >fold_append >fold_append >fold_append >fold_append |
---|
1175 | normalize |
---|
1176 | cut (eqrel (f x (foldr ? acctype f (f y (foldr ? acctype f acc c)) b)) |
---|
1177 | (f y (foldr ? acctype f (f x (foldr ? acctype f acc c)) b))) |
---|
1178 | [ 1: |
---|
1179 | elim c |
---|
1180 | [ 1: normalize elim b |
---|
1181 | [ 1: normalize @(Hcomm x y) |
---|
1182 | | 2: #hdb #tlb #Hind normalize |
---|
1183 | lapply (Hinj hdb ?? Hind) #Hind' |
---|
1184 | lapply (T … Hind' (Hcomm ???)) #Hind'' |
---|
1185 | @S @(triangle_diagram ? eqrel R T S … Hind'') @Hcomm ] |
---|
1186 | | 2: #hdc #tlc #Hind normalize elim b |
---|
1187 | [ 1: normalize @(Hcomm x y) |
---|
1188 | | 2: #hdb #tlb #Hind normalize |
---|
1189 | lapply (Hinj hdb ?? Hind) #Hind' |
---|
1190 | lapply (T … Hind' (Hcomm ???)) #Hind'' |
---|
1191 | @S @(triangle_diagram ? eqrel R T S … Hind'') @Hcomm ] |
---|
1192 | ] ] |
---|
1193 | #Hind @(inj_to_fold_inj … eqrel R T S ? Hinj … Hind) |
---|
1194 | | 2: #a #x #b #HeqA #HeqB #_ #acc >HeqB in H; #H @(T … (H acc)) |
---|
1195 | >fold_append >fold_append >fold_append >(fold_append ?? ([x]@[x])) |
---|
1196 | normalize @(inj_to_fold_inj … eqrel R T S ? Hinj) @S @Hidem |
---|
1197 | | 3: #a #x #b #HeqA #HeqB #_ #acc >HeqB in H; #H @(T … (H acc)) |
---|
1198 | >fold_append >(fold_append ?? ([x]@[x])) >fold_append >fold_append |
---|
1199 | normalize @(inj_to_fold_inj … eqrel R T S ? Hinj) @Hidem |
---|
1200 | ] qed. |
---|
1201 | |
---|
1202 | (* Prepare some well-founded induction principles on lists. The idea is to perform |
---|
1203 | an induction on the sequence of filterees of a list : taking the first element, |
---|
1204 | filtering it out of the tail, etc. We give such principles for pairs of lists |
---|
1205 | and isolated lists. *) |
---|
1206 | |
---|
1207 | (* The two lists [l1,l2] share at least the head of l1. *) |
---|
1208 | definition head_shared ≝ λA. λl1,l2 : list A. |
---|
1209 | match l1 with |
---|
1210 | [ nil ⇒ l2 = (nil ?) |
---|
1211 | | cons hd _ ⇒ mem … hd l2 |
---|
1212 | ]. |
---|
1213 | |
---|
1214 | (* Relation on pairs of lists, as described above. *) |
---|
1215 | definition filtered_lists : ∀A:DeqSet. relation (list (carr A)×(list (carr A))) ≝ |
---|
1216 | λA:DeqSet. λll1,ll2. |
---|
1217 | let 〈la1,lb1〉 ≝ ll1 in |
---|
1218 | let 〈la2,lb2〉 ≝ ll2 in |
---|
1219 | match la2 with |
---|
1220 | [ nil ⇒ False |
---|
1221 | | cons hda2 tla2 ⇒ |
---|
1222 | head_shared ? la2 lb2 ∧ |
---|
1223 | la1 = filter … (λx.¬(x==hda2)) tla2 ∧ |
---|
1224 | lb1 = filter … (λx.¬(x==hda2)) lb2 |
---|
1225 | ]. |
---|
1226 | |
---|
1227 | (* Notice the absence of plural : this relation works on a simple list, not a pair. *) |
---|
1228 | definition filtered_list : ∀A:DeqSet. relation (list (carr A)) ≝ |
---|
1229 | λA:DeqSet. λl1,l2. |
---|
1230 | match l2 with |
---|
1231 | [ nil ⇒ False |
---|
1232 | | cons hd2 tl2 ⇒ |
---|
1233 | l1 = filter … (λx.¬(x==hd2)) l2 |
---|
1234 | ]. |
---|
1235 | |
---|
1236 | (* Relation on lists based on their lengths. We know this one is well-founded. *) |
---|
1237 | definition length_lt : ∀A:DeqSet. relation (list (carr A)) ≝ |
---|
1238 | λA:DeqSet. λl1,l2. lt (length ? l1) (length ? l2). |
---|
1239 | |
---|
1240 | (* length_lt can be extended on pairs by just measuring the first component *) |
---|
1241 | definition length_fst_lt : ∀A:DeqSet. relation (list (carr A) × (list (carr A))) ≝ |
---|
1242 | λA:DeqSet. λll1,ll2. lt (length ? (\fst ll1)) (length ? (\fst ll2)). |
---|
1243 | |
---|
1244 | lemma filter_length : ∀A. ∀l. ∀f. |filter A f l| ≤ |l|. |
---|
1245 | #A #l #f elim l // |
---|
1246 | #hd #tl #Hind whd in match (filter ???); cases (f hd) normalize nodelta |
---|
1247 | [ 1: /2 by le_S_S/ |
---|
1248 | | 2: @le_S @Hind |
---|
1249 | ] qed. |
---|
1250 | |
---|
1251 | (* The order on lists defined by their length is wf *) |
---|
1252 | lemma length_lt_wf : ∀A. ∀l. WF (list (carr A)) (length_lt A) l. |
---|
1253 | #A #l % elim l |
---|
1254 | [ 1: #a normalize in ⊢ (% → ?); #H lapply (le_S_O_absurd ? H) @False_ind |
---|
1255 | | 2: #hd #tl #Hind #a #Hlt % #a' #Hlt' @Hind whd in Hlt Hlt'; whd |
---|
1256 | @(transitive_le … Hlt') @(monotonic_pred … Hlt) |
---|
1257 | qed. |
---|
1258 | |
---|
1259 | (* Order on pairs of list by measuring the first proj *) |
---|
1260 | lemma length_fst_lt_wf : ∀A. ∀ll. WF ? (length_fst_lt A) ll. |
---|
1261 | #A * #l1 #l2 % elim l1 |
---|
1262 | [ 1: * #a1 #a2 normalize in ⊢ (% → ?); #H lapply (le_S_O_absurd ? H) @False_ind |
---|
1263 | | 2: #hd1 #tl1 #Hind #a #Hlt % #a' #Hlt' @Hind whd in Hlt Hlt'; whd |
---|
1264 | @(transitive_le … Hlt') @(monotonic_pred … Hlt) |
---|
1265 | qed. |
---|
1266 | |
---|
1267 | lemma length_to_filtered_lists : ∀A. subR ? (filtered_lists A) (length_fst_lt A). |
---|
1268 | #A whd * #a1 #a2 * #b1 #b2 elim b1 |
---|
1269 | [ 1: @False_ind |
---|
1270 | | 2: #hd1 #tl1 #Hind whd in ⊢ (% → ?); * * #Hmem #Ha1 #Ha2 whd |
---|
1271 | >Ha1 whd in match (length ??) in ⊢ (??%); @le_S_S @filter_length |
---|
1272 | ] qed. |
---|
1273 | |
---|
1274 | lemma length_to_filtered_list : ∀A. subR ? (filtered_list A) (length_lt A). |
---|
1275 | #A whd #a #b elim b |
---|
1276 | [ 1: @False_ind |
---|
1277 | | 2: #hd #tl #Hind whd in ⊢ (% → ?); whd in match (filter ???); |
---|
1278 | lapply (eqb_true ? hd hd) * #_ #H >(H (refl ??)) normalize in match (notb ?); |
---|
1279 | normalize nodelta #Ha whd @le_S_S >Ha @filter_length ] |
---|
1280 | qed. |
---|
1281 | |
---|
1282 | (* Prove well-foundedness by embedding in lt *) |
---|
1283 | lemma filtered_lists_wf : ∀A. ∀ll. WF ? (filtered_lists A) ll. |
---|
1284 | #A #ll @(WF_antimonotonic … (length_to_filtered_lists A)) @length_fst_lt_wf |
---|
1285 | qed. |
---|
1286 | |
---|
1287 | lemma filtered_list_wf : ∀A. ∀l. WF ? (filtered_list A) l. |
---|
1288 | #A #l @(WF_antimonotonic … (length_to_filtered_list A)) @length_lt_wf |
---|
1289 | qed. |
---|
1290 | |
---|
1291 | definition Acc_elim : ∀A,R,x. WF A R x → (∀a. R a x → WF A R a) ≝ |
---|
1292 | λA,R,x,acc. |
---|
1293 | match acc with |
---|
1294 | [ wf _ a0 ⇒ a0 ]. |
---|
1295 | |
---|
1296 | (* Stolen from Coq. Warped to avoid prop-to-type restriction. *) |
---|
1297 | let rec WF_rect |
---|
1298 | (A : Type[0]) |
---|
1299 | (R : A → A → Prop) |
---|
1300 | (P : A → Type[0]) |
---|
1301 | (f : ∀ x : A. |
---|
1302 | (∀ y : A. R y x → WF ? R y) → |
---|
1303 | (∀ y : A. R y x → P y) → P x) |
---|
1304 | (x : A) |
---|
1305 | (a : WF A R x) on a : P x ≝ |
---|
1306 | f x (Acc_elim … a) (λy:A. λRel:R y x. WF_rect A R P f y ((Acc_elim … a) y Rel)). |
---|
1307 | |
---|
1308 | lemma lset_eq_concrete_filter : ∀A:DeqSet.∀tl.∀hd. |
---|
1309 | lset_eq_concrete A (hd :: (filter ? (λy.notb (eqb A y hd)) tl)) (hd :: tl). |
---|
1310 | #A #tl elim tl |
---|
1311 | [ 1: #hd // |
---|
1312 | | 2: #hd' #tl' #Hind #hd >filter_cons_unfold |
---|
1313 | lapply (eqb_true A hd' hd) |
---|
1314 | cases (hd'==hd) |
---|
1315 | [ 2: #_ normalize in match (notb false); normalize nodelta |
---|
1316 | >cons_to_append >(cons_to_append … hd') |
---|
1317 | >cons_to_append in ⊢ (???%); >(cons_to_append … hd') in ⊢ (???%); |
---|
1318 | @(transpose_lset_eq_concrete ? hd' hd [ ] [ ] (filter A (λy:A.¬y==hd) tl') [ ] [ ] tl') |
---|
1319 | >nil_append >nil_append >nil_append >nil_append |
---|
1320 | @lset_eq_concrete_cons >nil_append >nil_append |
---|
1321 | @Hind |
---|
1322 | | 1: * #H1 #_ lapply (H1 (refl ??)) #Heq normalize in match (notb ?); normalize nodelta |
---|
1323 | >Heq >cons_to_append >cons_to_append in ⊢ (???%); >cons_to_append in ⊢ (???(???%)); |
---|
1324 | @(square_lset_eq_concrete A ([hd]@filter A (λy:A.¬y==hd) tl') ([hd]@tl')) |
---|
1325 | [ 1: @Hind |
---|
1326 | | 2: %2 |
---|
1327 | | 3: %1{???? ? (lset_refl ??)} /2 by lset_weaken/ ] |
---|
1328 | ] |
---|
1329 | ] qed. |
---|
1330 | |
---|
1331 | |
---|
1332 | (* The "abstract", observational definition of set equivalence implies the concrete one. *) |
---|
1333 | |
---|
1334 | lemma lset_eq_to_lset_eq_concrete_aux : |
---|
1335 | ∀A,ll. |
---|
1336 | head_shared … (\fst ll) (\snd ll) → |
---|
1337 | lset_eq (carr A) (\fst ll) (\snd ll) → |
---|
1338 | lset_eq_concrete A (\fst ll) (\snd ll). |
---|
1339 | #A #ll @(WF_ind ????? (filtered_lists_wf A ll)) |
---|
1340 | * * |
---|
1341 | [ 1: #b2 #Hwf #Hind_wf whd in ⊢ (% → ?); #Hb2 >Hb2 #_ %2 |
---|
1342 | | 2: #hd1 #tl1 #b2 #Hwf #Hind_wf whd in ⊢ (% → ?); #Hmem |
---|
1343 | lapply (list_mem_split ??? Hmem) * #l2A * #l2B #Hb2 #Heq |
---|
1344 | destruct |
---|
1345 | lapply (Hind_wf 〈filter … (λx.¬x==hd1) tl1,filter … (λx.¬x==hd1) (l2A@l2B)〉) |
---|
1346 | cut (filtered_lists A 〈filter A (λx:A.¬x==hd1) tl1,filter A (λx:A.¬x==hd1) (l2A@l2B)〉 〈hd1::tl1,l2A@[hd1]@l2B〉) |
---|
1347 | [ @conj try @conj try @refl whd |
---|
1348 | [ 1: /2 by / |
---|
1349 | | 2: >filter_append in ⊢ (???%); >filter_append in ⊢ (???%); |
---|
1350 | whd in match (filter ?? [hd1]); |
---|
1351 | elim (eqb_true A hd1 hd1) #_ #H >(H (refl ??)) normalize in match (notb ?); |
---|
1352 | normalize nodelta <filter_append @refl ] ] |
---|
1353 | #Hfiltered #Hind_aux lapply (Hind_aux Hfiltered) -Hind_aux |
---|
1354 | cut (lset_eq A (filter A (λx:A.¬x==hd1) tl1) (filter A (λx:A.¬x==hd1) (l2A@l2B))) |
---|
1355 | [ 1: lapply (lset_eq_filter_monotonic … Heq hd1) |
---|
1356 | >filter_cons_unfold >filter_append >(filter_append … [hd1]) |
---|
1357 | whd in match (filter ?? [hd1]); |
---|
1358 | elim (eqb_true A hd1 hd1) #_ #H >(H (refl ??)) normalize in match (notb ?); |
---|
1359 | normalize nodelta <filter_append #Hsol @Hsol ] |
---|
1360 | #Hset_eq |
---|
1361 | cut (head_shared A (filter A (λx:A.¬x==hd1) tl1) (filter A (λx:A.¬x==hd1) (l2A@l2B))) |
---|
1362 | [ 1: lapply Hset_eq cases (filter A (λx:A.¬x==hd1) tl1) |
---|
1363 | [ 1: whd in ⊢ (% → ?); * #_ elim (filter A (λx:A.¬x==hd1) (l2A@l2B)) // |
---|
1364 | #hd' #tl' normalize #Hind * @False_ind |
---|
1365 | | 2: #hd' #tl' * #Hincl1 #Hincl2 whd elim Hincl1 #Hsol #_ @Hsol ] ] |
---|
1366 | #Hshared #Hind_aux lapply (Hind_aux Hshared Hset_eq) |
---|
1367 | #Hconcrete_set_eq |
---|
1368 | >cons_to_append |
---|
1369 | @(transitive_lset_eq_concrete ? ([hd1]@tl1) ([hd1]@l2A@l2B) (l2A@[hd1]@l2B)) |
---|
1370 | [ 2: @symmetric_lset_eq_concrete @switch_lset_eq_concrete ] |
---|
1371 | lapply (lset_eq_concrete_cons … Hconcrete_set_eq hd1) #Hconcrete_cons_eq |
---|
1372 | @(square_lset_eq_concrete … Hconcrete_cons_eq) |
---|
1373 | [ 1: @(lset_eq_concrete_filter ? tl1 hd1) |
---|
1374 | | 2: @(lset_eq_concrete_filter ? (l2A@l2B) hd1) ] |
---|
1375 | ] qed. |
---|
1376 | |
---|
1377 | lemma lset_eq_to_lset_eq_concrete : ∀A,l1,l2. lset_eq (carr A) l1 l2 → lset_eq_concrete A l1 l2. |
---|
1378 | #A * |
---|
1379 | [ 1: #l2 #Heq >(lset_eq_empty_inv … (symmetric_lset_eq … Heq)) // |
---|
1380 | | 2: #hd1 #tl1 #l2 #Hincl lapply Hincl lapply (lset_eq_to_lset_eq_concrete_aux ? 〈hd1::tl1,l2〉) #H @H |
---|
1381 | whd elim Hincl * // |
---|
1382 | ] qed. |
---|
1383 | |
---|
1384 | |
---|
1385 | (* The concrete one implies the abstract one. *) |
---|
1386 | lemma lset_eq_concrete_to_lset_eq : ∀A,l1,l2. lset_eq_concrete A l1 l2 → lset_eq A l1 l2. |
---|
1387 | #A #l1 #l2 #Hconcrete |
---|
1388 | elim Hconcrete try // |
---|
1389 | #a #b #c #Hstep #Heq_bc_concrete #Heq_bc |
---|
1390 | cut (lset_eq A a b) |
---|
1391 | [ 1: elim Hstep |
---|
1392 | [ 1: #a' elim a' |
---|
1393 | [ 2: #hda #tla #Hind #x #b' #y #c' >cons_to_append |
---|
1394 | >(associative_append ? [hda] tla ?) |
---|
1395 | >(associative_append ? [hda] tla ?) |
---|
1396 | @cons_monotonic_eq >nil_append >nil_append @Hind |
---|
1397 | | 1: #x #b' #y #c' >nil_append >nil_append |
---|
1398 | elim b' try // |
---|
1399 | #hdb #tlc #Hind >append_cons >append_cons in ⊢ (???%); |
---|
1400 | >associative_append >associative_append |
---|
1401 | lapply (cons_monotonic_eq … Hind hdb) #Hind' |
---|
1402 | @(transitive_lset_eq ? ([x]@[hdb]@tlc@[y]@c') ([hdb]@[x]@tlc@[y]@c')) |
---|
1403 | /2 by transitive_lset_eq/ ] |
---|
1404 | | 2: #a' elim a' |
---|
1405 | [ 2: #hda #tla #Hind #x #b' >cons_to_append |
---|
1406 | >(associative_append ? [hda] tla ?) |
---|
1407 | >(associative_append ? [hda] tla ?) |
---|
1408 | @cons_monotonic_eq >nil_append >nil_append @Hind |
---|
1409 | | 1: #x #b' >nil_append >nil_append @conj normalize |
---|
1410 | [ 1: @conj [ 1: %1 @refl ] elim b' |
---|
1411 | [ 1: @I |
---|
1412 | | 2: #hdb #tlb #Hind normalize @conj |
---|
1413 | [ 1: %2 %2 %1 @refl |
---|
1414 | | 2: @(All_mp … Hind) #a0 * |
---|
1415 | [ 1: #Heq %1 @Heq |
---|
1416 | | 2: * /2 by or_introl, or_intror/ ] ] ] |
---|
1417 | #H %2 %2 %2 @H |
---|
1418 | | 2: @conj try @conj try /2 by or_introl, or_intror/ elim b' |
---|
1419 | [ 1: @I |
---|
1420 | | 2: #hdb #tlb #Hind normalize @conj |
---|
1421 | [ 1: %2 %1 @refl |
---|
1422 | | 2: @(All_mp … Hind) #a0 * |
---|
1423 | [ 1: #Heq %1 @Heq |
---|
1424 | | 2: #H %2 %2 @H ] ] ] ] ] |
---|
1425 | | 3: #a #x #b elim a try @lset_eq_contract |
---|
1426 | #hda #tla #Hind @cons_monotonic_eq @Hind ] ] |
---|
1427 | #Heq_ab @(transitive_lset_eq … Heq_ab Heq_bc) |
---|
1428 | qed. |
---|
1429 | |
---|
1430 | lemma lset_eq_fold : |
---|
1431 | ∀A : DeqSet. |
---|
1432 | ∀acctype : Type[0]. |
---|
1433 | ∀eqrel : acctype → acctype → Prop. |
---|
1434 | ∀refl_eqrel : reflexive ? eqrel. |
---|
1435 | ∀trans_eqrel : transitive ? eqrel. |
---|
1436 | ∀sym_eqrel : symmetric ? eqrel. |
---|
1437 | ∀f:carr A → acctype → acctype. |
---|
1438 | (∀x,acc1,acc2. eqrel acc1 acc2 → eqrel (f x acc1) (f x acc2)) → |
---|
1439 | (∀x1,x2. ∀acc. eqrel (f x1 (f x2 acc)) (f x2 (f x1 acc))) → |
---|
1440 | (∀x.∀acc. eqrel (f x (f x acc)) (f x acc)) → |
---|
1441 | ∀l1,l2 : list (carr A). |
---|
1442 | lset_eq A l1 l2 → |
---|
1443 | ∀acc. eqrel (foldr ?? f acc l1) (foldr ?? f acc l2). |
---|
1444 | #A #acctype #eqrel #refl_eqrel #trans_eqrel #sym_eqrel #f #Hinj #Hsym #Hcontract #l1 #l2 #Heq #acc |
---|
1445 | lapply (lset_eq_to_lset_eq_concrete … Heq) #Heq_concrete |
---|
1446 | @(lset_eq_concrete_fold_ext A acctype eqrel refl_eqrel trans_eqrel sym_eqrel f Hinj Hsym Hcontract l1 l2 Heq_concrete acc) |
---|
1447 | qed. |
---|
1448 | |
---|
1449 | (* Additional lemmas on lsets *) |
---|
1450 | |
---|
1451 | lemma lset_difference_empty : |
---|
1452 | ∀A : DeqSet. |
---|
1453 | ∀s1. lset_difference A s1 [ ] = s1. |
---|
1454 | #A #s1 elim s1 try // |
---|
1455 | #hd #tl #Hind >lset_difference_unfold >Hind @refl |
---|
1456 | qed. |
---|
1457 | |
---|
1458 | lemma lset_not_mem_difference : |
---|
1459 | ∀A : DeqSet. ∀s1,s2,s3. lset_inclusion (carr A) s1 (lset_difference ? s2 s3) → ∀x. mem ? x s1 → ¬(mem ? x s3). |
---|
1460 | #A #s1 #s2 #s3 #Hincl #x #Hmem |
---|
1461 | lapply (lset_difference_disjoint ? s3 s2) whd in ⊢ (% → ?); #Hdisjoint % #Hmem_s3 |
---|
1462 | elim s1 in Hincl Hmem; |
---|
1463 | [ 1: #_ * |
---|
1464 | | 2: #hd #tl #Hind whd in ⊢ (% → %); * #Hmem_hd #Hall * |
---|
1465 | [ 2: #Hmem_x_tl @Hind assumption |
---|
1466 | | 1: #Heq lapply (Hdisjoint … Hmem_s3 Hmem_hd) * #H @H @Heq ] |
---|
1467 | ] qed. |
---|
1468 | |
---|
1469 | lemma lset_mem_inclusion_mem : |
---|
1470 | ∀A,s1,s2,elt. |
---|
1471 | mem A elt s1 → lset_inclusion ? s1 s2 → mem ? elt s2. |
---|
1472 | #A #s1 elim s1 |
---|
1473 | [ 1: #s2 #elt * |
---|
1474 | | 2: #hd #tl #Hind #s2 #elt * |
---|
1475 | [ 1: #Heq destruct * // |
---|
1476 | | 2: #Hmem_tl * #Hmem #Hall elim tl in Hall Hmem_tl; |
---|
1477 | [ 1: #_ * |
---|
1478 | | 2: #hd' #tl' #Hind * #Hmem' #Hall * |
---|
1479 | [ 1: #Heq destruct @Hmem' |
---|
1480 | | 2: #Hmem'' @Hind assumption ] ] ] ] |
---|
1481 | qed. |
---|
1482 | |
---|
1483 | lemma lset_remove_inclusion : |
---|
1484 | ∀A : DeqSet. ∀s,elt. |
---|
1485 | lset_inclusion A (lset_remove ? s elt) s. |
---|
1486 | #A #s elim s try // qed. |
---|
1487 | |
---|
1488 | lemma lset_difference_remove_inclusion : |
---|
1489 | ∀A : DeqSet. ∀s1,s2,elt. |
---|
1490 | lset_inclusion A |
---|
1491 | (lset_difference ? (lset_remove ? s1 elt) s2) |
---|
1492 | (lset_difference ? s1 s2). |
---|
1493 | #A #s elim s try // qed. |
---|
1494 | |
---|
1495 | lemma lset_difference_permute : |
---|
1496 | ∀A : DeqSet. ∀s1,s2,s3. |
---|
1497 | lset_difference A s1 (s2 @ s3) = |
---|
1498 | lset_difference A s1 (s3 @ s2). |
---|
1499 | #A #s1 #s2 elim s2 try // |
---|
1500 | #hd #tl #Hind #s3 >lset_difference_unfold2 >lset_difference_lset_remove_commute |
---|
1501 | >Hind elim s3 try // |
---|
1502 | #hd' #tl' #Hind' >cons_to_append >associative_append |
---|
1503 | >associative_append >(cons_to_append … hd tl) |
---|
1504 | >lset_difference_unfold2 >lset_difference_lset_remove_commute >nil_append |
---|
1505 | >lset_difference_unfold2 >lset_difference_lset_remove_commute >nil_append |
---|
1506 | <Hind' generalize in match (lset_difference ???); #foo |
---|
1507 | whd in match (lset_remove ???); whd in match (lset_remove ???) in ⊢ (??(?????%)?); |
---|
1508 | whd in match (lset_remove ???) in ⊢ (???%); whd in match (lset_remove ???) in ⊢ (???(?????%)); |
---|
1509 | elim foo |
---|
1510 | [ 1: normalize @refl |
---|
1511 | | 2: #hd'' #tl'' #Hind normalize |
---|
1512 | @(match (hd''==hd') return λx. ((hd''==hd') = x) → ? with |
---|
1513 | [ true ⇒ λH. ? |
---|
1514 | | false ⇒ λH. ? |
---|
1515 | ] (refl ? (hd''==hd'))) |
---|
1516 | @(match (hd''==hd) return λx. ((hd''==hd) = x) → ? with |
---|
1517 | [ true ⇒ λH'. ? |
---|
1518 | | false ⇒ λH'. ? |
---|
1519 | ] (refl ? (hd''==hd))) |
---|
1520 | normalize nodelta |
---|
1521 | try @Hind |
---|
1522 | [ 1: normalize >H normalize nodelta @Hind |
---|
1523 | | 2: normalize >H' normalize nodelta @Hind |
---|
1524 | | 3: normalize >H >H' normalize nodelta >Hind @refl |
---|
1525 | ] qed. |
---|
1526 | |
---|
1527 | |
---|
1528 | |
---|
1529 | lemma lset_disjoint_dec : |
---|
1530 | ∀A : DeqSet. |
---|
1531 | ∀s1,elt,s2. |
---|
1532 | mem ? elt s1 ∨ mem ? elt (lset_difference A (elt :: s2) s1). |
---|
1533 | #A #s1 #elt #s2 |
---|
1534 | @(match elt ∈ s1 return λx. ((elt ∈ s1) = x) → ? |
---|
1535 | with |
---|
1536 | [ false ⇒ λHA. ? |
---|
1537 | | true ⇒ λHA. ? ] (refl ? (elt ∈ s1))) |
---|
1538 | [ 1: lapply (memb_to_mem … HA) #Hmem |
---|
1539 | %1 @Hmem |
---|
1540 | | 2: %2 elim s1 in HA; |
---|
1541 | [ 1: #_ whd %1 @refl |
---|
1542 | | 2: #hd1 #tl1 #Hind normalize in ⊢ (% → ?); |
---|
1543 | >lset_difference_unfold |
---|
1544 | >lset_difference_unfold2 |
---|
1545 | lapply (eqb_true ? elt hd1) whd in match (memb ???) in ⊢ (? → ? → %); |
---|
1546 | cases (elt==hd1) normalize nodelta |
---|
1547 | [ 1: #_ #Habsurd destruct |
---|
1548 | | 2: #HA #HB >HB normalize nodelta %1 @refl ] ] ] |
---|
1549 | qed. |
---|
1550 | |
---|
1551 | lemma mem_filter : ∀A : DeqSet. ∀l,elt1,elt2. |
---|
1552 | mem A elt1 (filter A (λx:A.¬x==elt2) l) → mem A elt1 l. |
---|
1553 | #A #l elim l try // #hd #tl #Hind #elt1 #elt2 /2 by lset_mem_inclusion_mem/ |
---|
1554 | qed. |
---|
1555 | |
---|
1556 | lemma lset_inclusion_difference_aux : |
---|
1557 | ∀A : DeqSet. ∀s1,s2. |
---|
1558 | lset_inclusion A s1 s2 → |
---|
1559 | (lset_eq A s2 (s1@lset_difference A s2 s1)). |
---|
1560 | #A #s1 |
---|
1561 | @(WF_ind ????? (filtered_list_wf A s1)) |
---|
1562 | * |
---|
1563 | [ 1: #_ #_ #s2 #_ >nil_append >lset_difference_empty @reflexive_lset_eq |
---|
1564 | | 2: #hd1 #tl1 #Hwf #Hind #s2 * #Hmem #Hincl |
---|
1565 | lapply (Hind (filter ? (λx.¬x==hd1) tl1) ?) |
---|
1566 | [ 1: whd normalize |
---|
1567 | >(proj2 … (eqb_true ? hd1 hd1) (refl ??)) normalize nodelta @refl ] |
---|
1568 | #Hind_wf |
---|
1569 | elim (list_mem_split ??? Hmem) #s2A * #s2B #Heq >Heq |
---|
1570 | >cons_to_append in ⊢ (???%); >associative_append |
---|
1571 | >lset_difference_unfold2 |
---|
1572 | >nil_append |
---|
1573 | >lset_remove_split >lset_remove_split |
---|
1574 | normalize in match (lset_remove ? [hd1] hd1); |
---|
1575 | >(proj2 … (eqb_true ? hd1 hd1) (refl ??)) normalize nodelta |
---|
1576 | >nil_append <lset_remove_split >lset_difference_lset_remove_commute |
---|
1577 | lapply (Hind_wf (lset_remove A (s2A@s2B) hd1) ?) |
---|
1578 | [ 1: lapply (lset_inclusion_remove … Hincl hd1) |
---|
1579 | >Heq @lset_inclusion_eq2 |
---|
1580 | >lset_remove_split >lset_remove_split >lset_remove_split |
---|
1581 | normalize in match (lset_remove ? [hd1] hd1); |
---|
1582 | >(proj2 … (eqb_true ? hd1 hd1) (refl ??)) normalize nodelta |
---|
1583 | >nil_append @reflexive_lset_eq ] |
---|
1584 | #Hind >lset_difference_lset_remove_commute in Hind; <lset_remove_split #Hind |
---|
1585 | @lset_eq_concrete_to_lset_eq |
---|
1586 | lapply (lset_eq_to_lset_eq_concrete … (cons_monotonic_eq … Hind hd1)) #Hindc |
---|
1587 | @(square_lset_eq_concrete ????? Hindc) -Hindc -Hind |
---|
1588 | [ 1: @(transitive_lset_eq_concrete ?? ([hd1]@s2A@s2B) (s2A@[hd1]@s2B)) |
---|
1589 | [ 2: @symmetric_lset_eq_concrete @switch_lset_eq_concrete |
---|
1590 | | 1: @lset_eq_to_lset_eq_concrete @lset_eq_filter ] |
---|
1591 | | 2: @lset_eq_to_lset_eq_concrete @(transitive_lset_eq A … (lset_eq_filter ? ? hd1 …)) |
---|
1592 | elim (s2A@s2B) |
---|
1593 | [ 1: normalize in match (lset_difference ???); @reflexive_lset_eq |
---|
1594 | | 2: #hd2 #tl2 #Hind >lset_difference_unfold >lset_difference_unfold |
---|
1595 | @(match (hd2∈filter A (λx:A.¬x==hd1) tl1) |
---|
1596 | return λx. ((hd2∈filter A (λx:A.¬x==hd1) tl1) = x) → ? |
---|
1597 | with |
---|
1598 | [ false ⇒ λH. ? |
---|
1599 | | true ⇒ λH. ? |
---|
1600 | ] (refl ? (hd2∈filter A (λx:A.¬x==hd1) tl1))) normalize nodelta |
---|
1601 | [ 1: lapply (memb_to_mem … H) #Hfilter >(mem_to_memb … (mem_filter … Hfilter)) |
---|
1602 | normalize nodelta @Hind |
---|
1603 | | 2: @(match (hd2∈tl1) |
---|
1604 | return λx. ((hd2∈tl1) = x) → ? |
---|
1605 | with |
---|
1606 | [ false ⇒ λH'. ? |
---|
1607 | | true ⇒ λH'. ? |
---|
1608 | ] (refl ? (hd2∈tl1))) normalize nodelta |
---|
1609 | [ 1: (* We have hd2 = hd1 *) |
---|
1610 | cut (hd2 = hd1) |
---|
1611 | [ elim tl1 in H H'; |
---|
1612 | [ 1: normalize #_ #Habsurd destruct (Habsurd) |
---|
1613 | | 2: #hdtl1 #tltl1 #Hind normalize in ⊢ (% → % → ?); |
---|
1614 | lapply (eqb_true ? hdtl1 hd1) |
---|
1615 | cases (hdtl1==hd1) normalize nodelta |
---|
1616 | [ 1: * #H >(H (refl ??)) #_ |
---|
1617 | lapply (eqb_true ? hd2 hd1) |
---|
1618 | cases (hd2==hd1) normalize nodelta * |
---|
1619 | [ 1: #H' >(H' (refl ??)) #_ #_ #_ @refl |
---|
1620 | | 2: #_ #_ @Hind ] |
---|
1621 | | 2: #_ normalize lapply (eqb_true ? hd2 hdtl1) |
---|
1622 | cases (hd2 == hdtl1) normalize nodelta * |
---|
1623 | [ 1: #_ #_ #Habsurd destruct (Habsurd) |
---|
1624 | | 2: #_ #_ @Hind ] ] ] ] |
---|
1625 | #Heq_hd2hd1 destruct (Heq_hd2hd1) |
---|
1626 | @lset_eq_concrete_to_lset_eq lapply (lset_eq_to_lset_eq_concrete … Hind) |
---|
1627 | #Hind' @(square_lset_eq_concrete … Hind') |
---|
1628 | [ 2: @lset_refl |
---|
1629 | | 1: >cons_to_append >cons_to_append in ⊢ (???%); |
---|
1630 | @(transitive_lset_eq_concrete … ([hd1]@[hd1]@tl1@lset_difference A tl2 (filter A (λx:A.¬x==hd1) tl1))) |
---|
1631 | [ 1: @lset_eq_to_lset_eq_concrete @symmetric_lset_eq @lset_eq_contract |
---|
1632 | | 2: >(cons_to_append … hd1 (lset_difference ???)) |
---|
1633 | @lset_eq_concrete_cons >nil_append >nil_append |
---|
1634 | @symmetric_lset_eq_concrete @switch_lset_eq_concrete ] ] |
---|
1635 | | 2: @(match hd2 == hd1 |
---|
1636 | return λx. ((hd2 == hd1) = x) → ? |
---|
1637 | with |
---|
1638 | [ true ⇒ λH''. ? |
---|
1639 | | false ⇒ λH''. ? |
---|
1640 | ] (refl ? (hd2 == hd1))) |
---|
1641 | [ 1: whd in match (lset_remove ???) in ⊢ (???%); |
---|
1642 | >H'' normalize nodelta >((proj1 … (eqb_true …)) H'') |
---|
1643 | @(transitive_lset_eq … Hind) |
---|
1644 | @(transitive_lset_eq … (hd1::hd1::tl1@lset_difference A tl2 (filter A (λx:A.¬x==hd1) tl1))) |
---|
1645 | [ 2: @lset_eq_contract ] |
---|
1646 | @lset_eq_concrete_to_lset_eq @lset_eq_concrete_cons |
---|
1647 | @switch_lset_eq_concrete |
---|
1648 | | 2: whd in match (lset_remove ???) in ⊢ (???%); |
---|
1649 | >H'' >notb_false normalize nodelta |
---|
1650 | @lset_eq_concrete_to_lset_eq |
---|
1651 | lapply (lset_eq_to_lset_eq_concrete … Hind) #Hindc |
---|
1652 | lapply (lset_eq_concrete_cons … Hindc hd2) #Hindc' -Hindc |
---|
1653 | @(square_lset_eq_concrete … Hindc') |
---|
1654 | [ 1: @symmetric_lset_eq_concrete |
---|
1655 | >cons_to_append >cons_to_append in ⊢ (???%); |
---|
1656 | >(cons_to_append … hd2) >(cons_to_append … hd1) in ⊢ (???%); |
---|
1657 | @(switch_lset_eq_concrete ? ([hd1]@tl1) hd2 ?) |
---|
1658 | | 2: @symmetric_lset_eq_concrete @(switch_lset_eq_concrete ? ([hd1]@tl1) hd2 ?) |
---|
1659 | ] |
---|
1660 | ] |
---|
1661 | ] |
---|
1662 | ] |
---|
1663 | ] |
---|
1664 | ] |
---|
1665 | ] qed. |
---|
1666 | |
---|
1667 | lemma lset_inclusion_difference : |
---|
1668 | ∀A : DeqSet. |
---|
1669 | ∀s1,s2 : lset (carr A). |
---|
1670 | lset_inclusion ? s1 s2 → |
---|
1671 | ∃s2'. lset_eq ? s2 (s1 @ s2') ∧ |
---|
1672 | lset_disjoint ? s1 s2' ∧ |
---|
1673 | lset_eq ? s2' (lset_difference ? s2 s1). |
---|
1674 | #A #s1 #s2 #Hincl %{(lset_difference A s2 s1)} @conj try @conj |
---|
1675 | [ 1: @lset_inclusion_difference_aux @Hincl |
---|
1676 | | 2: /2 by lset_difference_disjoint/ |
---|
1677 | | 3,4: @reflexive_lset_inclusion ] |
---|
1678 | qed. |
---|
1679 | |
---|
1680 | (* --------------------------------------------------------------------------- *) |
---|
1681 | (* Stuff on bitvectors, previously in memoryInjections.ma *) |
---|
1682 | (* --------------------------------------------------------------------------- *) |
---|
1683 | (* --------------------------------------------------------------------------- *) |
---|
1684 | (* Some general lemmas on bitvectors (offsets /are/ bitvectors) *) |
---|
1685 | (* --------------------------------------------------------------------------- *) |
---|
1686 | |
---|
1687 | lemma add_with_carries_n_O : ∀n,bv. add_with_carries n bv (zero n) false = 〈bv,zero n〉. |
---|
1688 | #n #bv whd in match (add_with_carries ????); elim bv // |
---|
1689 | #n #hd #tl #Hind whd in match (fold_right2_i ????????); |
---|
1690 | >Hind normalize |
---|
1691 | cases n in tl; |
---|
1692 | [ 1: #tl cases hd normalize @refl |
---|
1693 | | 2: #n' #tl cases hd normalize @refl ] |
---|
1694 | qed. |
---|
1695 | |
---|
1696 | lemma addition_n_0 : ∀n,bv. addition_n n bv (zero n) = bv. |
---|
1697 | #n #bv whd in match (addition_n ???); |
---|
1698 | >add_with_carries_n_O // |
---|
1699 | qed. |
---|
1700 | |
---|
1701 | lemma replicate_Sn : ∀A,sz,elt. |
---|
1702 | replicate A (S sz) elt = elt ::: (replicate A sz elt). |
---|
1703 | // qed. |
---|
1704 | |
---|
1705 | lemma zero_Sn : ∀n. zero (S n) = false ::: (zero n). // qed. |
---|
1706 | |
---|
1707 | lemma negation_bv_Sn : ∀n. ∀xa. ∀a : BitVector n. negation_bv … (xa ::: a) = (notb xa) ::: (negation_bv … a). |
---|
1708 | #n #xa #a normalize @refl qed. |
---|
1709 | |
---|
1710 | (* useful facts on carry_of *) |
---|
1711 | lemma carry_of_TT : ∀x. carry_of true true x = true. // qed. |
---|
1712 | lemma carry_of_TF : ∀x. carry_of true false x = x. // qed. |
---|
1713 | lemma carry_of_FF : ∀x. carry_of false false x = false. // qed. |
---|
1714 | lemma carry_of_lcomm : ∀x,y,z. carry_of x y z = carry_of y x z. * * * // qed. |
---|
1715 | lemma carry_of_rcomm : ∀x,y,z. carry_of x y z = carry_of x z y. * * * // qed. |
---|
1716 | |
---|
1717 | |
---|
1718 | |
---|
1719 | definition one_bv ≝ λn. (\fst (add_with_carries … (zero n) (zero n) true)). |
---|
1720 | |
---|
1721 | lemma one_bv_Sn_aux : ∀n. ∀bits,flags : BitVector (S n). |
---|
1722 | add_with_carries … (zero (S n)) (zero (S n)) true = 〈bits, flags〉 → |
---|
1723 | add_with_carries … (zero (S (S n))) (zero (S (S n))) true = 〈false ::: bits, false ::: flags〉. |
---|
1724 | #n elim n |
---|
1725 | [ 1: #bits #flags elim (BitVector_Sn … bits) #hd_bits * #tl_bits #Heq_bits |
---|
1726 | elim (BitVector_Sn … flags) #hd_flags * #tl_flags #Heq_flags |
---|
1727 | >(BitVector_O … tl_flags) >(BitVector_O … tl_bits) |
---|
1728 | normalize #Heq destruct (Heq) @refl |
---|
1729 | | 2: #n' #Hind #bits #flags elim (BitVector_Sn … bits) #hd_bits * #tl_bits #Heq_bits |
---|
1730 | destruct #Hind >add_with_carries_Sn >replicate_Sn |
---|
1731 | whd in match (zero ?) in Hind; lapply Hind |
---|
1732 | elim (add_with_carries (S (S n')) |
---|
1733 | (false:::replicate bool (S n') false) |
---|
1734 | (false:::replicate bool (S n') false) true) #bits #flags #Heq destruct |
---|
1735 | normalize >add_with_carries_Sn in Hind; |
---|
1736 | elim (add_with_carries (S n') (replicate bool (S n') false) |
---|
1737 | (replicate bool (S n') false) true) #flags' #bits' |
---|
1738 | normalize |
---|
1739 | cases (match bits' in Vector return λsz:ℕ.(λfoo:Vector bool sz.bool) with |
---|
1740 | [VEmpty⇒true|VCons (sz:ℕ) (cy:bool) (tl:(Vector bool sz))⇒cy]) |
---|
1741 | normalize #Heq destruct @refl |
---|
1742 | ] qed. |
---|
1743 | |
---|
1744 | lemma one_bv_Sn : ∀n. one_bv (S (S n)) = false ::: (one_bv (S n)). |
---|
1745 | #n lapply (one_bv_Sn_aux n) |
---|
1746 | whd in match (one_bv ?) in ⊢ (? → (??%%)); |
---|
1747 | elim (add_with_carries (S n) (zero (S n)) (zero (S n)) true) #bits #flags |
---|
1748 | #H lapply (H bits flags (refl ??)) #H2 >H2 @refl |
---|
1749 | qed. |
---|
1750 | |
---|
1751 | lemma increment_to_addition_n_aux : ∀n. ∀a : BitVector n. |
---|
1752 | add_with_carries ? a (zero n) true = add_with_carries ? a (one_bv n) false. |
---|
1753 | #n |
---|
1754 | elim n |
---|
1755 | [ 1: #a >(BitVector_O … a) normalize @refl |
---|
1756 | | 2: #n' cases n' |
---|
1757 | [ 1: #Hind #a elim (BitVector_Sn ? a) #xa * #tl #Heq destruct |
---|
1758 | >(BitVector_O … tl) normalize cases xa @refl |
---|
1759 | | 2: #n'' #Hind #a elim (BitVector_Sn ? a) #xa * #tl #Heq destruct |
---|
1760 | >one_bv_Sn >zero_Sn |
---|
1761 | lapply (Hind tl) |
---|
1762 | >add_with_carries_Sn >add_with_carries_Sn |
---|
1763 | #Hind >Hind elim (add_with_carries (S n'') tl (one_bv (S n'')) false) #bits #flags |
---|
1764 | normalize nodelta elim (BitVector_Sn … flags) #flags_hd * #flags_tl #Hflags_eq >Hflags_eq |
---|
1765 | normalize nodelta @refl |
---|
1766 | ] qed. |
---|
1767 | |
---|
1768 | (* In order to use associativity on increment, we hide it under addition_n. *) |
---|
1769 | lemma increment_to_addition_n : ∀n. ∀a : BitVector n. increment ? a = addition_n ? a (one_bv n). |
---|
1770 | #n |
---|
1771 | whd in match (increment ??) in ⊢ (∀_.??%?); |
---|
1772 | whd in match (addition_n ???) in ⊢ (∀_.???%); |
---|
1773 | #a lapply (increment_to_addition_n_aux n a) |
---|
1774 | #Heq >Heq cases (add_with_carries n a (one_bv n) false) #bits #flags @refl |
---|
1775 | qed. |
---|
1776 | |
---|
1777 | (* Explicit formulation of addition *) |
---|
1778 | |
---|
1779 | (* Explicit formulation of the last carry bit *) |
---|
1780 | let rec ith_carry (n : nat) (a,b : BitVector n) (init : bool) on n : bool ≝ |
---|
1781 | match n return λx. BitVector x → BitVector x → bool with |
---|
1782 | [ O ⇒ λ_,_. init |
---|
1783 | | S x ⇒ λa',b'. |
---|
1784 | let hd_a ≝ head' … a' in |
---|
1785 | let hd_b ≝ head' … b' in |
---|
1786 | let tl_a ≝ tail … a' in |
---|
1787 | let tl_b ≝ tail … b' in |
---|
1788 | carry_of hd_a hd_b (ith_carry x tl_a tl_b init) |
---|
1789 | ] a b. |
---|
1790 | |
---|
1791 | lemma ith_carry_unfold : ∀n. ∀init. ∀a,b : BitVector (S n). |
---|
1792 | ith_carry ? a b init = (carry_of (head' … a) (head' … b) (ith_carry ? (tail … a) (tail … b) init)). |
---|
1793 | #n #init #a #b @refl qed. |
---|
1794 | |
---|
1795 | lemma ith_carry_Sn : ∀n. ∀init. ∀xa,xb. ∀a,b : BitVector n. |
---|
1796 | ith_carry ? (xa ::: a) (xb ::: b) init = (carry_of xa xb (ith_carry ? a b init)). // qed. |
---|
1797 | |
---|
1798 | (* correction of [ith_carry] *) |
---|
1799 | lemma ith_carry_ok : ∀n. ∀init. ∀a,b,res_ab,flags_ab : BitVector (S n). |
---|
1800 | 〈res_ab,flags_ab〉 = add_with_carries ? a b init → |
---|
1801 | head' … flags_ab = ith_carry ? a b init. |
---|
1802 | #n elim n |
---|
1803 | [ 1: #init #a #b #res_ab #flags_ab |
---|
1804 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a |
---|
1805 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b |
---|
1806 | elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res |
---|
1807 | elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags |
---|
1808 | destruct |
---|
1809 | >(BitVector_O … tl_a) >(BitVector_O … tl_b) |
---|
1810 | cases hd_a cases hd_b cases init normalize #Heq destruct (Heq) |
---|
1811 | @refl |
---|
1812 | | 2: #n' #Hind #init #a #b #res_ab #flags_ab |
---|
1813 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a |
---|
1814 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b |
---|
1815 | elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res |
---|
1816 | elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags |
---|
1817 | destruct |
---|
1818 | lapply (Hind … init tl_a tl_b tl_res tl_flags) |
---|
1819 | >add_with_carries_Sn >(ith_carry_Sn (S n')) |
---|
1820 | elim (add_with_carries (S n') tl_a tl_b init) #res_ab #flags_ab |
---|
1821 | elim (BitVector_Sn … flags_ab) #hd_flags_ab * #tl_flags_ab #Heq_flags_ab >Heq_flags_ab |
---|
1822 | normalize nodelta cases hd_flags_ab normalize nodelta |
---|
1823 | whd in match (head' ? (S n') ?); #H1 #H2 |
---|
1824 | destruct (H2) lapply (H1 (refl ??)) whd in match (head' ???); #Heq <Heq @refl |
---|
1825 | ] qed. |
---|
1826 | |
---|
1827 | (* Explicit formulation of ith bit of an addition, with explicit initial carry bit. *) |
---|
1828 | definition ith_bit ≝ λ(n : nat).λ(a,b : BitVector n).λinit. |
---|
1829 | match n return λx. BitVector x → BitVector x → bool with |
---|
1830 | [ O ⇒ λ_,_. init |
---|
1831 | | S x ⇒ λa',b'. |
---|
1832 | let hd_a ≝ head' … a' in |
---|
1833 | let hd_b ≝ head' … b' in |
---|
1834 | let tl_a ≝ tail … a' in |
---|
1835 | let tl_b ≝ tail … b' in |
---|
1836 | xorb (xorb hd_a hd_b) (ith_carry x tl_a tl_b init) |
---|
1837 | ] a b. |
---|
1838 | |
---|
1839 | lemma ith_bit_unfold : ∀n. ∀init. ∀a,b : BitVector (S n). |
---|
1840 | ith_bit ? a b init = xorb (xorb (head' … a) (head' … b)) (ith_carry ? (tail … a) (tail … b) init). |
---|
1841 | #n #a #b // qed. |
---|
1842 | |
---|
1843 | lemma ith_bit_Sn : ∀n. ∀init. ∀xa,xb. ∀a,b : BitVector n. |
---|
1844 | ith_bit ? (xa ::: a) (xb ::: b) init = xorb (xorb xa xb) (ith_carry ? a b init). // qed. |
---|
1845 | |
---|
1846 | (* correction of ith_bit *) |
---|
1847 | lemma ith_bit_ok : ∀n. ∀init. ∀a,b,res_ab,flags_ab : BitVector (S n). |
---|
1848 | 〈res_ab,flags_ab〉 = add_with_carries ? a b init → |
---|
1849 | head' … res_ab = ith_bit ? a b init. |
---|
1850 | #n |
---|
1851 | cases n |
---|
1852 | [ 1: #init #a #b #res_ab #flags_ab |
---|
1853 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a |
---|
1854 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b |
---|
1855 | elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res |
---|
1856 | elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags |
---|
1857 | destruct |
---|
1858 | >(BitVector_O … tl_a) >(BitVector_O … tl_b) |
---|
1859 | >(BitVector_O … tl_flags) >(BitVector_O … tl_res) |
---|
1860 | normalize cases init cases hd_a cases hd_b normalize #Heq destruct @refl |
---|
1861 | | 2: #n' #init #a #b #res_ab #flags_ab |
---|
1862 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a |
---|
1863 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b |
---|
1864 | elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res |
---|
1865 | elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags |
---|
1866 | destruct |
---|
1867 | lapply (ith_carry_ok … init tl_a tl_b tl_res tl_flags) |
---|
1868 | #Hcarry >add_with_carries_Sn elim (add_with_carries ? tl_a tl_b init) in Hcarry; |
---|
1869 | #res #flags normalize nodelta elim (BitVector_Sn … flags) #hd_flags' * #tl_flags' #Heq_flags' |
---|
1870 | >Heq_flags' normalize nodelta cases hd_flags' normalize nodelta #H1 #H2 destruct (H2) |
---|
1871 | cases hd_a cases hd_b >ith_bit_Sn whd in match (head' ???) in H1 ⊢ %; |
---|
1872 | <(H1 (refl ??)) @refl |
---|
1873 | ] qed. |
---|
1874 | |
---|
1875 | (* Transform a function from bit-vectors to bits into a vector by folding *) |
---|
1876 | let rec bitvector_fold (n : nat) (v : BitVector n) (f : ∀sz. BitVector sz → bool) on v : BitVector n ≝ |
---|
1877 | match v with |
---|
1878 | [ VEmpty ⇒ VEmpty ? |
---|
1879 | | VCons sz elt tl ⇒ |
---|
1880 | let bit ≝ f ? v in |
---|
1881 | bit ::: (bitvector_fold ? tl f) |
---|
1882 | ]. |
---|
1883 | |
---|
1884 | (* Two-arguments version *) |
---|
1885 | let rec bitvector_fold2 (n : nat) (v1, v2 : BitVector n) (f : ∀sz. BitVector sz → BitVector sz → bool) on v1 : BitVector n ≝ |
---|
1886 | match v1 with |
---|
1887 | [ VEmpty ⇒ λ_. VEmpty ? |
---|
1888 | | VCons sz elt tl ⇒ λv2'. |
---|
1889 | let bit ≝ f ? v1 v2 in |
---|
1890 | bit ::: (bitvector_fold2 ? tl (tail … v2') f) |
---|
1891 | ] v2. |
---|
1892 | |
---|
1893 | lemma bitvector_fold2_Sn : ∀n,x1,x2,v1,v2,f. |
---|
1894 | bitvector_fold2 (S n) (x1 ::: v1) (x2 ::: v2) f = (f ? (x1 ::: v1) (x2 ::: v2)) ::: (bitvector_fold2 … v1 v2 f). // qed. |
---|
1895 | |
---|
1896 | (* These functions pack all the relevant information (including carries) directly. *) |
---|
1897 | definition addition_n_direct ≝ λn,v1,v2,init. bitvector_fold2 n v1 v2 (λn,v1,v2. ith_bit n v1 v2 init). |
---|
1898 | |
---|
1899 | lemma addition_n_direct_Sn : ∀n,x1,x2,v1,v2,init. |
---|
1900 | addition_n_direct (S n) (x1 ::: v1) (x2 ::: v2) init = (ith_bit ? (x1 ::: v1) (x2 ::: v2) init) ::: (addition_n_direct … v1 v2 init). // qed. |
---|
1901 | |
---|
1902 | lemma tail_Sn : ∀n. ∀x. ∀a : BitVector n. tail … (x ::: a) = a. // qed. |
---|
1903 | |
---|
1904 | (* Prove the equivalence of addition_n_direct with add_with_carries *) |
---|
1905 | lemma addition_n_direct_ok : ∀n,carry,v1,v2. |
---|
1906 | (\fst (add_with_carries n v1 v2 carry)) = addition_n_direct n v1 v2 carry. |
---|
1907 | #n elim n |
---|
1908 | [ 1: #carry #v1 #v2 >(BitVector_O … v1) >(BitVector_O … v2) normalize @refl |
---|
1909 | | 2: #n' #Hind #carry #v1 #v2 |
---|
1910 | elim (BitVector_Sn … v1) #hd1 * #tl1 #Heq1 |
---|
1911 | elim (BitVector_Sn … v2) #hd2 * #tl2 #Heq2 |
---|
1912 | lapply (Hind carry tl1 tl2) |
---|
1913 | lapply (ith_bit_ok ? carry v1 v2) |
---|
1914 | lapply (ith_carry_ok ? carry v1 v2) |
---|
1915 | destruct |
---|
1916 | #Hind >addition_n_direct_Sn |
---|
1917 | >ith_bit_Sn >add_with_carries_Sn |
---|
1918 | elim (add_with_carries n' tl1 tl2 carry) #bits #flags normalize nodelta |
---|
1919 | cases (match flags in Vector return λsz:ℕ.(λfoo:Vector bool sz.bool) with |
---|
1920 | [VEmpty⇒carry|VCons (sz:ℕ) (cy:bool) (tl:(Vector bool sz))⇒cy]) |
---|
1921 | normalize nodelta #Hcarry' lapply (Hcarry' ?? (refl ??)) |
---|
1922 | whd in match head'; normalize nodelta |
---|
1923 | #H1 #H2 >H1 >H2 @refl |
---|
1924 | ] qed. |
---|
1925 | |
---|
1926 | lemma addition_n_direct_ok2 : ∀n,carry,v1,v2. |
---|
1927 | (let 〈a,b〉 ≝ add_with_carries n v1 v2 carry in a) = addition_n_direct n v1 v2 carry. |
---|
1928 | #n #carry #v1 #v2 <addition_n_direct_ok |
---|
1929 | cases (add_with_carries ????) // |
---|
1930 | qed. |
---|
1931 | |
---|
1932 | (* trivially lift associativity to our new setting *) |
---|
1933 | lemma associative_addition_n_direct : ∀n. ∀carry1,carry2. ∀v1,v2,v3 : BitVector n. |
---|
1934 | addition_n_direct ? (addition_n_direct ? v1 v2 carry1) v3 carry2 = |
---|
1935 | addition_n_direct ? v1 (addition_n_direct ? v2 v3 carry1) carry2. |
---|
1936 | #n #carry1 #carry2 #v1 #v2 #v3 |
---|
1937 | <addition_n_direct_ok <addition_n_direct_ok |
---|
1938 | <addition_n_direct_ok <addition_n_direct_ok |
---|
1939 | lapply (associative_add_with_carries … carry1 carry2 v1 v2 v3) |
---|
1940 | elim (add_with_carries n v2 v3 carry1) #bits #carries normalize nodelta |
---|
1941 | elim (add_with_carries n v1 v2 carry1) #bits' #carries' normalize nodelta |
---|
1942 | #H @(sym_eq … H) |
---|
1943 | qed. |
---|
1944 | |
---|
1945 | lemma commutative_addition_n_direct : ∀n. ∀v1,v2 : BitVector n. |
---|
1946 | addition_n_direct ? v1 v2 false = addition_n_direct ? v2 v1 false. |
---|
1947 | #n #v1 #v2 /by associative_addition_n, addition_n_direct_ok/ |
---|
1948 | qed. |
---|
1949 | |
---|
1950 | definition increment_direct ≝ λn,v. addition_n_direct n v (one_bv ?) false. |
---|
1951 | definition twocomp_neg_direct ≝ λn,v. increment_direct n (negation_bv n v). |
---|
1952 | |
---|
1953 | |
---|
1954 | (* fold andb on a bitvector. *) |
---|
1955 | let rec andb_fold (n : nat) (b : BitVector n) on b : bool ≝ |
---|
1956 | match b with |
---|
1957 | [ VEmpty ⇒ true |
---|
1958 | | VCons sz elt tl ⇒ |
---|
1959 | andb elt (andb_fold ? tl) |
---|
1960 | ]. |
---|
1961 | |
---|
1962 | lemma andb_fold_Sn : ∀n. ∀x. ∀b : BitVector n. andb_fold (S n) (x ::: b) = andb x (andb_fold … n b). // qed. |
---|
1963 | |
---|
1964 | lemma andb_fold_inversion : ∀n. ∀elt,x. andb_fold (S n) (elt ::: x) = true → elt = true ∧ andb_fold n x = true. |
---|
1965 | #n #elt #x cases elt normalize #H @conj destruct (H) try assumption @refl |
---|
1966 | qed. |
---|
1967 | |
---|
1968 | lemma ith_increment_carry : ∀n. ∀a : BitVector (S n). |
---|
1969 | ith_carry … a (one_bv ?) false = andb_fold … a. |
---|
1970 | #n elim n |
---|
1971 | [ 1: #a elim (BitVector_Sn … a) #hd * #tl #Heq >Heq >(BitVector_O … tl) |
---|
1972 | cases hd normalize @refl |
---|
1973 | | 2: #n' #Hind #a |
---|
1974 | elim (BitVector_Sn … a) #hd * #tl #Heq >Heq |
---|
1975 | lapply (Hind … tl) #Hind >one_bv_Sn |
---|
1976 | >ith_carry_Sn whd in match (andb_fold ??); |
---|
1977 | cases hd >Hind @refl |
---|
1978 | ] qed. |
---|
1979 | |
---|
1980 | lemma ith_increment_bit : ∀n. ∀a : BitVector (S n). |
---|
1981 | ith_bit … a (one_bv ?) false = xorb (head' … a) (andb_fold … (tail … a)). |
---|
1982 | #n #a |
---|
1983 | elim (BitVector_Sn … a) #hd * #tl #Heq >Heq |
---|
1984 | whd in match (head' ???); |
---|
1985 | -a cases n in tl; |
---|
1986 | [ 1: #tl >(BitVector_O … tl) cases hd normalize try // |
---|
1987 | | 2: #n' #tl >one_bv_Sn >ith_bit_Sn |
---|
1988 | >ith_increment_carry >tail_Sn |
---|
1989 | cases hd try // |
---|
1990 | ] qed. |
---|
1991 | |
---|
1992 | (* Lemma used to prove involutivity of two-complement negation *) |
---|
1993 | lemma twocomp_neg_involutive_aux : ∀n. ∀v : BitVector (S n). |
---|
1994 | (andb_fold (S n) (negation_bv (S n) v) = |
---|
1995 | andb_fold (S n) (negation_bv (S n) (addition_n_direct (S n) (negation_bv (S n) v) (one_bv (S n)) false))). |
---|
1996 | #n elim n |
---|
1997 | [ 1: #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq >(BitVector_O … tl) cases hd @refl |
---|
1998 | | 2: #n' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq |
---|
1999 | lapply (Hind tl) -Hind #Hind >negation_bv_Sn >one_bv_Sn >addition_n_direct_Sn |
---|
2000 | >andb_fold_Sn >ith_bit_Sn >negation_bv_Sn >andb_fold_Sn <Hind |
---|
2001 | cases hd normalize nodelta |
---|
2002 | [ 1: >xorb_false >(xorb_comm false ?) >xorb_false |
---|
2003 | | 2: >xorb_false >(xorb_comm true ?) >xorb_true ] |
---|
2004 | >ith_increment_carry |
---|
2005 | cases (andb_fold (S n') (negation_bv (S n') tl)) @refl |
---|
2006 | ] qed. |
---|
2007 | |
---|
2008 | (* Test of the 'direct' approach: proof of the involutivity of two-complement negation. *) |
---|
2009 | lemma twocomp_neg_involutive : ∀n. ∀v : BitVector n. twocomp_neg_direct ? (twocomp_neg_direct ? v) = v. |
---|
2010 | #n elim n |
---|
2011 | [ 1: #v >(BitVector_O v) @refl |
---|
2012 | | 2: #n' cases n' |
---|
2013 | [ 1: #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq |
---|
2014 | >(BitVector_O … tl) normalize cases hd @refl |
---|
2015 | | 2: #n'' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq |
---|
2016 | lapply (Hind tl) -Hind #Hind <Hind in ⊢ (???%); |
---|
2017 | whd in match twocomp_neg_direct; normalize nodelta |
---|
2018 | whd in match increment_direct; normalize nodelta |
---|
2019 | >(negation_bv_Sn ? hd tl) >one_bv_Sn >(addition_n_direct_Sn ? (¬hd) false ??) |
---|
2020 | >ith_bit_Sn >negation_bv_Sn >addition_n_direct_Sn >ith_bit_Sn |
---|
2021 | generalize in match (addition_n_direct (S n'') |
---|
2022 | (negation_bv (S n'') |
---|
2023 | (addition_n_direct (S n'') (negation_bv (S n'') tl) (one_bv (S n'')) false)) |
---|
2024 | (one_bv (S n'')) false); #tail |
---|
2025 | >ith_increment_carry >ith_increment_carry |
---|
2026 | cases hd normalize nodelta |
---|
2027 | [ 1: normalize in match (xorb false false); >(xorb_comm false ?) >xorb_false >xorb_false |
---|
2028 | | 2: normalize in match (xorb true false); >(xorb_comm true ?) >xorb_true >xorb_false ] |
---|
2029 | <twocomp_neg_involutive_aux |
---|
2030 | cases (andb_fold (S n'') (negation_bv (S n'') tl)) @refl |
---|
2031 | ] |
---|
2032 | ] qed. |
---|
2033 | |
---|
2034 | lemma bitvector_cons_inj_inv : ∀n. ∀a,b. ∀va,vb : BitVector n. a ::: va = b ::: vb → a =b ∧ va = vb. |
---|
2035 | #n #a #b #va #vb #H destruct (H) @conj @refl qed. |
---|
2036 | |
---|
2037 | lemma bitvector_cons_eq : ∀n. ∀a,b. ∀v : BitVector n. a = b → a ::: v = b ::: v. // qed. |
---|
2038 | |
---|
2039 | (* Injectivity of increment *) |
---|
2040 | lemma increment_inj : ∀n. ∀a,b : BitVector n. |
---|
2041 | increment_direct ? a = increment_direct ? b → |
---|
2042 | a = b ∧ (ith_carry n a (one_bv n) false = ith_carry n b (one_bv n) false). |
---|
2043 | #n whd in match increment_direct; normalize nodelta elim n |
---|
2044 | [ 1: #a #b >(BitVector_O … a) >(BitVector_O … b) normalize #_ @conj // |
---|
2045 | | 2: #n' cases n' |
---|
2046 | [ 1: #_ #a #b |
---|
2047 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a >Heq_a |
---|
2048 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b >Heq_b |
---|
2049 | >(BitVector_O … tl_a) >(BitVector_O … tl_b) cases hd_a cases hd_b |
---|
2050 | normalize #H @conj try // |
---|
2051 | | 2: #n'' #Hind #a #b |
---|
2052 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a >Heq_a |
---|
2053 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b >Heq_b |
---|
2054 | lapply (Hind … tl_a tl_b) -Hind #Hind |
---|
2055 | >one_bv_Sn >addition_n_direct_Sn >ith_bit_Sn |
---|
2056 | >addition_n_direct_Sn >ith_bit_Sn >xorb_false >xorb_false |
---|
2057 | #H elim (bitvector_cons_inj_inv … H) #Heq1 #Heq2 |
---|
2058 | lapply (Hind Heq2) * #Heq3 #Heq4 |
---|
2059 | cut (hd_a = hd_b) |
---|
2060 | [ 1: >Heq4 in Heq1; #Heq5 lapply (xorb_inj (ith_carry ? tl_b (one_bv ?) false) hd_a hd_b) |
---|
2061 | * #Heq6 #_ >xorb_comm in Heq6; >(xorb_comm ? hd_b) #Heq6 >(Heq6 Heq5) |
---|
2062 | @refl ] |
---|
2063 | #Heq5 @conj [ 1: >Heq3 >Heq5 @refl ] |
---|
2064 | >ith_carry_Sn >ith_carry_Sn >Heq4 >Heq5 @refl |
---|
2065 | ] qed. |
---|
2066 | |
---|
2067 | (* Inverse of injecivity of increment, does not lose information (cf increment_inj) *) |
---|
2068 | lemma increment_inj_inv : ∀n. ∀a,b : BitVector n. |
---|
2069 | a = b → increment_direct ? a = increment_direct ? b. // qed. |
---|
2070 | |
---|
2071 | (* A more general result. *) |
---|
2072 | lemma addition_n_direct_inj : ∀n. ∀x,y,delta: BitVector n. |
---|
2073 | addition_n_direct ? x delta false = addition_n_direct ? y delta false → |
---|
2074 | x = y ∧ (ith_carry n x delta false = ith_carry n y delta false). |
---|
2075 | #n elim n |
---|
2076 | [ 1: #x #y #delta >(BitVector_O … x) >(BitVector_O … y) >(BitVector_O … delta) * @conj @refl |
---|
2077 | | 2: #n' #Hind #x #y #delta |
---|
2078 | elim (BitVector_Sn … x) #hdx * #tlx #Heqx >Heqx |
---|
2079 | elim (BitVector_Sn … y) #hdy * #tly #Heqy >Heqy |
---|
2080 | elim (BitVector_Sn … delta) #hdd * #tld #Heqd >Heqd |
---|
2081 | >addition_n_direct_Sn >ith_bit_Sn |
---|
2082 | >addition_n_direct_Sn >ith_bit_Sn |
---|
2083 | >ith_carry_Sn >ith_carry_Sn |
---|
2084 | lapply (Hind … tlx tly tld) -Hind #Hind #Heq |
---|
2085 | elim (bitvector_cons_inj_inv … Heq) #Heq_hd #Heq_tl |
---|
2086 | lapply (Hind Heq_tl) -Hind * #HindA #HindB |
---|
2087 | >HindA >HindB >HindB in Heq_hd; #Heq_hd |
---|
2088 | cut (hdx = hdy) |
---|
2089 | [ 1: cases hdd in Heq_hd; cases (ith_carry n' tly tld false) |
---|
2090 | cases hdx cases hdy normalize #H try @H try @refl |
---|
2091 | >H try @refl ] |
---|
2092 | #Heq_hd >Heq_hd @conj @refl |
---|
2093 | ] qed. |
---|
2094 | |
---|
2095 | (* We also need it the other way around. *) |
---|
2096 | lemma addition_n_direct_inj_inv : ∀n. ∀x,y,delta: BitVector n. |
---|
2097 | x ≠ y → (* ∧ (ith_carry n x delta false = ith_carry n y delta false). *) |
---|
2098 | addition_n_direct ? x delta false ≠ addition_n_direct ? y delta false. |
---|
2099 | #n elim n |
---|
2100 | [ 1: #x #y #delta >(BitVector_O … x) >(BitVector_O … y) >(BitVector_O … delta) * #H @(False_ind … (H (refl ??))) |
---|
2101 | | 2: #n' #Hind #x #y #delta |
---|
2102 | elim (BitVector_Sn … x) #hdx * #tlx #Heqx >Heqx |
---|
2103 | elim (BitVector_Sn … y) #hdy * #tly #Heqy >Heqy |
---|
2104 | elim (BitVector_Sn … delta) #hdd * #tld #Heqd >Heqd |
---|
2105 | #Hneq |
---|
2106 | cut (hdx ≠ hdy ∨ tlx ≠ tly) |
---|
2107 | [ @(eq_bv_elim … tlx tly) |
---|
2108 | [ #Heq_tl >Heq_tl >Heq_tl in Hneq; |
---|
2109 | #Hneq cut (hdx ≠ hdy) [ % #Heq_hd >Heq_hd in Hneq; * |
---|
2110 | #H @H @refl ] |
---|
2111 | #H %1 @H |
---|
2112 | | #H %2 @H ] ] |
---|
2113 | -Hneq #Hneq |
---|
2114 | >addition_n_direct_Sn >addition_n_direct_Sn |
---|
2115 | >ith_bit_Sn >ith_bit_Sn cases Hneq |
---|
2116 | [ 1: #Hneq_hd |
---|
2117 | lapply (addition_n_direct_inj … tlx tly tld) |
---|
2118 | @(eq_bv_elim … (addition_n_direct ? tlx tld false) (addition_n_direct ? tly tld false)) |
---|
2119 | [ 1: #Heq -Hind #Hind elim (Hind Heq) #Heq_tl >Heq_tl #Heq_carry >Heq_carry |
---|
2120 | % #Habsurd elim (bitvector_cons_inj_inv … Habsurd) -Habsurd |
---|
2121 | lapply Hneq_hd |
---|
2122 | cases hdx cases hdd cases hdy cases (ith_carry ? tly tld false) |
---|
2123 | normalize in ⊢ (? → % → ?); #Hneq_hd #Heq_hd #_ |
---|
2124 | try @(absurd … Heq_hd Hneq_hd) |
---|
2125 | elim Hneq_hd -Hneq_hd #Hneq_hd @Hneq_hd |
---|
2126 | try @refl try assumption try @(sym_eq … Heq_hd) |
---|
2127 | | 2: #Htl_not_eq #_ % #Habsurd elim (bitvector_cons_inj_inv … Habsurd) #_ |
---|
2128 | elim Htl_not_eq -Htl_not_eq #HA #HB @HA @HB ] |
---|
2129 | | 2: #Htl_not_eq lapply (Hind tlx tly tld Htl_not_eq) -Hind #Hind |
---|
2130 | % #Habsurd elim (bitvector_cons_inj_inv … Habsurd) #_ |
---|
2131 | elim Hind -Hind #HA #HB @HA @HB ] |
---|
2132 | ] qed. |
---|
2133 | |
---|
2134 | lemma carry_notb : ∀a,b,c. notb (carry_of a b c) = carry_of (notb a) (notb b) (notb c). * * * @refl qed. |
---|
2135 | |
---|
2136 | lemma increment_to_carry_aux : ∀n. ∀a : BitVector (S n). |
---|
2137 | ith_carry (S n) a (one_bv (S n)) false |
---|
2138 | = ith_carry (S n) a (zero (S n)) true. |
---|
2139 | #n elim n |
---|
2140 | [ 1: #a elim (BitVector_Sn ? a) #hd_a * #tl_a #Heq >Heq >(BitVector_O … tl_a) @refl |
---|
2141 | | 2: #n' #Hind #a elim (BitVector_Sn ? a) #hd_a * #tl_a #Heq >Heq |
---|
2142 | lapply (Hind tl_a) #Hind |
---|
2143 | >one_bv_Sn >zero_Sn >ith_carry_Sn >ith_carry_Sn >Hind @refl |
---|
2144 | ] qed. |
---|
2145 | |
---|
2146 | lemma neutral_addition_n_direct_aux : ∀n. ∀v. ith_carry n v (zero n) false = false. |
---|
2147 | #n elim n // |
---|
2148 | #n' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq >zero_Sn |
---|
2149 | >ith_carry_Sn >(Hind tl) cases hd @refl. |
---|
2150 | qed. |
---|
2151 | |
---|
2152 | lemma neutral_addition_n_direct : ∀n. ∀v : BitVector n. |
---|
2153 | addition_n_direct ? v (zero ?) false = v. |
---|
2154 | #n elim n |
---|
2155 | [ 1: #v >(BitVector_O … v) normalize @refl |
---|
2156 | | 2: #n' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq |
---|
2157 | lapply (Hind … tl) #H >zero_Sn >addition_n_direct_Sn |
---|
2158 | >ith_bit_Sn >H >xorb_false >neutral_addition_n_direct_aux |
---|
2159 | >xorb_false @refl |
---|
2160 | ] qed. |
---|
2161 | |
---|
2162 | lemma increment_to_carry_zero : ∀n. ∀a : BitVector n. addition_n_direct ? a (one_bv ?) false = addition_n_direct ? a (zero ?) true. |
---|
2163 | #n elim n |
---|
2164 | [ 1: #a >(BitVector_O … a) normalize @refl |
---|
2165 | | 2: #n' cases n' |
---|
2166 | [ 1: #_ #a elim (BitVector_Sn … a) #hd_a * #tl_a #Heq >Heq >(BitVector_O … tl_a) cases hd_a @refl |
---|
2167 | | 2: #n'' #Hind #a |
---|
2168 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heq >Heq |
---|
2169 | lapply (Hind tl_a) -Hind #Hind |
---|
2170 | >one_bv_Sn >zero_Sn >addition_n_direct_Sn >ith_bit_Sn |
---|
2171 | >addition_n_direct_Sn >ith_bit_Sn |
---|
2172 | >xorb_false >Hind @bitvector_cons_eq |
---|
2173 | >increment_to_carry_aux @refl |
---|
2174 | ] |
---|
2175 | ] qed. |
---|
2176 | |
---|
2177 | lemma increment_to_carry : ∀n. ∀a,b : BitVector n. |
---|
2178 | addition_n_direct ? a (addition_n_direct ? b (one_bv ?) false) false = addition_n_direct ? a b true. |
---|
2179 | #n #a #b >increment_to_carry_zero <associative_addition_n_direct |
---|
2180 | >neutral_addition_n_direct @refl |
---|
2181 | qed. |
---|
2182 | |
---|
2183 | lemma increment_direct_ok : ∀n,v. increment_direct n v = increment n v. |
---|
2184 | #n #v whd in match (increment ??); |
---|
2185 | >addition_n_direct_ok <increment_to_carry_zero @refl |
---|
2186 | qed. |
---|
2187 | |
---|
2188 | (* Prove -(a + b) = -a + -b *) |
---|
2189 | lemma twocomp_neg_plus : ∀n. ∀a,b : BitVector n. |
---|
2190 | twocomp_neg_direct ? (addition_n_direct ? a b false) = addition_n_direct ? (twocomp_neg_direct … a) (twocomp_neg_direct … b) false. |
---|
2191 | whd in match twocomp_neg_direct; normalize nodelta |
---|
2192 | lapply increment_inj_inv |
---|
2193 | whd in match increment_direct; normalize nodelta |
---|
2194 | #H #n #a #b |
---|
2195 | <associative_addition_n_direct @H |
---|
2196 | >associative_addition_n_direct >(commutative_addition_n_direct ? (one_bv n)) |
---|
2197 | >increment_to_carry |
---|
2198 | -H lapply b lapply a -b -a |
---|
2199 | cases n |
---|
2200 | [ 1: #a #b >(BitVector_O … a) >(BitVector_O … b) @refl |
---|
2201 | | 2: #n' #a #b |
---|
2202 | cut (negation_bv ? (addition_n_direct ? a b false) |
---|
2203 | = addition_n_direct ? (negation_bv ? a) (negation_bv ? b) true ∧ |
---|
2204 | notb (ith_carry ? a b false) = (ith_carry ? (negation_bv ? a) (negation_bv ? b) true)) |
---|
2205 | [ -n lapply b lapply a elim n' |
---|
2206 | [ 1: #a #b elim (BitVector_Sn … a) #hd_a * #tl_a #Heqa >Heqa >(BitVector_O … tl_a) |
---|
2207 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heqb >Heqb >(BitVector_O … tl_b) |
---|
2208 | cases hd_a cases hd_b normalize @conj @refl |
---|
2209 | | 2: #n #Hind #a #b |
---|
2210 | elim (BitVector_Sn … a) #hd_a * #tl_a #Heqa >Heqa |
---|
2211 | elim (BitVector_Sn … b) #hd_b * #tl_b #Heqb >Heqb |
---|
2212 | lapply (Hind tl_a tl_b) * #H1 #H2 |
---|
2213 | @conj |
---|
2214 | [ 2: >ith_carry_Sn >negation_bv_Sn >negation_bv_Sn >ith_carry_Sn |
---|
2215 | >carry_notb >H2 @refl |
---|
2216 | | 1: >addition_n_direct_Sn >ith_bit_Sn >negation_bv_Sn |
---|
2217 | >negation_bv_Sn >negation_bv_Sn |
---|
2218 | >addition_n_direct_Sn >ith_bit_Sn >H1 @bitvector_cons_eq |
---|
2219 | >xorb_lneg >xorb_rneg >notb_notb |
---|
2220 | <xorb_rneg >H2 @refl |
---|
2221 | ] |
---|
2222 | ] ] |
---|
2223 | * #H1 #H2 @H1 |
---|
2224 | ] qed. |
---|
2225 | |
---|
2226 | lemma addition_n_direct_neg : ∀n. ∀a. |
---|
2227 | (addition_n_direct n a (negation_bv n a) false) = replicate ?? true |
---|
2228 | ∧ (ith_carry n a (negation_bv n a) false = false). |
---|
2229 | #n elim n |
---|
2230 | [ 1: #a >(BitVector_O … a) @conj @refl |
---|
2231 | | 2: #n' #Hind #a elim (BitVector_Sn … a) #hd * #tl #Heq >Heq |
---|
2232 | lapply (Hind … tl) -Hind * #HA #HB |
---|
2233 | @conj |
---|
2234 | [ 2: >negation_bv_Sn >ith_carry_Sn >HB cases hd @refl |
---|
2235 | | 1: >negation_bv_Sn >addition_n_direct_Sn |
---|
2236 | >ith_bit_Sn >HB >xorb_false >HA |
---|
2237 | @bitvector_cons_eq elim hd @refl |
---|
2238 | ] |
---|
2239 | ] qed. |
---|
2240 | |
---|
2241 | (* -a + a = 0 *) |
---|
2242 | lemma bitvector_opp_direct : ∀n. ∀a : BitVector n. addition_n_direct ? a (twocomp_neg_direct ? a) false = (zero ?). |
---|
2243 | whd in match twocomp_neg_direct; |
---|
2244 | whd in match increment_direct; |
---|
2245 | normalize nodelta |
---|
2246 | #n #a <associative_addition_n_direct |
---|
2247 | elim (addition_n_direct_neg … a) #H #_ >H |
---|
2248 | -H -a |
---|
2249 | cases n try // |
---|
2250 | #n' |
---|
2251 | cut ((addition_n_direct (S n') (replicate bool ? true) (one_bv ?) false = (zero (S n'))) |
---|
2252 | ∧ (ith_carry ? (replicate bool (S n') true) (one_bv (S n')) false = true)) |
---|
2253 | [ elim n' |
---|
2254 | [ 1: @conj @refl |
---|
2255 | | 2: #n' * #HA #HB @conj |
---|
2256 | [ 1: >replicate_Sn >one_bv_Sn >addition_n_direct_Sn |
---|
2257 | >ith_bit_Sn >HA >zero_Sn @bitvector_cons_eq >HB @refl |
---|
2258 | | 2: >replicate_Sn >one_bv_Sn >ith_carry_Sn >HB @refl ] |
---|
2259 | ] |
---|
2260 | ] * #H1 #H2 @H1 |
---|
2261 | qed. |
---|
2262 | |
---|
2263 | (* Lift back the previous result to standard operations. *) |
---|
2264 | lemma twocomp_neg_direct_ok : ∀n. ∀v. twocomp_neg_direct ? v = two_complement_negation n v. |
---|
2265 | #n #v whd in match twocomp_neg_direct; normalize nodelta |
---|
2266 | whd in match increment_direct; normalize nodelta |
---|
2267 | whd in match two_complement_negation; normalize nodelta |
---|
2268 | >increment_to_addition_n <addition_n_direct_ok |
---|
2269 | whd in match addition_n; normalize nodelta |
---|
2270 | elim (add_with_carries ????) #a #b @refl |
---|
2271 | qed. |
---|
2272 | |
---|
2273 | lemma two_complement_negation_plus : ∀n. ∀a,b : BitVector n. |
---|
2274 | two_complement_negation ? (addition_n ? a b) = addition_n ? (two_complement_negation ? a) (two_complement_negation ? b). |
---|
2275 | #n #a #b |
---|
2276 | lapply (twocomp_neg_plus ? a b) |
---|
2277 | >twocomp_neg_direct_ok >twocomp_neg_direct_ok >twocomp_neg_direct_ok |
---|
2278 | <addition_n_direct_ok <addition_n_direct_ok |
---|
2279 | whd in match addition_n; normalize nodelta |
---|
2280 | elim (add_with_carries n a b false) #bits #flags normalize nodelta |
---|
2281 | elim (add_with_carries n (two_complement_negation n a) (two_complement_negation n b) false) #bits' #flags' |
---|
2282 | normalize nodelta #H @H |
---|
2283 | qed. |
---|
2284 | |
---|
2285 | lemma bitvector_opp_addition_n : ∀n. ∀a : BitVector n. addition_n ? a (two_complement_negation ? a) = (zero ?). |
---|
2286 | #n #a lapply (bitvector_opp_direct ? a) |
---|
2287 | >twocomp_neg_direct_ok <addition_n_direct_ok |
---|
2288 | whd in match (addition_n ???); |
---|
2289 | elim (add_with_carries n a (two_complement_negation n a) false) #bits #flags #H @H |
---|
2290 | qed. |
---|
2291 | |
---|
2292 | lemma neutral_addition_n : ∀n. ∀a : BitVector n. addition_n ? a (zero ?) = a. |
---|
2293 | #n #a |
---|
2294 | lapply (neutral_addition_n_direct n a) |
---|
2295 | <addition_n_direct_ok |
---|
2296 | whd in match (addition_n ???); |
---|
2297 | elim (add_with_carries n a (zero n) false) #bits #flags #H @H |
---|
2298 | qed. |
---|
2299 | |
---|
2300 | lemma injective_addition_n : ∀n. ∀x,y,delta : BitVector n. |
---|
2301 | addition_n ? x delta = addition_n ? y delta → x = y. |
---|
2302 | #n #x #y #delta |
---|
2303 | lapply (addition_n_direct_inj … x y delta) |
---|
2304 | <addition_n_direct_ok <addition_n_direct_ok |
---|
2305 | whd in match addition_n; normalize nodelta |
---|
2306 | elim (add_with_carries n x delta false) #bitsx #flagsx |
---|
2307 | elim (add_with_carries n y delta false) #bitsy #flagsy |
---|
2308 | normalize #H1 #H2 elim (H1 H2) #Heq #_ @Heq |
---|
2309 | qed. |
---|
2310 | |
---|
2311 | lemma injective_inv_addition_n : ∀n. ∀x,y,delta : BitVector n. |
---|
2312 | x ≠ y → addition_n ? x delta ≠ addition_n ? y delta. |
---|
2313 | #n #x #y #delta |
---|
2314 | lapply (addition_n_direct_inj_inv … x y delta) |
---|
2315 | <addition_n_direct_ok <addition_n_direct_ok |
---|
2316 | whd in match addition_n; normalize nodelta |
---|
2317 | elim (add_with_carries n x delta false) #bitsx #flagsx |
---|
2318 | elim (add_with_carries n y delta false) #bitsy #flagsy |
---|
2319 | normalize #H1 #H2 @(H1 H2) |
---|
2320 | qed. |
---|
2321 | |
---|