(* Various small homeless results. *) include "Clight/TypeComparison.ma". include "Clight/Csem.ma". include "common/Pointers.ma". include "basics/jmeq.ma". include "basics/star.ma". (* well-founded relations *) include "common/IOMonad.ma". include "common/IO.ma". include "basics/lists/listb.ma". include "basics/lists/list.ma". (* --------------------------------------------------------------------------- *) (* [cthulhu] plays the same role as daemon. It should be droppable. *) (* --------------------------------------------------------------------------- *) axiom cthulhu : ∀A:Prop. A. (* Because of the nightmares. *) (* --------------------------------------------------------------------------- *) (* Misc. *) (* --------------------------------------------------------------------------- *) lemma eq_intsize_identity : ∀id. eq_intsize id id = true. * normalize // qed. lemma sz_eq_identity : ∀s. sz_eq_dec s s = inl ?? (refl ? s). * normalize // qed. lemma type_eq_identity : ∀t. type_eq t t = true. #t normalize elim (type_eq_dec t t) [ 1: #Heq normalize // | 2: #H destruct elim H #Hcontr elim (Hcontr (refl ? t)) ] qed. lemma type_neq_not_identity : ∀t1, t2. t1 ≠ t2 → type_eq t1 t2 = false. #t1 #t2 #Hneq normalize elim (type_eq_dec t1 t2) [ 1: #Heq destruct elim Hneq #Hcontr elim (Hcontr (refl ? t2)) | 2: #Hneq' normalize // ] qed. lemma le_S_O_absurd : ∀x:nat. S x ≤ 0 → False. /2 by absurd/ qed. lemma some_inj : ∀A : Type[0]. ∀a,b : A. Some ? a = Some ? b → a = b. #A #a #b #H destruct (H) @refl qed. lemma prod_eq_lproj : ∀A,B : Type[0]. ∀a : A. ∀b : B. ∀c : A × B. 〈a,b〉 = c → a = \fst c. #A #B #a #b * #a' #b' #H destruct @refl qed. lemma prod_eq_rproj : ∀A,B : Type[0]. ∀a : A. ∀b : B. ∀c : A × B. 〈a,b〉 = c → b = \snd c. #A #B #a #b * #a' #b' #H destruct @refl qed. 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. // qed. lemma bindIO_Value : ∀v,f. bindIO io_out io_in (val×trace) (trace×state) (Value … v) f = (f v). // qed. lemma bindIO_elim : ∀A. ∀P : (IO io_out io_in A) → Prop. ∀e : res A. (*IO io_out io_in A.*) ∀f. (∀v. (e:IO io_out io_in A) = OK … v → P (f v)) → (∀err. (e:IO io_out io_in A) = Error … err → P (Wrong ??? err)) → P (bindIO io_out io_in ? A (e:IO io_out io_in A) f). #A #P * try /2/ qed. lemma opt_to_io_Value : ∀A,B,C,err,x,res. opt_to_io A B C err x = return res → x = Some ? res. #A #B #C #err #x cases x normalize [ 1: #res #Habsurd destruct | 2: #c #res #Heq destruct @refl ] qed. lemma some_inversion : ∀A,B:Type[0]. ∀e:option A. ∀res:B. ∀f:A → option B. match e with [ None ⇒ None ? | Some x ⇒ f x ] = Some ? res → ∃x. e = Some ? x ∧ f x = Some ? res. #A #B #e #res #f cases e normalize nodelta [ 1: #Habsurd destruct (Habsurd) | 2: #a #Heq %{a} @conj >Heq @refl ] qed. lemma res_inversion : ∀A,B:Type[0]. ∀e:option A. ∀errmsg. ∀result:B. ∀f:A → res B. match e with [ None ⇒ Error ? errmsg | Some x ⇒ f x ] = OK ? result → ∃x. e = Some ? x ∧ f x = OK ? result. #A #B #e #errmsg #result #f cases e normalize nodelta [ 1: #Habsurd destruct (Habsurd) | 2: #a #Heq %{a} @conj >Heq @refl ] qed. lemma cons_inversion : ∀A,B:Type[0]. ∀e:list A. ∀res:B. ∀f:A → list A → option B. match e with [ nil ⇒ None ? | cons hd tl ⇒ f hd tl ] = Some ? res → ∃hd,tl. e = hd::tl ∧ f hd tl = Some ? res. #A #B #e #res #f cases e normalize nodelta [ 1: #Habsurd destruct (Habsurd) | 2: #hd #tl #Heq %{hd} %{tl} @conj >Heq @refl ] qed. lemma if_opt_inversion : ∀A:Type[0]. ∀x. ∀y:A. ∀e:bool. (if e then x else None ?) = Some ? y → e = true ∧ x = Some ? y. #A #x #y * normalize #H destruct @conj @refl qed. lemma opt_to_res_inversion : ∀A:Type[0]. ∀errmsg. ∀opt. ∀val. opt_to_res A errmsg opt = OK ? val → opt = Some ? val. #A #errmsg * [ 1: #val normalize #Habsurd destruct | 2: #res #val normalize #Heq destruct @refl ] qed. lemma andb_inversion : ∀a,b : bool. andb a b = true → a = true ∧ b = true. * * normalize /2 by conj, refl/ qed. lemma identifier_eq_i_i : ∀tag,i. ∃pf. identifier_eq tag i i = inl … pf. #tag #i cases (identifier_eq ? i i) [ 1: #H %{H} @refl | 2: * #Habsurd @(False_ind … (Habsurd … (refl ? i))) ] qed. lemma intsize_eq_inversion : ∀sz,sz'. ∀A:Type[0]. ∀ok,not_ok. intsize_eq_elim' sz sz' (λsz,sz'. res A) (OK ? ok) (Error ? not_ok) = (OK ? ok) → sz = sz'. * * try // normalize #A #ok #not_ok #Habsurd destruct qed. lemma intsize_eq_elim_dec : ∀sz1,sz2. ∀P: ∀sz1,sz2. Type[0]. ((∀ifok,iferr. intsize_eq_elim' sz1 sz1 P ifok iferr = ifok) ∧ sz1 = sz2) ∨ ((∀ifok,iferr. intsize_eq_elim' sz1 sz2 P ifok iferr = iferr) ∧ sz1 ≠ sz2). * * #P normalize try /3 by or_introl, conj, refl/ %2 @conj try // % #H destruct qed. lemma typ_eq_elim : ∀t1,t2. ∀(P: (t1=t2)+(t1≠t2) → Prop). (∀H:t1 = t2. P (inl ?? H)) → (∀H:t1 ≠ t2. P (inr ?? H)) → P (typ_eq t1 t2). #t1 #t2 #P #Hl #Hr @(match typ_eq t1 t2 with [ inl H ⇒ Hl H | inr H ⇒ Hr H ]) qed. lemma eq_nat_dec_true : ∀n. eq_nat_dec n n = inl ?? (refl ? n). #n elim n try // #n' #Hind whd in ⊢ (??%?); >Hind @refl qed. lemma type_eq_dec_true : ∀ty. type_eq_dec ty ty = inl ?? (refl ? ty). #ty cases (type_eq_dec ty ty) #H destruct (H) try @refl @False_ind cases H #J @J @refl qed. lemma typ_eq_refl : ∀t. typ_eq t t = inl ?? (refl ? t). * [ * * normalize @refl | @refl ] qed. lemma intsize_eq_elim_inversion : ∀A:Type[0]. ∀sz1,sz2. ∀elt1,f,errmsg,res. intsize_eq_elim ? sz1 sz2 bvint elt1 f (Error A errmsg) = OK ? res → ∃H:sz1 = sz2. OK ? res = (f (eq_rect_r ? sz1 sz2 (sym_eq ??? H) ? elt1)). #A * * #elt1 #f #errmsg #res normalize #H destruct (H) %{(refl ??)} normalize nodelta >H @refl qed. lemma inttyp_eq_elim_true' : ∀sz,sg,P,p1,p2. inttyp_eq_elim' sz sz sg sg P p1 p2 = p1. * * #P #p1 #p2 normalize try @refl qed. (* --------------------------------------------------------------------------- *) (* Useful facts on various boolean operations. *) (* --------------------------------------------------------------------------- *) lemma andb_lsimpl_true : ∀x. andb true x = x. // qed. lemma andb_lsimpl_false : ∀x. andb false x = false. normalize // qed. lemma andb_comm : ∀x,y. andb x y = andb y x. // qed. lemma notb_true : notb true = false. // qed. lemma notb_false : notb false = true. % #H destruct qed. lemma notb_fold : ∀x. if x then false else true = (¬x). // qed. (* --------------------------------------------------------------------------- *) (* Useful facts on Z. *) (* --------------------------------------------------------------------------- *) lemma Zltb_to_Zleb_true : ∀a,b. Zltb a b = true → Zleb a b = true. #a #b #Hltb lapply (Zltb_true_to_Zlt … Hltb) #Hlt @Zle_to_Zleb_true /3 by Zlt_to_Zle, transitive_Zle/ qed. lemma Zle_to_Zle_to_eq : ∀a,b. Zle a b → Zle b a → a = b. #a #b elim b [ 1: elim a // #a' normalize [ 1: @False_ind | 2: #_ @False_ind ] ] #b' elim a normalize [ 1: #_ @False_ind | 2: #a' #H1 #H2 >(le_to_le_to_eq … H1 H2) @refl | 3: #a' #_ @False_ind | 4: @False_ind | 5: #a' @False_ind | 6: #a' #H2 #H1 >(le_to_le_to_eq … H1 H2) @refl ] qed. lemma Zleb_true_to_Zleb_true_to_eq : ∀a,b. (Zleb a b = true) → (Zleb b a = true) → a = b. #a #b #H1 #H2 /3 by Zle_to_Zle_to_eq, Zleb_true_to_Zle/ qed. lemma Zltb_dec : ∀a,b. (Zltb a b = true) ∨ (Zltb a b = false ∧ Zleb b a = true). #a #b lapply (Z_compare_to_Prop … a b) cases a [ 1: | 2,3: #a' ] cases b whd in match (Z_compare OZ OZ); normalize nodelta [ 2,3,5,6,8,9: #b' ] whd in match (Zleb ? ?); try /3 by or_introl, or_intror, conj, I, refl/ whd in match (Zltb ??); whd in match (Zleb ??); #_ [ 1: cases (decidable_le (succ a') b') [ 1: #H lapply (le_to_leb_true … H) #H %1 assumption | 2: #Hnotle lapply (not_le_to_lt … Hnotle) #Hlt %2 @conj try @le_to_leb_true /2 by monotonic_pred/ @(not_le_to_leb_false … Hnotle) ] | 2: cases (decidable_le (succ b') a') [ 1: #H lapply (le_to_leb_true … H) #H %1 assumption | 2: #Hnotle lapply (not_le_to_lt … Hnotle) #Hlt %2 @conj try @le_to_leb_true /2 by monotonic_pred/ @(not_le_to_leb_false … Hnotle) ] ] qed. lemma Zleb_unsigned_OZ : ∀bv. Zleb OZ (Z_of_unsigned_bitvector 16 bv) = true. #bv elim bv try // #n' * #tl normalize /2/ qed. lemma Zltb_unsigned_OZ : ∀bv. Zltb (Z_of_unsigned_bitvector 16 bv) OZ = false. #bv elim bv try // #n' * #tl normalize /2/ qed. lemma Z_of_unsigned_not_neg : ∀bv. (Z_of_unsigned_bitvector 16 bv = OZ) ∨ (∃p. Z_of_unsigned_bitvector 16 bv = pos p). #bv elim bv [ 1: normalize %1 @refl | 2: #n #hd #tl #Hind cases hd [ 1: normalize %2 /2 by ex_intro/ | 2: cases Hind normalize [ 1: #H %1 @H | 2: * #p #H >H %2 %{p} @refl ] ] ] qed. lemma free_not_in_bounds : ∀x. if Zleb (pos one) x then Zltb x OZ  else false = false. #x lapply (Zltb_to_Zleb_true x OZ) elim (Zltb_dec … x OZ) [ 1: #Hlt >Hlt #H lapply (H (refl ??)) elim x [ 2,3: #x' ] normalize in ⊢ (% → ?); [ 1: #Habsurd destruct (Habsurd) | 2,3: #_ @refl ] | 2: * #Hlt #Hle #_ >Hlt cases (Zleb (pos one) x) @refl ] qed. lemma free_not_valid : ∀x. if Zleb (pos one) x then Zltb x OZ  else false = false. #x cut (Zle (pos one) x ∧ Zlt x OZ → False) [ * #H1 #H2 lapply (transitive_Zle … (monotonic_Zle_Zsucc … H1) (Zlt_to_Zle … H2)) #H @H ] #Hguard cut (Zleb (pos one) x = true ∧ Zltb x OZ = true → False) [ * #H1 #H2 @Hguard @conj /2 by Zleb_true_to_Zle/ @Zltb_true_to_Zlt assumption ] cases (Zleb (pos one) x) cases (Zltb x OZ) /2 by refl/ #H @(False_ind … (H (conj … (refl ??) (refl ??)))) qed. (* follows from (0 ≤ a < b → mod a b = a) *) axiom Z_of_unsigned_bitvector_of_small_Z : ∀z. OZ ≤ z → z < Z_two_power_nat 16 → Z_of_unsigned_bitvector 16 (bitvector_of_Z 16 z) = z. theorem Zle_to_Zlt_to_Zlt: ∀n,m,p:Z. n ≤ m → m < p → n < p. #n #m #p #Hle #Hlt /4 by monotonic_Zle_Zplus_r, Zle_to_Zlt, Zlt_to_Zle, transitive_Zle/ qed. (* --------------------------------------------------------------------------- *) (* Useful facts on blocks. *) (* --------------------------------------------------------------------------- *) lemma eq_block_to_refl : ∀b1,b2. eq_block b1 b2 = true → b1 = b2. #b1 #b2 @(eq_block_elim … b1 b2) [ // | #_ #Habsurd destruct ] qed. lemma not_eq_block : ∀b1,b2. b1 ≠ b2 → eq_block b1 b2 = false. #b1 #b2 #Hneq @(eq_block_elim … b1 b2) [ 1: #Heq destruct elim Hneq #H @(False_ind … (H (refl ??))) | 2: #_ @refl ] qed. lemma not_eq_block_rev : ∀b1,b2. b2 ≠ b1 → eq_block b1 b2 = false. #b1 #b2 #Hneq @(eq_block_elim … b1 b2) [ 1: #Heq destruct elim Hneq #H @(False_ind … (H (refl ??))) | 2: #_ @refl ] qed. definition block_DeqSet : DeqSet ≝ mk_DeqSet block eq_block ?. * #r1 #id1 * #r2 #id2 @(eqZb_elim … id1 id2) [ 1: #Heq >Heq cases r1 cases r2 normalize >eqZb_z_z normalize try // @conj #H destruct (H) try @refl | 2: #Hneq cases r1 cases r2 normalize >(eqZb_false … Hneq) normalize @conj #H destruct (H) elim Hneq #H @(False_ind … (H (refl ??))) ] qed. (* --------------------------------------------------------------------------- *) (* General results on lists. *) (* --------------------------------------------------------------------------- *) let rec mem_assoc_env (i : ident) (l : list (ident×type)) on l : bool ≝ match l with [ nil ⇒ false | cons hd tl ⇒ let 〈id, ty〉 ≝ hd in match identifier_eq SymbolTag i id with [ inl Hid_eq ⇒ true | inr Hid_neq ⇒ mem_assoc_env i tl ] ]. (* If mem succeeds in finding an element, then the list can be partitioned around this element. *) lemma list_mem_split : ∀A. ∀l : list A. ∀x : A. mem … x l → ∃l1,l2. l = l1 @ [x] @ l2. #A #l elim l [ 1: normalize #x @False_ind | 2: #hd #tl #Hind #x whd in ⊢ (% → ?); * [ 1: #Heq %{(nil ?)} %{tl} destruct @refl | 2: #Hmem lapply (Hind … Hmem) * #l1 * #l2 #Heq_tl >Heq_tl %{(hd :: l1)} %{l2} @refl ] ] qed. lemma cons_to_append : ∀A. ∀hd : A. ∀l : list A. hd :: l = [hd] @ l. #A #hd #l @refl qed. lemma fold_append : ∀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. #A #B #l1 elim l1 // #hd #tl #Hind #l2 #f #seed normalize >Hind @refl qed. lemma filter_append : ∀A:Type[0]. ∀l1,l2 : list A. ∀f. filter ? f (l1 @ l2) = (filter ? f l1) @ (filter ? f l2). #A #l1 elim l1 // #hd #tl #Hind #l2 #f >cons_to_append >associative_append normalize cases (f hd) normalize filter_cons_unfold lapply (eqb_true A hd elt) cases (hd==elt) normalize nodelta [ 2: #_ #Habsurd destruct | 1: * #H1 #H2 #Heq lapply (Hind Heq) #Hall whd @conj // @H1 @refl ] ] qed. lemma nil_append : ∀A. ∀(l : list A). [ ] @ l = l. // qed. alias id "meml" = "cic:/matita/basics/lists/list/mem.fix(0,2,1)". lemma mem_append : ∀A:Type[0]. ∀elt : A. ∀l1,l2. mem … elt (l1 @ l2) ↔ (mem … elt l1) ∨ (mem … elt l2). #A #elt #l1 elim l1 [ 1: #l2 normalize @conj [ 1: #H %2 @H | 2: * [ 1: @False_ind | 2: #H @H ] ] | 2: #hd #tl #Hind #l2 @conj [ 1: whd in match (meml ???); * [ 1: #Heq >Heq %1 normalize %1 @refl | 2: #H1 lapply (Hind l2) * #HA #HB normalize elim (HA H1) [ 1: #H %1 %2 assumption | 2: #H %2 assumption ] ] | 2: normalize * [ 1: * /2 by or_introl, or_intror/ #H %2 elim (Hind l2) #_ #H' @H' %1 @H | 2: #H %2 elim (Hind l2) #_ #H' @H' %2 @H ] ] ] qed. lemma mem_append_forward : ∀A:Type[0]. ∀elt : A. ∀l1,l2. mem … elt (l1 @ l2) → (mem … elt l1) ∨ (mem … elt l2). #A #elt #l1 #l2 #H elim (mem_append … elt l1 l2) #H' #_ @H' @H qed. lemma mem_append_backwards : ∀A:Type[0]. ∀elt : A. ∀l1,l2. (mem … elt l1) ∨ (mem … elt l2) → mem … elt (l1 @ l2) . #A #elt #l1 #l2 #H elim (mem_append … elt l1 l2) #_ #H' @H' @H qed. (* "Observational" equivalence on list implies concrete equivalence. Useful to * prove equality from some reasoning on indexings. Needs a particular induction * principle. *) let rec double_list_ind (A : Type[0]) (P : list A → list A → Prop) (base_nil : P [ ] [ ]) (base_l1 : ∀hd1,l1. P (hd1::l1) [ ]) (base_l2 : ∀hd2,l2. P [ ] (hd2::l2)) (ind : ∀hd1,hd2,tl1,tl2. P tl1 tl2 → P (hd1 :: tl1) (hd2 :: tl2)) (l1, l2 : list A) on l1 ≝ match l1 with [ nil ⇒ match l2 with [ nil ⇒ base_nil | cons hd2 tl2 ⇒ base_l2 hd2 tl2 ] | cons hd1 tl1 ⇒ match l2 with [ nil ⇒ base_l1 hd1 tl1 | cons hd2 tl2 ⇒ ind hd1 hd2 tl1 tl2 (double_list_ind A P base_nil base_l1 base_l2 ind tl1 tl2) ] ]. lemma nth_eq_tl : ∀A:Type[0]. ∀l1,l2:list A. ∀hd1,hd2. (∀i. nth_opt A i (hd1::l1) = nth_opt A i (hd2::l2)) → (∀i. nth_opt A i l1 = nth_opt A i l2). #A #l1 #l2 @(double_list_ind … l1 l2) [ 1: #hd1 #hd2 #_ #i elim i try /2/ | 2: #hd1' #l1' #hd1 #hd2 #H lapply (H 1) normalize #Habsurd destruct | 3: #hd2' #l2' #hd1 #hd2 #H lapply (H 1) normalize #Habsurd destruct | 4: #hd1' #hd2' #tl1' #tl2' #H #hd1 #hd2 #Hind @(λi. Hind (S i)) ] qed. lemma nth_eq_to_eq : ∀A:Type[0]. ∀l1,l2:list A. (∀i. nth_opt A i l1 = nth_opt A i l2) → l1 = l2. #A #l1 elim l1 [ 1: #l2 #H lapply (H 0) normalize cases l2 [ 1: // | 2: #hd2 #tl2 normalize #Habsurd destruct ] | 2: #hd1 #tl1 #Hind * [ 1: #H lapply (H 0) normalize #Habsurd destruct | 2: #hd2 #tl2 #H lapply (H 0) normalize #Heq destruct (Heq) >(Hind tl2) try @refl @(nth_eq_tl … H) ] ] qed. (* --------------------------------------------------------------------------- *) (* General results on vectors. *) (* --------------------------------------------------------------------------- *) (* copied from AssemblyProof, TODO get rid of the redundant stuff. *) lemma Vector_O: ∀A. ∀v: Vector A 0. v ≃ VEmpty A. #A #v generalize in match (refl … 0); cases v in ⊢ (??%? → ?%%??); // #n #hd #tl #abs @⊥ destruct (abs) qed. lemma Vector_Sn: ∀A. ∀n.∀v:Vector A (S n). ∃hd.∃tl.v ≃ VCons A n hd tl. #A #n #v generalize in match (refl … (S n)); cases v in ⊢ (??%? → ??(λ_.??(λ_.?%%??))); [ #abs @⊥ destruct (abs) | #m #hd #tl #EQ <(injective_S … EQ) %[@hd] %[@tl] // ] qed. lemma vector_append_zero: ∀A,m. ∀v: Vector A m. ∀q: Vector A 0. v = q@@v. #A #m #v #q >(Vector_O A q) % qed. corollary prod_vector_zero_eq_left: ∀A, n. ∀q: Vector A O. ∀r: Vector A n. 〈q, r〉 = 〈[[ ]], r〉. #A #n #q #r generalize in match (Vector_O A q …); #hyp >hyp in ⊢ (??%?); % qed. lemma vsplit_eq : ∀A. ∀m,n. ∀v : Vector A ((S m) + n). ∃v1:Vector A (S m). ∃v2:Vector A n. v = v1 @@ v2. # A #m #n elim m [ 1: normalize #v elim (Vector_Sn ?? v) #hd * #tl #Eq @(ex_intro … (hd ::: [[]])) @(ex_intro … tl) >Eq normalize // | 2: #n' #Hind #v elim (Vector_Sn ?? v) #hd * #tl #Eq elim (Hind tl) #tl1 * #tl2 #Eq_tl @(ex_intro … (hd ::: tl1)) @(ex_intro … tl2) destruct normalize // ] qed. lemma vsplit_zero: ∀A,m. ∀v: Vector A m. 〈[[]], v〉 = vsplit A 0 m v. #A #m #v elim v [ % | #n #hd #tl #ih normalize in ⊢ (???%); % ] qed. lemma vsplit_zero2: ∀A,m. ∀v: Vector A m. 〈[[]], v〉 = vsplit' A 0 m v. #A #m #v elim v [ % | #n #hd #tl #ih normalize in ⊢ (???%); % ] qed. lemma vsplit_eq2 : ∀A. ∀m,n : nat. ∀v : Vector A (m + n). ∃v1:Vector A m. ∃v2:Vector A n. v = v1 @@ v2. # A #m #n elim m [ 1: normalize #v @(ex_intro … (VEmpty ?)) @(ex_intro … v) normalize // | 2: #n' #Hind #v elim (Vector_Sn ?? v) #hd * #tl #Eq elim (Hind tl) #tl1 * #tl2 #Eq_tl @(ex_intro … (hd ::: tl1)) @(ex_intro … tl2) destruct normalize // ] qed. (* This is not very nice. Note that this axiom was copied verbatim from ASM/AssemblyProof.ma. * TODO sync with AssemblyProof.ma, in a better world we shouldn't have to copy all of this. *) axiom vsplit_succ: ∀A, m, n. ∀l: Vector A m. ∀r: Vector A n. ∀v: Vector A (m + n). ∀hd. v = l@@r → (〈l, r〉 = vsplit ? m n v → 〈hd:::l, r〉 = vsplit ? (S m) n (hd:::v)). axiom vsplit_succ2: ∀A, m, n. ∀l: Vector A m. ∀r: Vector A n. ∀v: Vector A (m + n). ∀hd. v = l@@r → (〈l, r〉 = vsplit' ? m n v → 〈hd:::l, r〉 = vsplit' ? (S m) n (hd:::v)). lemma vsplit_prod2: ∀A,m,n. ∀p: Vector A (m + n). ∀v: Vector A m. ∀q: Vector A n. p = v@@q → 〈v, q〉 = vsplit' A m n p. #A #m elim m [ #n #p #v #q #hyp >hyp <(vector_append_zero A n q v) >(prod_vector_zero_eq_left A …) @vsplit_zero2 | #r #ih #n #p #v #q #hyp >hyp cases (Vector_Sn A r v) #hd #exists cases exists #tl #jmeq >jmeq @vsplit_succ2 [1: % |2: @ih % ] ] qed. lemma vsplit_prod: ∀A,m,n. ∀p: Vector A (m + n). ∀v: Vector A m. ∀q: Vector A n. p = v@@q → 〈v, q〉 = vsplit A m n p. #A #m elim m [ #n #p #v #q #hyp >hyp <(vector_append_zero A n q v) >(prod_vector_zero_eq_left A …) @vsplit_zero | #r #ih #n #p #v #q #hyp >hyp cases (Vector_Sn A r v) #hd #exists cases exists #tl #jmeq >jmeq @vsplit_succ [1: % |2: @ih % ] ] qed. (* --------------------------------------------------------------------------- *) (* Some more stuff on bitvectors. *) (* --------------------------------------------------------------------------- *) axiom commutative_multiplication : ∀n. ∀v1,v2:BitVector n. multiplication ? v1 v2 = multiplication ? v2 v1. lemma commutative_short_multiplication : ∀n. ∀v1,v2:BitVector n. short_multiplication ? v1 v2 = short_multiplication ? v2 v1. #n #v1 #v2 whd in ⊢ (??%%); >commutative_multiplication @refl qed. lemma sign_ext_same_size : ∀n,v. sign_ext n n v = v. #n #v whd in match (sign_ext ???); >nat_compare_eq @refl qed. lemma zero_ext_same_size : ∀n,v. zero_ext n n v = v. #n #v whd in match (zero_ext ???); >nat_compare_eq @refl qed. axiom sign_ext_zero : ∀sz1,sz2. sign_ext sz1 sz2 (zero sz1) = zero sz2. axiom zero_ext_zero : ∀sz1,sz2. zero_ext sz1 sz2 (zero sz1) = zero sz2. (* notice that we restrict source and target sizes to be ≠ 0 *) axiom zero_ext_one : ∀sz1,sz2. zero_ext (bitsize_of_intsize sz1) (bitsize_of_intsize sz2) (repr sz1 1) = (repr sz2 1). axiom multiplication_zero : ∀n:nat. ∀v : BitVector n. multiplication … (zero ?) v = (zero ?). axiom short_multiplication_zero : ∀n. ∀v:BitVector n. short_multiplication ? (zero ?) v = (zero ?). (* dividing zero by something eq zero, not the other way around ofc. *) axiom division_u_zero : ∀sz.∀v:BitVector ?. division_u sz (bv_zero ?) v = Some ? (bv_zero ?). (* lemma eq_v_to_eq_Z : ∀n. ∀v1,v2:BitVector n. (Z_of_bitvector … v1) = (Z_of_bitvector eq_bv … v1 v2. *) (* --------------------------------------------------------------------------- *) (* Generic properties of equivalence relations *) (* --------------------------------------------------------------------------- *) lemma triangle_diagram : ∀acctype : Type[0]. ∀eqrel : acctype → acctype → Prop. ∀refl_eqrel : reflexive ? eqrel. ∀trans_eqrel : transitive ? eqrel. ∀sym_eqrel : symmetric ? eqrel. ∀a,b,c. eqrel a b → eqrel a c → eqrel b c. #acctype #eqrel #R #T #S #a #b #c #H1 #H2 @(T … (S … H1) H2) qed. lemma cotriangle_diagram : ∀acctype : Type[0]. ∀eqrel : acctype → acctype → Prop. ∀refl_eqrel : reflexive ? eqrel. ∀trans_eqrel : transitive ? eqrel. ∀sym_eqrel : symmetric ? eqrel. ∀a,b,c. eqrel b a → eqrel c a → eqrel b c. #acctype #eqrel #R #T #S #a #b #c #H1 #H2 @S @(T … H2 (S … H1)) qed. (* --------------------------------------------------------------------------- *) (* Quick and dirty implementation of finite sets relying on lists. The * implementation is split in two: an abstract equivalence defined using inclusion * of lists, and a concrete one where equivalence is defined as the closure of * duplication, contraction and transposition of elements. We rely on the latter * to prove stuff on folds over sets. *) (* --------------------------------------------------------------------------- *) definition lset ≝ λA:Type[0]. list A. (* The empty set. *) definition empty_lset ≝ λA:Type[0]. nil A. (* Standard operations. *) definition lset_union ≝ λA:Type[0].λl1,l2 : lset A. l1 @ l2. definition lset_remove ≝ λA:DeqSet.λl:lset (carr A).λelt:carr A. (filter … (λx.¬x==elt) l). definition lset_difference ≝ λA:DeqSet.λl1,l2:lset (carr A). (filter … (λx.¬ (memb ? x l2)) l1). (* Standard predicates on sets *) definition lset_in ≝ λA:Type[0].λx : A. λl : lset A. mem … x l. definition lset_disjoint ≝ λA:Type[0].λl1, l2 : lset A. ∀x,y. mem … x l1 → mem … y l2 → x ≠ y. definition lset_inclusion ≝ λA:Type[0].λl1,l2. All A (λx1. mem … x1 l2) l1. (* Definition of abstract set equivalence. *) definition lset_eq ≝ λA:Type[0].λl1,l2. lset_inclusion A l1 l2 ∧ lset_inclusion A l2 l1. (* Properties of inclusion. *) lemma reflexive_lset_inclusion : ∀A. ∀l. lset_inclusion A l l. #A #l elim l try // #hd #tl #Hind whd @conj [ 1: %1 @refl | 2: whd in Hind; @(All_mp … Hind) #a #H whd %2 @H ] qed. lemma transitive_lset_inclusion : ∀A. ∀l1,l2,l3. lset_inclusion A l1 l2 → lset_inclusion A l2 l3 → lset_inclusion A l1 l3 . #A #l1 #l2 #l3 #Hincl1 #Hincl2 @(All_mp … Hincl1) whd in Hincl2; #a elim l2 in Hincl2; [ 1: normalize #_ @False_ind | 2: #hd #tl #Hind whd in ⊢ (% → ?); * #Hmem #Hincl_tl whd in ⊢ (% → ?); * [ 1: #Heq destruct @Hmem | 2: #Hmem_tl @Hind assumption ] ] qed. lemma cons_preserves_inclusion : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀x. lset_inclusion A l1 (x::l2). #A #l1 #l2 #Hincl #x @(All_mp … Hincl) #a #Hmem whd %2 @Hmem qed. lemma cons_monotonic_inclusion : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀x. lset_inclusion A (x::l1) (x::l2). #A #l1 #l2 #Hincl #x whd @conj [ 1: /2 by or_introl/ | 2: @(All_mp … Hincl) #a #Hmem whd %2 @Hmem ] qed. lemma lset_inclusion_concat : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀l3. lset_inclusion A l1 (l3 @ l2). #A #l1 #l2 #Hincl #l3 elim l3 try /2 by cons_preserves_inclusion/ qed. lemma lset_inclusion_concat_monotonic : ∀A. ∀l1,l2. lset_inclusion A l1 l2 → ∀l3. lset_inclusion A (l3 @ l1) (l3 @ l2). #A #l1 #l2 #Hincl #l3 elim l3 try @Hincl #hd #tl #Hind @cons_monotonic_inclusion @Hind qed. (* lset_eq is an equivalence relation. *) lemma reflexive_lset_eq : ∀A. ∀l. lset_eq A l l. /2 by conj/ qed. lemma transitive_lset_eq : ∀A. ∀l1,l2,l3. lset_eq A l1 l2 → lset_eq A l2 l3 → lset_eq A l1 l3. #A #l1 #l2 #l3 * #Hincl1 #Hincl2 * #Hincl3 #Hincl4 @conj @(transitive_lset_inclusion ??l2) assumption qed. lemma symmetric_lset_eq : ∀A. ∀l1,l2. lset_eq A l1 l2 → lset_eq A l2 l1. #A #l1 #l2 * #Hincl1 #Hincl2 @conj assumption qed. (* Properties of inclusion vs union and equality. *) lemma lset_union_inclusion : ∀A:Type[0]. ∀a,b,c. lset_inclusion A a c → lset_inclusion ? b c → lset_inclusion ? (lset_union ? a b) c. #A #a #b #c #H1 #H2 whd whd in match (lset_union ???); @All_append assumption qed. lemma lset_inclusion_union : ∀A:Type[0]. ∀a,b,c. lset_inclusion A a b ∨ lset_inclusion A a c → lset_inclusion ? a (lset_union ? b c). #A #a #b #c * [ 1: @All_mp #elt #Hmem @mem_append_backwards %1 @Hmem | 2: @All_mp #elt #Hmem @mem_append_backwards %2 @Hmem ] qed. lemma lset_inclusion_eq : ∀A : Type[0]. ∀a,b,c : lset A. lset_eq ? a b → lset_inclusion ? b c → lset_inclusion ? a c. #A #a #b #c * #H1 #H2 #H3 @(transitive_lset_inclusion … H1 H3) qed. lemma lset_inclusion_eq2 : ∀A : Type[0]. ∀a,b,c : lset A. lset_eq ? b c → lset_inclusion ? a b → lset_inclusion ? a c. #A #a #b #c * #H1 #H2 #H3 @(transitive_lset_inclusion … H3 H1) qed. (* Properties of lset_eq and mem *) lemma lset_eq_mem : ∀A:Type[0]. ∀s1,s2 : lset A. lset_eq ? s1 s2 → ∀b.mem ? b s1 → mem ? b s2. #A #s1 #s2 * #Hincl12 #_ #b whd in Hincl12; elim s1 in Hincl12; [ 1: normalize #_ * | 2: #hd #tl #Hind whd in ⊢ (% → % → ?); * #Hmem' #Hall * #Heq [ 1: destruct (Heq) assumption | 2: @Hind assumption ] ] qed. lemma lset_eq_memb : ∀A : DeqSet. ∀s1,s2 : lset (carr A). lset_eq ? s1 s2 → ∀b.memb ? b s1 = true → memb ? b s2 = true. #A #s1 #s2 #Heq #b lapply (memb_to_mem A s1 b) #H1 #H2 lapply (H1 H2) #Hmem lapply (lset_eq_mem … Heq ? Hmem) @mem_to_memb qed. lemma lset_eq_memb_eq : ∀A : DeqSet. ∀s1,s2 : lset (carr A). lset_eq ? s1 s2 → ∀b.memb ? b s1 = memb ? b s2. #A #s1 #s2 #Hlset_eq #b lapply (lset_eq_memb … Hlset_eq b) lapply (lset_eq_memb … (symmetric_lset_eq … Hlset_eq) b) cases (b∈s1) [ 1: #_ #H lapply (H (refl ??)) #Hmem >H @refl | 2: cases (b∈s2) // #H lapply (H (refl ??)) #Habsurd destruct ] qed. lemma lset_eq_filter_eq : ∀A : DeqSet. ∀s1,s2 : lset (carr A). lset_eq ? s1 s2 → ∀l. (filter ? (λwb.¬wb∈s1) l) = (filter ? (λwb.¬wb∈s2) l). #A #s1 #s2 #Heq #l elim l [ 1: @refl | 2: #hd #tl #Hind >filter_cons_unfold >filter_cons_unfold >(lset_eq_memb_eq … Heq) cases (hd∈s2) normalize in match (notb ?); normalize nodelta try @Hind >Hind @refl ] qed. lemma cons_monotonic_eq : ∀A. ∀l1,l2 : lset A. lset_eq A l1 l2 → ∀x. lset_eq A (x::l1) (x::l2). #A #l1 #l2 #Heq #x cases Heq #Hincl1 #Hincl2 @conj [ 1: @cons_monotonic_inclusion | 2: @cons_monotonic_inclusion ] assumption qed. (* Properties of difference and remove *) lemma lset_difference_unfold : ∀A : DeqSet. ∀s1, s2 : lset (carr A). ∀hd. lset_difference A (hd :: s1) s2 = if hd∈s2 then lset_difference A s1 s2 else hd :: (lset_difference A s1 s2). #A #s1 #s2 #hd normalize cases (hd∈s2) @refl qed. lemma lset_difference_unfold2 : ∀A : DeqSet. ∀s1, s2 : lset (carr A). ∀hd. lset_difference A s1 (hd :: s2) = lset_difference A (lset_remove ? s1 hd) s2. #A #s1 elim s1 [ 1: // | 2: #hd1 #tl1 #Hind #s2 #hd whd in match (lset_remove ???); whd in match (lset_difference A ??); whd in match (memb ???); lapply (eqb_true … hd1 hd) cases (hd1==hd) normalize nodelta [ 1: * #H #_ lapply (H (refl ??)) -H #H @Hind | 2: * #_ #Hguard >lset_difference_unfold cases (hd1∈s2) normalize in match (notb ?); normalize nodelta lset_difference_unfold2 #x #y whd in ⊢ (% → ?); * [ 2: @Hind | 1: #Heq >Heq elim s2 [ 1: normalize * | 2: #hd2 #tl2 #Hind2 whd in match (lset_remove ???); lapply (eqb_true … hd2 hd1) cases (hd2 == hd1) normalize in match (notb ?); normalize nodelta * #H1 #H2 [ 1: @Hind2 | 2: >lset_difference_unfold cases (hd2 ∈ tl1) normalize nodelta try @Hind2 whd in ⊢ (% → ?); * [ 1: #Hyhd2 destruct % #Heq lapply (H2 … (sym_eq … Heq)) #Habsurd destruct | 2: @Hind2 ] ] ] ] ] qed. 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). #A #l1 #l2 #elt /2 by filter_append/ qed. lemma lset_inclusion_remove : ∀A : DeqSet. ∀s1, s2 : lset A. lset_inclusion ? s1 s2 → ∀elt. lset_inclusion ? (lset_remove ? s1 elt) (lset_remove ? s2 elt). #A #s1 elim s1 [ 1: normalize // | 2: #hd1 #tl1 #Hind1 #s2 * #Hmem #Hincl elim (list_mem_split ??? Hmem) #s2A * #s2B #Hs2_eq destruct #elt whd in match (lset_remove ???); @(match (hd1 == elt) return λx. (hd1 == elt = x) → ? with [ true ⇒ λH. ? | false ⇒ λH. ? ] (refl ? (hd1 == elt))) normalize in match (notb ?); normalize nodelta [ 1: @Hind1 @Hincl | 2: whd @conj [ 2: @(Hind1 … Hincl) | 1: >lset_remove_split >lset_remove_split normalize in match (lset_remove A [hd1] elt); >H normalize nodelta @mem_append_backwards %2 @mem_append_backwards %1 normalize %1 @refl ] ] ] qed. lemma lset_difference_lset_eq : ∀A : DeqSet. ∀a,b,c. lset_eq A b c → lset_eq A (lset_difference A a b) (lset_difference A a c). #A #a #b #c #Heq whd in match (lset_difference ???) in ⊢ (??%%); elim a [ 1: normalize @conj @I | 2: #hda #tla #Hind whd in match (foldr ?????) in ⊢ (??%%); >(lset_eq_memb_eq … Heq hda) cases (hda∈c) normalize in match (notb ?); normalize nodelta try @Hind @cons_monotonic_eq @Hind ] qed. lemma lset_difference_lset_remove_commute : ∀A:DeqSet. ∀elt,s1,s2. (lset_difference A (lset_remove ? s1 elt) s2) = (lset_remove A (lset_difference ? s1 s2) elt). #A #elt #s1 #s2 elim s1 try // #hd #tl #Hind >lset_difference_unfold whd in match (lset_remove ???); @(match (hd==elt) return λx. (hd==elt) = x → ? with [ true ⇒ λHhd. ? | false ⇒ λHhd. ? ] (refl ? (hd==elt))) @(match (hd∈s2) return λx. (hd∈s2) = x → ? with [ true ⇒ λHmem. ? | false ⇒ λHmem. ? ] (refl ? (hd∈s2))) >notb_true >notb_false normalize nodelta try // try @Hind [ 1: whd in match (lset_remove ???) in ⊢ (???%); >Hhd elim (eqb_true ? elt elt) #_ #H >(H (refl ??)) normalize in match (notb ?); normalize nodelta @Hind | 2: >lset_difference_unfold >Hmem @Hind | 3: whd in match (lset_remove ???) in ⊢ (???%); >lset_difference_unfold >Hhd >Hmem normalize in match (notb ?); normalize nodelta >Hind @refl ] qed. (* Inversion lemma on emptyness *) lemma lset_eq_empty_inv : ∀A. ∀l. lset_eq A l [ ] → l = [ ]. #A #l elim l // #hd' #tl' normalize #Hind * * @False_ind qed. (* Inversion lemma on singletons *) lemma lset_eq_singleton_inv : ∀A,hd,l. lset_eq A [hd] (hd::l) → All ? (λx.x=hd) l. #A #hd #l * #Hincl1 whd in ⊢ (% → ?); * #_ @All_mp normalize #a * [ 1: #H @H | 2: @False_ind ] qed. (* Permutation of two elements on top of the list is ok. *) lemma lset_eq_permute : ∀A,l,x1,x2. lset_eq A (x1 :: x2 :: l) (x2 :: x1 :: l). #A #l #x1 #x2 @conj normalize [ 1: /5 by cons_preserves_inclusion, All_mp, or_introl, or_intror, conj/ | 2: /5 by cons_preserves_inclusion, All_mp, or_introl, or_intror, conj/ ] qed. (* "contraction" of an element. *) lemma lset_eq_contract : ∀A,l,x. lset_eq A (x :: x :: l) (x :: l). #A #l #x @conj [ 1: /5 by or_introl, conj, transitive_lset_inclusion/ | 2: /5 by cons_preserves_inclusion, cons_monotonic_inclusion/ ] qed. (* We don't need more than one instance of each element. *) lemma lset_eq_filter : ∀A:DeqSet.∀tl.∀hd. lset_eq A (hd :: (filter ? (λy.notb (eqb A y hd)) tl)) (hd :: tl). #A #tl elim tl [ 1: #hd normalize /4 by or_introl, conj, I/ | 2: #hd' #tl' #Hind #hd >filter_cons_unfold lapply (eqb_true A hd' hd) cases (hd'==hd) [ 2: #_ normalize in match (notb ?); normalize nodelta lapply (cons_monotonic_eq … (Hind hd) hd') #H lapply (lset_eq_permute ? tl' hd' hd) #H' @(transitive_lset_eq ? ? (hd'::hd::tl') ? … H') @(transitive_lset_eq ? ?? (hd'::hd::tl') … H) @lset_eq_permute | 1: * #Heq #_ >(Heq (refl ??)) normalize in match (notb ?); normalize nodelta lapply (Hind hd) #H @(transitive_lset_eq ? ? (hd::tl') (hd::hd::tl') H) @conj [ 1: whd @conj /2 by or_introl/ @cons_preserves_inclusion @cons_preserves_inclusion @reflexive_lset_inclusion | 2: whd @conj /2 by or_introl/ ] ] ] qed. lemma lset_inclusion_filter_self : ∀A:DeqSet.∀l,pred. lset_inclusion A (filter ? pred l) l. #A #l #pred elim l [ 1: normalize @I | 2: #hd #tl #Hind whd in match (filter ???); cases (pred hd) normalize nodelta [ 1: @cons_monotonic_inclusion @Hind | 2: @cons_preserves_inclusion @Hind ] ] qed. lemma lset_inclusion_filter_monotonic : ∀A:DeqSet.∀l1,l2. lset_inclusion (carr A) l1 l2 → ∀elt. lset_inclusion A (filter ? (λx.¬(x==elt)) l1) (filter ? (λx.¬(x==elt)) l2). #A #l1 elim l1 [ 1: #l2 normalize // | 2: #hd1 #tl1 #Hind #l2 whd in ⊢ (% → ?); * #Hmem1 #Htl1_incl #elt whd >filter_cons_unfold lapply (eqb_true A hd1 elt) cases (hd1==elt) [ 1: * #H1 #_ lapply (H1 (refl ??)) #H1_eq >H1_eq in Hmem1; #Hmem normalize in match (notb ?); normalize nodelta @Hind assumption | 2: * #_ #Hneq normalize in match (notb ?); normalize nodelta whd @conj [ 1: elim l2 in Hmem1; try // #hd2 #tl2 #Hincl whd in ⊢ (% → ?); * [ 1: #Heq >Heq in Hneq; normalize lapply (eqb_true A hd2 elt) cases (hd2==elt) [ 1: * #H #_ #H2 lapply (H2 (H (refl ??))) #Habsurd destruct (Habsurd) | 2: #_ normalize nodelta #_ /2 by or_introl/ ] | 2: #H lapply (Hincl H) #Hok normalize cases (hd2==elt) normalize nodelta [ 1: @Hok | 2: %2 @Hok ] ] | 2: @Hind assumption ] ] ] qed. (* removing an element of two equivalent sets conserves equivalence. *) lemma lset_eq_filter_monotonic : ∀A:DeqSet.∀l1,l2. lset_eq (carr A) l1 l2 → ∀elt. lset_eq A (filter ? (λx.¬(x==elt)) l1) (filter ? (λx.¬(x==elt)) l2). #A #l1 #l2 * #Hincl1 #Hincl2 #elt @conj /2 by lset_inclusion_filter_monotonic/ qed. (* ---------------- Concrete implementation of sets --------------------- *) (* We can give an explicit characterization of equivalent sets: they are permutations with repetitions, i,e. a composition of transpositions and duplications. We restrict ourselves to DeqSets. *) inductive iso_step_lset (A : DeqSet) : lset A → lset A → Prop ≝ | lset_transpose : ∀a,x,b,y,c. iso_step_lset A (a @ [x] @ b @ [y] @ c) (a @ [y] @ b @ [x] @ c) | lset_weaken : ∀a,x,b. iso_step_lset A (a @ [x] @ b) (a @ [x] @ [x] @ b) | lset_contract : ∀a,x,b. iso_step_lset A (a @ [x] @ [x] @ b) (a @ [x] @ b). (* The equivalence is the reflexive, transitive and symmetric closure of iso_step_lset *) inductive lset_eq_concrete (A : DeqSet) : lset A → lset A → Prop ≝ | lset_trans : ∀a,b,c. iso_step_lset A a b → lset_eq_concrete A b c → lset_eq_concrete A a c | lset_refl : ∀a. lset_eq_concrete A a a. (* lset_eq_concrete is symmetric and transitive *) 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. #A #l1 #l2 #l3 #Hequiv elim Hequiv // #a #b #c #Hstep #Hequiv #Hind #Hcl3 lapply (Hind Hcl3) #Hbl3 @(lset_trans ???? Hstep Hbl3) qed. lemma symmetric_step : ∀A,l1,l2. iso_step_lset A l1 l2 → iso_step_lset A l2 l1. #A #l1 #l2 * /2/ qed. lemma symmetric_lset_eq_concrete : ∀A,l1,l2. lset_eq_concrete A l1 l2 → lset_eq_concrete A l2 l1. #A #l1 #l2 #H elim H // #a #b #c #Hab #Hbc #Hcb @(transitive_lset_eq_concrete ???? Hcb ?) @(lset_trans … (symmetric_step ??? Hab) (lset_refl …)) qed. (* lset_eq_concrete is conserved by cons. *) lemma equivalent_step_cons : ∀A,l1,l2. iso_step_lset A l1 l2 → ∀x. iso_step_lset A (x :: l1) (x :: l2). #A #l1 #l2 * // qed. (* That // was impressive. *) lemma lset_eq_concrete_cons : ∀A,l1,l2. lset_eq_concrete A l1 l2 → ∀x. lset_eq_concrete A (x :: l1) (x :: l2). #A #l1 #l2 #Hequiv elim Hequiv try // #a #b #c #Hab #Hbc #Hind #x %1{(equivalent_step_cons ??? Hab x) (Hind x)} qed. lemma absurd_list_eq_append : ∀A.∀x.∀l1,l2:list A. [ ] = l1 @ [x] @ l2 → False. #A #x #l1 #l2 elim l1 normalize [ 1: #Habsurd destruct | 2: #hd #tl #_ #Habsurd destruct ] qed. (* Inversion lemma for emptyness, step case *) lemma equivalent_iso_step_empty_inv : ∀A,l. iso_step_lset A l [] → l = [ ]. #A #l elim l // #hd #tl #Hind #H inversion H [ 1: #a #x #b #y #c #_ #Habsurd @(False_ind … (absurd_list_eq_append ? y … Habsurd)) | 2: #a #x #b #_ #Habsurd @(False_ind … (absurd_list_eq_append ? x … Habsurd)) | 3: #a #x #b #_ #Habsurd @(False_ind … (absurd_list_eq_append ? x … Habsurd)) ] qed. (* Same thing for non-emptyness *) lemma equivalent_iso_step_cons_inv : ∀A,l1,l2. l1 ≠ [ ] → iso_step_lset A l1 l2 → l2 ≠ [ ]. #A #l1 elim l1 [ 1: #l2 * #H @(False_ind … (H (refl ??))) | 2: #hd #tl #H #l2 #_ #Hstep % #Hl2 >Hl2 in Hstep; #Hstep lapply (equivalent_iso_step_empty_inv … Hstep) #Habsurd destruct ] qed. lemma lset_eq_concrete_cons_inv : ∀A,l1,l2. l1 ≠ [ ] → lset_eq_concrete A l1 l2 → l2 ≠ [ ]. #A #l1 #l2 #Hl1 #H lapply Hl1 elim H [ 2: #a #H @H | 1: #a #b #c #Hab #Hbc #H #Ha lapply (equivalent_iso_step_cons_inv … Ha Hab) #Hb @H @Hb ] qed. lemma lset_eq_concrete_empty_inv : ∀A,l1,l2. l1 = [ ] → lset_eq_concrete A l1 l2 → l2 = [ ]. #A #l1 #l2 #Hl1 #H lapply Hl1 elim H // #a #b #c #Hab #Hbc #Hbc_eq #Ha >Ha in Hab; #H_b lapply (equivalent_iso_step_empty_inv … ?? (symmetric_step … H_b)) #Hb @Hbc_eq @Hb qed. (* Square equivalence diagram *) lemma square_lset_eq_concrete : ∀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'. #A #a #b #a' #b' #H1 #H2 #H3 @(transitive_lset_eq_concrete ???? (symmetric_lset_eq_concrete … H2)) @(transitive_lset_eq_concrete ???? H1) @H3 qed. (* Make the transposition of elements visible at top-level *) lemma transpose_lset_eq_concrete : ∀A. ∀x,y,a,b,c,a',b',c'. lset_eq_concrete A (a @ [x] @ b @ [y] @ c) (a' @ [x] @ b' @ [y] @ c') → lset_eq_concrete A (a @ [y] @ b @ [x] @ c) (a' @ [y] @ b' @ [x] @ c'). #A #x #y #a #b #c #a' #b' #c #H @(square_lset_eq_concrete ????? H) /2 by lset_trans, lset_refl, lset_transpose/ qed. lemma switch_lset_eq_concrete : ∀A. ∀a,b,c. lset_eq_concrete A (a@[b]@c) ([b]@a@c). #A #a elim a // #hda #tla #Hind #b #c lapply (Hind hda c) #H lapply (lset_eq_concrete_cons … H b) #H' normalize in H' ⊢ %; @symmetric_lset_eq_concrete /3 by lset_transpose, lset_trans, symmetric_lset_eq_concrete/ qed. (* Folding a commutative and idempotent function on equivalent sets yields the same result. *) lemma lset_eq_concrete_fold : ∀A : DeqSet. ∀acctype : Type[0]. ∀l1,l2 : list (carr A). lset_eq_concrete A l1 l2 → ∀f:carr A → acctype → acctype. (∀x1,x2. ∀acc. f x1 (f x2 acc) = f x2 (f x1 acc)) → (∀x.∀acc. f x (f x acc) = f x acc) → ∀acc. foldr ?? f acc l1 = foldr ?? f acc l2. #A #acctype #l1 #l2 #Heq #f #Hcomm #Hidem elim Heq try // #a' #b' #c' #Hstep #Hbc #H #acc fold_append >fold_append >fold_append >fold_append >fold_append >fold_append >fold_append >fold_append normalize cut (f x (foldr A acctype f (f y (foldr A acctype f acc c)) b) = f y (foldr A acctype f (f x (foldr A acctype f acc c)) b)) [ elim c [ 1: normalize elim b [ 1: normalize >(Hcomm x y) @refl | 2: #hdb #tlb #Hind normalize >(Hcomm x hdb) >(Hcomm y hdb) >Hind @refl ] | 2: #hdc #tlc #Hind normalize elim b [ 1: normalize >(Hcomm x y) @refl | 2: #hdb #tlb #Hind normalize >(Hcomm x hdb) >(Hcomm y hdb) >Hind @refl ] ] ] #Hind >Hind @refl | 2: #a #x #b >fold_append >fold_append >fold_append >(fold_append ?? ([x]@[x])) normalize >Hidem @refl | 3: #a #x #b >fold_append >(fold_append ?? ([x]@[x])) >fold_append >fold_append normalize >Hidem @refl ] qed. (* Folding over equivalent sets yields equivalent results, for any equivalence. *) lemma inj_to_fold_inj : ∀A,acctype : Type[0]. ∀eqrel : acctype → acctype → Prop. ∀refl_eqrel : reflexive ? eqrel. ∀trans_eqrel : transitive ? eqrel. ∀sym_eqrel : symmetric ? eqrel. ∀f : A → acctype → acctype. (∀x:A.∀acc1:acctype.∀acc2:acctype.eqrel acc1 acc2→eqrel (f x acc1) (f x acc2)) → ∀l : list A. ∀acc1, acc2 : acctype. eqrel acc1 acc2 → eqrel (foldr … f acc1 l) (foldr … f acc2 l). #A #acctype #eqrel #R #T #S #f #Hinj #l elim l // #hd #tl #Hind #acc1 #acc2 #Heq normalize @Hinj @Hind @Heq qed. (* We need to extend the above proof to arbitrary equivalence relation instead of just standard equality. *) lemma lset_eq_concrete_fold_ext : ∀A : DeqSet. ∀acctype : Type[0]. ∀eqrel : acctype → acctype → Prop. ∀refl_eqrel : reflexive ? eqrel. ∀trans_eqrel : transitive ? eqrel. ∀sym_eqrel : symmetric ? eqrel. ∀f:carr A → acctype → acctype. (∀x,acc1,acc2. eqrel acc1 acc2 → eqrel (f x acc1) (f x acc2)) → (∀x1,x2. ∀acc. eqrel (f x1 (f x2 acc)) (f x2 (f x1 acc))) → (∀x.∀acc. eqrel (f x (f x acc)) (f x acc)) → ∀l1,l2 : list (carr A). lset_eq_concrete A l1 l2 → ∀acc. eqrel (foldr ?? f acc l1) (foldr ?? f acc l2). #A #acctype #eqrel #R #T #S #f #Hinj #Hcomm #Hidem #l1 #l2 #Heq elim Heq try // #a' #b' #c' #Hstep #Hbc #H inversion Hstep [ 1: #a #x #b #y #c #HlA #HlB #_ #acc >HlB in H; #H @(T … ? (H acc)) >fold_append >fold_append >fold_append >fold_append >fold_append >fold_append >fold_append >fold_append normalize cut (eqrel (f x (foldr ? acctype f (f y (foldr ? acctype f acc c)) b)) (f y (foldr ? acctype f (f x (foldr ? acctype f acc c)) b))) [ 1: elim c [ 1: normalize elim b [ 1: normalize @(Hcomm x y) | 2: #hdb #tlb #Hind normalize lapply (Hinj hdb ?? Hind) #Hind' lapply (T … Hind' (Hcomm ???)) #Hind'' @S @(triangle_diagram ? eqrel R T S … Hind'') @Hcomm ] | 2: #hdc #tlc #Hind normalize elim b [ 1: normalize @(Hcomm x y) | 2: #hdb #tlb #Hind normalize lapply (Hinj hdb ?? Hind) #Hind' lapply (T … Hind' (Hcomm ???)) #Hind'' @S @(triangle_diagram ? eqrel R T S … Hind'') @Hcomm ] ] ] #Hind @(inj_to_fold_inj … eqrel R T S ? Hinj … Hind) | 2: #a #x #b #HeqA #HeqB #_ #acc >HeqB in H; #H @(T … (H acc)) >fold_append >fold_append >fold_append >(fold_append ?? ([x]@[x])) normalize @(inj_to_fold_inj … eqrel R T S ? Hinj) @S @Hidem | 3: #a #x #b #HeqA #HeqB #_ #acc >HeqB in H; #H @(T … (H acc)) >fold_append >(fold_append ?? ([x]@[x])) >fold_append >fold_append normalize @(inj_to_fold_inj … eqrel R T S ? Hinj) @Hidem ] qed. (* Prepare some well-founded induction principles on lists. The idea is to perform an induction on the sequence of filterees of a list : taking the first element, filtering it out of the tail, etc. We give such principles for pairs of lists and isolated lists. *) (* The two lists [l1,l2] share at least the head of l1. *) definition head_shared ≝ λA. λl1,l2 : list A. match l1 with [ nil ⇒ l2 = (nil ?) | cons hd _ ⇒ mem … hd l2 ]. (* Relation on pairs of lists, as described above. *) definition filtered_lists : ∀A:DeqSet. relation (list (carr A)×(list (carr A))) ≝ λA:DeqSet. λll1,ll2. let 〈la1,lb1〉 ≝ ll1 in let 〈la2,lb2〉 ≝ ll2 in match la2 with [ nil ⇒ False | cons hda2 tla2 ⇒ head_shared ? la2 lb2 ∧ la1 = filter … (λx.¬(x==hda2)) tla2 ∧ lb1 = filter … (λx.¬(x==hda2)) lb2 ]. (* Notice the absence of plural : this relation works on a simple list, not a pair. *) definition filtered_list : ∀A:DeqSet. relation (list (carr A)) ≝ λA:DeqSet. λl1,l2. match l2 with [ nil ⇒ False | cons hd2 tl2 ⇒ l1 = filter … (λx.¬(x==hd2)) l2 ]. (* Relation on lists based on their lengths. We know this one is well-founded. *) definition length_lt : ∀A:DeqSet. relation (list (carr A)) ≝ λA:DeqSet. λl1,l2. lt (length ? l1) (length ? l2). (* length_lt can be extended on pairs by just measuring the first component *) definition length_fst_lt : ∀A:DeqSet. relation (list (carr A) × (list (carr A))) ≝ λA:DeqSet. λll1,ll2. lt (length ? (\fst ll1)) (length ? (\fst ll2)). lemma filter_length : ∀A. ∀l. ∀f. |filter A f l| ≤ |l|. #A #l #f elim l // #hd #tl #Hind whd in match (filter ???); cases (f hd) normalize nodelta [ 1: /2 by le_S_S/ | 2: @le_S @Hind ] qed. (* The order on lists defined by their length is wf *) lemma length_lt_wf : ∀A. ∀l. WF (list (carr A)) (length_lt A) l. #A #l % elim l [ 1: #a normalize in ⊢ (% → ?); #H lapply (le_S_O_absurd ? H) @False_ind | 2: #hd #tl #Hind #a #Hlt % #a' #Hlt' @Hind whd in Hlt Hlt'; whd @(transitive_le … Hlt') @(monotonic_pred … Hlt) qed. (* Order on pairs of list by measuring the first proj *) lemma length_fst_lt_wf : ∀A. ∀ll. WF ? (length_fst_lt A) ll. #A * #l1 #l2 % elim l1 [ 1: * #a1 #a2 normalize in ⊢ (% → ?); #H lapply (le_S_O_absurd ? H) @False_ind | 2: #hd1 #tl1 #Hind #a #Hlt % #a' #Hlt' @Hind whd in Hlt Hlt'; whd @(transitive_le … Hlt') @(monotonic_pred … Hlt) qed. lemma length_to_filtered_lists : ∀A. subR ? (filtered_lists A) (length_fst_lt A). #A whd * #a1 #a2 * #b1 #b2 elim b1 [ 1: @False_ind | 2: #hd1 #tl1 #Hind whd in ⊢ (% → ?); * * #Hmem #Ha1 #Ha2 whd >Ha1 whd in match (length ??) in ⊢ (??%); @le_S_S @filter_length ] qed. lemma length_to_filtered_list : ∀A. subR ? (filtered_list A) (length_lt A). #A whd #a #b elim b [ 1: @False_ind | 2: #hd #tl #Hind whd in ⊢ (% → ?); whd in match (filter ???); lapply (eqb_true ? hd hd) * #_ #H >(H (refl ??)) normalize in match (notb ?); normalize nodelta #Ha whd @le_S_S >Ha @filter_length ] qed. (* Prove well-foundedness by embedding in lt *) lemma filtered_lists_wf : ∀A. ∀ll. WF ? (filtered_lists A) ll. #A #ll @(WF_antimonotonic … (length_to_filtered_lists A)) @length_fst_lt_wf qed. lemma filtered_list_wf : ∀A. ∀l. WF ? (filtered_list A) l. #A #l @(WF_antimonotonic … (length_to_filtered_list A)) @length_lt_wf qed. definition Acc_elim : ∀A,R,x. WF A R x → (∀a. R a x → WF A R a) ≝ λA,R,x,acc. match acc with [ wf _ a0 ⇒ a0 ]. (* Stolen from Coq. Warped to avoid prop-to-type restriction. *) let rec WF_rect (A : Type[0]) (R : A → A → Prop) (P : A → Type[0]) (f : ∀ x : A. (∀ y : A. R y x → WF ? R y) → (∀ y : A. R y x → P y) → P x) (x : A) (a : WF A R x) on a : P x ≝ f x (Acc_elim … a) (λy:A. λRel:R y x. WF_rect A R P f y ((Acc_elim … a) y Rel)). lemma lset_eq_concrete_filter : ∀A:DeqSet.∀tl.∀hd. lset_eq_concrete A (hd :: (filter ? (λy.notb (eqb A y hd)) tl)) (hd :: tl). #A #tl elim tl [ 1: #hd // | 2: #hd' #tl' #Hind #hd >filter_cons_unfold lapply (eqb_true A hd' hd) cases (hd'==hd) [ 2: #_ normalize in match (notb false); normalize nodelta >cons_to_append >(cons_to_append … hd') >cons_to_append in ⊢ (???%); >(cons_to_append … hd') in ⊢ (???%); @(transpose_lset_eq_concrete ? hd' hd [ ] [ ] (filter A (λy:A.¬y==hd) tl') [ ] [ ] tl') >nil_append >nil_append >nil_append >nil_append @lset_eq_concrete_cons >nil_append >nil_append @Hind | 1: * #H1 #_ lapply (H1 (refl ??)) #Heq normalize in match (notb ?); normalize nodelta >Heq >cons_to_append >cons_to_append in ⊢ (???%); >cons_to_append in ⊢ (???(???%)); @(square_lset_eq_concrete A ([hd]@filter A (λy:A.¬y==hd) tl') ([hd]@tl')) [ 1: @Hind | 2: %2 | 3: %1{???? ? (lset_refl ??)} /2 by lset_weaken/ ] ] ] qed. (* The "abstract", observational definition of set equivalence implies the concrete one. *) lemma lset_eq_to_lset_eq_concrete_aux : ∀A,ll. head_shared … (\fst ll) (\snd ll) → lset_eq (carr A) (\fst ll) (\snd ll) → lset_eq_concrete A (\fst ll) (\snd ll). #A #ll @(WF_ind ????? (filtered_lists_wf A ll)) * * [ 1: #b2 #Hwf #Hind_wf whd in ⊢ (% → ?); #Hb2 >Hb2 #_ %2 | 2: #hd1 #tl1 #b2 #Hwf #Hind_wf whd in ⊢ (% → ?); #Hmem lapply (list_mem_split ??? Hmem) * #l2A * #l2B #Hb2 #Heq destruct lapply (Hind_wf 〈filter … (λx.¬x==hd1) tl1,filter … (λx.¬x==hd1) (l2A@l2B)〉) cut (filtered_lists A 〈filter A (λx:A.¬x==hd1) tl1,filter A (λx:A.¬x==hd1) (l2A@l2B)〉 〈hd1::tl1,l2A@[hd1]@l2B〉) [ @conj try @conj try @refl whd [ 1: /2 by / | 2: >filter_append in ⊢ (???%); >filter_append in ⊢ (???%); whd in match (filter ?? [hd1]); elim (eqb_true A hd1 hd1) #_ #H >(H (refl ??)) normalize in match (notb ?); normalize nodelta filter_cons_unfold >filter_append >(filter_append … [hd1]) whd in match (filter ?? [hd1]); elim (eqb_true A hd1 hd1) #_ #H >(H (refl ??)) normalize in match (notb ?); normalize nodelta cons_to_append @(transitive_lset_eq_concrete ? ([hd1]@tl1) ([hd1]@l2A@l2B) (l2A@[hd1]@l2B)) [ 2: @symmetric_lset_eq_concrete @switch_lset_eq_concrete ] lapply (lset_eq_concrete_cons … Hconcrete_set_eq hd1) #Hconcrete_cons_eq @(square_lset_eq_concrete … Hconcrete_cons_eq) [ 1: @(lset_eq_concrete_filter ? tl1 hd1) | 2: @(lset_eq_concrete_filter ? (l2A@l2B) hd1) ] ] qed. lemma lset_eq_to_lset_eq_concrete : ∀A,l1,l2. lset_eq (carr A) l1 l2 → lset_eq_concrete A l1 l2. #A * [ 1: #l2 #Heq >(lset_eq_empty_inv … (symmetric_lset_eq … Heq)) // | 2: #hd1 #tl1 #l2 #Hincl lapply Hincl lapply (lset_eq_to_lset_eq_concrete_aux ? 〈hd1::tl1,l2〉) #H @H whd elim Hincl * // ] qed. (* The concrete one implies the abstract one. *) lemma lset_eq_concrete_to_lset_eq : ∀A,l1,l2. lset_eq_concrete A l1 l2 → lset_eq A l1 l2. #A #l1 #l2 #Hconcrete elim Hconcrete try // #a #b #c #Hstep #Heq_bc_concrete #Heq_bc cut (lset_eq A a b) [ 1: elim Hstep [ 1: #a' elim a' [ 2: #hda #tla #Hind #x #b' #y #c' >cons_to_append >(associative_append ? [hda] tla ?) >(associative_append ? [hda] tla ?) @cons_monotonic_eq >nil_append >nil_append @Hind | 1: #x #b' #y #c' >nil_append >nil_append elim b' try // #hdb #tlc #Hind >append_cons >append_cons in ⊢ (???%); >associative_append >associative_append lapply (cons_monotonic_eq … Hind hdb) #Hind' @(transitive_lset_eq ? ([x]@[hdb]@tlc@[y]@c') ([hdb]@[x]@tlc@[y]@c')) /2 by transitive_lset_eq/ ] | 2: #a' elim a' [ 2: #hda #tla #Hind #x #b' >cons_to_append >(associative_append ? [hda] tla ?) >(associative_append ? [hda] tla ?) @cons_monotonic_eq >nil_append >nil_append @Hind | 1: #x #b' >nil_append >nil_append @conj normalize [ 1: @conj [ 1: %1 @refl ] elim b' [ 1: @I | 2: #hdb #tlb #Hind normalize @conj [ 1: %2 %2 %1 @refl | 2: @(All_mp … Hind) #a0 * [ 1: #Heq %1 @Heq | 2: * /2 by or_introl, or_intror/ ] ] ] #H %2 %2 %2 @H | 2: @conj try @conj try /2 by or_introl, or_intror/ elim b' [ 1: @I | 2: #hdb #tlb #Hind normalize @conj [ 1: %2 %1 @refl | 2: @(All_mp … Hind) #a0 * [ 1: #Heq %1 @Heq | 2: #H %2 %2 @H ] ] ] ] ] | 3: #a #x #b elim a try @lset_eq_contract #hda #tla #Hind @cons_monotonic_eq @Hind ] ] #Heq_ab @(transitive_lset_eq … Heq_ab Heq_bc) qed. lemma lset_eq_fold : ∀A : DeqSet. ∀acctype : Type[0]. ∀eqrel : acctype → acctype → Prop. ∀refl_eqrel : reflexive ? eqrel. ∀trans_eqrel : transitive ? eqrel. ∀sym_eqrel : symmetric ? eqrel. ∀f:carr A → acctype → acctype. (∀x,acc1,acc2. eqrel acc1 acc2 → eqrel (f x acc1) (f x acc2)) → (∀x1,x2. ∀acc. eqrel (f x1 (f x2 acc)) (f x2 (f x1 acc))) → (∀x.∀acc. eqrel (f x (f x acc)) (f x acc)) → ∀l1,l2 : list (carr A). lset_eq A l1 l2 → ∀acc. eqrel (foldr ?? f acc l1) (foldr ?? f acc l2). #A #acctype #eqrel #refl_eqrel #trans_eqrel #sym_eqrel #f #Hinj #Hsym #Hcontract #l1 #l2 #Heq #acc lapply (lset_eq_to_lset_eq_concrete … Heq) #Heq_concrete @(lset_eq_concrete_fold_ext A acctype eqrel refl_eqrel trans_eqrel sym_eqrel f Hinj Hsym Hcontract l1 l2 Heq_concrete acc) qed. (* Additional lemmas on lsets *) lemma lset_difference_empty : ∀A : DeqSet. ∀s1. lset_difference A s1 [ ] = s1. #A #s1 elim s1 try // #hd #tl #Hind >lset_difference_unfold >Hind @refl qed. lemma lset_not_mem_difference : ∀A : DeqSet. ∀s1,s2,s3. lset_inclusion (carr A) s1 (lset_difference ? s2 s3) → ∀x. mem ? x s1 → ¬(mem ? x s3). #A #s1 #s2 #s3 #Hincl #x #Hmem lapply (lset_difference_disjoint ? s3 s2) whd in ⊢ (% → ?); #Hdisjoint % #Hmem_s3 elim s1 in Hincl Hmem; [ 1: #_ * | 2: #hd #tl #Hind whd in ⊢ (% → %); * #Hmem_hd #Hall * [ 2: #Hmem_x_tl @Hind assumption | 1: #Heq lapply (Hdisjoint … Hmem_s3 Hmem_hd) * #H @H @Heq ] ] qed. lemma lset_mem_inclusion_mem : ∀A,s1,s2,elt. mem A elt s1 → lset_inclusion ? s1 s2 → mem ? elt s2. #A #s1 elim s1 [ 1: #s2 #elt * | 2: #hd #tl #Hind #s2 #elt * [ 1: #Heq destruct * // | 2: #Hmem_tl * #Hmem #Hall elim tl in Hall Hmem_tl; [ 1: #_ * | 2: #hd' #tl' #Hind * #Hmem' #Hall * [ 1: #Heq destruct @Hmem' | 2: #Hmem'' @Hind assumption ] ] ] ] qed. lemma lset_remove_inclusion : ∀A : DeqSet. ∀s,elt. lset_inclusion A (lset_remove ? s elt) s. #A #s elim s try // qed. lemma lset_difference_remove_inclusion : ∀A : DeqSet. ∀s1,s2,elt. lset_inclusion A (lset_difference ? (lset_remove ? s1 elt) s2) (lset_difference ? s1 s2). #A #s elim s try // qed. lemma lset_difference_permute : ∀A : DeqSet. ∀s1,s2,s3. lset_difference A s1 (s2 @ s3) = lset_difference A s1 (s3 @ s2). #A #s1 #s2 elim s2 try // #hd #tl #Hind #s3 >lset_difference_unfold2 >lset_difference_lset_remove_commute >Hind elim s3 try // #hd' #tl' #Hind' >cons_to_append >associative_append >associative_append >(cons_to_append … hd tl) >lset_difference_unfold2 >lset_difference_lset_remove_commute >nil_append >lset_difference_unfold2 >lset_difference_lset_remove_commute >nil_append H normalize nodelta @Hind | 2: normalize >H' normalize nodelta @Hind | 3: normalize >H >H' normalize nodelta >Hind @refl ] qed. lemma lset_disjoint_dec : ∀A : DeqSet. ∀s1,elt,s2. mem ? elt s1 ∨ mem ? elt (lset_difference A (elt :: s2) s1). #A #s1 #elt #s2 @(match elt ∈ s1 return λx. ((elt ∈ s1) = x) → ? with [ false ⇒ λHA. ? | true ⇒ λHA. ? ] (refl ? (elt ∈ s1))) [ 1: lapply (memb_to_mem … HA) #Hmem %1 @Hmem | 2: %2 elim s1 in HA; [ 1: #_ whd %1 @refl | 2: #hd1 #tl1 #Hind normalize in ⊢ (% → ?); >lset_difference_unfold >lset_difference_unfold2 lapply (eqb_true ? elt hd1) whd in match (memb ???) in ⊢ (? → ? → %); cases (elt==hd1) normalize nodelta [ 1: #_ #Habsurd destruct | 2: #HA #HB >HB normalize nodelta %1 @refl ] ] ] qed. lemma mem_filter : ∀A : DeqSet. ∀l,elt1,elt2. mem A elt1 (filter A (λx:A.¬x==elt2) l) → mem A elt1 l. #A #l elim l try // #hd #tl #Hind #elt1 #elt2 /2 by lset_mem_inclusion_mem/ qed. lemma lset_inclusion_difference_aux : ∀A : DeqSet. ∀s1,s2. lset_inclusion A s1 s2 → (lset_eq A s2 (s1@lset_difference A s2 s1)). #A #s1 @(WF_ind ????? (filtered_list_wf A s1)) * [ 1: #_ #_ #s2 #_ >nil_append >lset_difference_empty @reflexive_lset_eq | 2: #hd1 #tl1 #Hwf #Hind #s2 * #Hmem #Hincl lapply (Hind (filter ? (λx.¬x==hd1) tl1) ?) [ 1: whd normalize >(proj2 … (eqb_true ? hd1 hd1) (refl ??)) normalize nodelta @refl ] #Hind_wf elim (list_mem_split ??? Hmem) #s2A * #s2B #Heq >Heq >cons_to_append in ⊢ (???%); >associative_append >lset_difference_unfold2 >nil_append >lset_remove_split >lset_remove_split normalize in match (lset_remove ? [hd1] hd1); >(proj2 … (eqb_true ? hd1 hd1) (refl ??)) normalize nodelta >nil_append lset_difference_lset_remove_commute lapply (Hind_wf (lset_remove A (s2A@s2B) hd1) ?) [ 1: lapply (lset_inclusion_remove … Hincl hd1) >Heq @lset_inclusion_eq2 >lset_remove_split >lset_remove_split >lset_remove_split normalize in match (lset_remove ? [hd1] hd1); >(proj2 … (eqb_true ? hd1 hd1) (refl ??)) normalize nodelta >nil_append @reflexive_lset_eq ] #Hind >lset_difference_lset_remove_commute in Hind; lset_difference_unfold >lset_difference_unfold @(match (hd2∈filter A (λx:A.¬x==hd1) tl1) return λx. ((hd2∈filter A (λx:A.¬x==hd1) tl1) = x) → ? with [ false ⇒ λH. ? | true ⇒ λH. ? ] (refl ? (hd2∈filter A (λx:A.¬x==hd1) tl1))) normalize nodelta [ 1: lapply (memb_to_mem … H) #Hfilter >(mem_to_memb … (mem_filter … Hfilter)) normalize nodelta @Hind | 2: @(match (hd2∈tl1) return λx. ((hd2∈tl1) = x) → ? with [ false ⇒ λH'. ? | true ⇒ λH'. ? ] (refl ? (hd2∈tl1))) normalize nodelta [ 1: (* We have hd2 = hd1 *) cut (hd2 = hd1) [ elim tl1 in H H'; [ 1: normalize #_ #Habsurd destruct (Habsurd) | 2: #hdtl1 #tltl1 #Hind normalize in ⊢ (% → % → ?); lapply (eqb_true ? hdtl1 hd1) cases (hdtl1==hd1) normalize nodelta [ 1: * #H >(H (refl ??)) #_ lapply (eqb_true ? hd2 hd1) cases (hd2==hd1) normalize nodelta * [ 1: #H' >(H' (refl ??)) #_ #_ #_ @refl | 2: #_ #_ @Hind ] | 2: #_ normalize lapply (eqb_true ? hd2 hdtl1) cases (hd2 == hdtl1) normalize nodelta * [ 1: #_ #_ #Habsurd destruct (Habsurd) | 2: #_ #_ @Hind ] ] ] ] #Heq_hd2hd1 destruct (Heq_hd2hd1) @lset_eq_concrete_to_lset_eq lapply (lset_eq_to_lset_eq_concrete … Hind) #Hind' @(square_lset_eq_concrete … Hind') [ 2: @lset_refl | 1: >cons_to_append >cons_to_append in ⊢ (???%); @(transitive_lset_eq_concrete … ([hd1]@[hd1]@tl1@lset_difference A tl2 (filter A (λx:A.¬x==hd1) tl1))) [ 1: @lset_eq_to_lset_eq_concrete @symmetric_lset_eq @lset_eq_contract | 2: >(cons_to_append … hd1 (lset_difference ???)) @lset_eq_concrete_cons >nil_append >nil_append @symmetric_lset_eq_concrete @switch_lset_eq_concrete ] ] | 2: @(match hd2 == hd1 return λx. ((hd2 == hd1) = x) → ? with [ true ⇒ λH''. ? | false ⇒ λH''. ? ] (refl ? (hd2 == hd1))) [ 1: whd in match (lset_remove ???) in ⊢ (???%); >H'' normalize nodelta >((proj1 … (eqb_true …)) H'') @(transitive_lset_eq … Hind) @(transitive_lset_eq … (hd1::hd1::tl1@lset_difference A tl2 (filter A (λx:A.¬x==hd1) tl1))) [ 2: @lset_eq_contract ] @lset_eq_concrete_to_lset_eq @lset_eq_concrete_cons @switch_lset_eq_concrete | 2: whd in match (lset_remove ???) in ⊢ (???%); >H'' >notb_false normalize nodelta @lset_eq_concrete_to_lset_eq lapply (lset_eq_to_lset_eq_concrete … Hind) #Hindc lapply (lset_eq_concrete_cons … Hindc hd2) #Hindc' -Hindc @(square_lset_eq_concrete … Hindc') [ 1: @symmetric_lset_eq_concrete >cons_to_append >cons_to_append in ⊢ (???%); >(cons_to_append … hd2) >(cons_to_append … hd1) in ⊢ (???%); @(switch_lset_eq_concrete ? ([hd1]@tl1) hd2 ?) | 2: @symmetric_lset_eq_concrete @(switch_lset_eq_concrete ? ([hd1]@tl1) hd2 ?) ] ] ] ] ] ] ] qed. lemma lset_inclusion_difference : ∀A : DeqSet. ∀s1,s2 : lset (carr A). lset_inclusion ? s1 s2 → ∃s2'. lset_eq ? s2 (s1 @ s2') ∧ lset_disjoint ? s1 s2' ∧ lset_eq ? s2' (lset_difference ? s2 s1). #A #s1 #s2 #Hincl %{(lset_difference A s2 s1)} @conj try @conj [ 1: @lset_inclusion_difference_aux @Hincl | 2: /2 by lset_difference_disjoint/ | 3,4: @reflexive_lset_inclusion ] qed. (* --------------------------------------------------------------------------- *) (* Stuff on bitvectors, previously in memoryInjections.ma *) (* --------------------------------------------------------------------------- *) (* --------------------------------------------------------------------------- *) (* Some general lemmas on bitvectors (offsets /are/ bitvectors) *) (* --------------------------------------------------------------------------- *) lemma add_with_carries_n_O : ∀n,bv. add_with_carries n bv (zero n) false = 〈bv,zero n〉. #n #bv whd in match (add_with_carries ????); elim bv // #n #hd #tl #Hind whd in match (fold_right2_i ????????); >Hind normalize cases n in tl; [ 1: #tl cases hd normalize @refl | 2: #n' #tl cases hd normalize @refl ] qed. lemma addition_n_0 : ∀n,bv. addition_n n bv (zero n) = bv. #n #bv whd in match (addition_n ???); >add_with_carries_n_O // qed. lemma replicate_Sn : ∀A,sz,elt. replicate A (S sz) elt = elt ::: (replicate A sz elt). // qed. lemma zero_Sn : ∀n. zero (S n) = false ::: (zero n). // qed. lemma negation_bv_Sn : ∀n. ∀xa. ∀a : BitVector n. negation_bv … (xa ::: a) = (notb xa) ::: (negation_bv … a). #n #xa #a normalize @refl qed. (* useful facts on carry_of *) lemma carry_of_TT : ∀x. carry_of true true x = true. // qed. lemma carry_of_TF : ∀x. carry_of true false x = x. // qed. lemma carry_of_FF : ∀x. carry_of false false x = false. // qed. lemma carry_of_lcomm : ∀x,y,z. carry_of x y z = carry_of y x z. * * * // qed. lemma carry_of_rcomm : ∀x,y,z. carry_of x y z = carry_of x z y. * * * // qed. definition one_bv ≝ λn. (\fst (add_with_carries … (zero n) (zero n) true)). lemma one_bv_Sn_aux : ∀n. ∀bits,flags : BitVector (S n). add_with_carries … (zero (S n)) (zero (S n)) true = 〈bits, flags〉 → add_with_carries … (zero (S (S n))) (zero (S (S n))) true = 〈false ::: bits, false ::: flags〉. #n elim n [ 1: #bits #flags elim (BitVector_Sn … bits) #hd_bits * #tl_bits #Heq_bits elim (BitVector_Sn … flags) #hd_flags * #tl_flags #Heq_flags >(BitVector_O … tl_flags) >(BitVector_O … tl_bits) normalize #Heq destruct (Heq) @refl | 2: #n' #Hind #bits #flags elim (BitVector_Sn … bits) #hd_bits * #tl_bits #Heq_bits destruct #Hind >add_with_carries_Sn >replicate_Sn whd in match (zero ?) in Hind; lapply Hind elim (add_with_carries (S (S n')) (false:::replicate bool (S n') false) (false:::replicate bool (S n') false) true) #bits #flags #Heq destruct normalize >add_with_carries_Sn in Hind; elim (add_with_carries (S n') (replicate bool (S n') false) (replicate bool (S n') false) true) #flags' #bits' normalize cases (match bits' in Vector return λsz:ℕ.(λfoo:Vector bool sz.bool) with  [VEmpty⇒true|VCons (sz:ℕ)   (cy:bool)   (tl:(Vector bool sz))⇒cy]) normalize #Heq destruct @refl ] qed. lemma one_bv_Sn : ∀n. one_bv (S (S n)) = false ::: (one_bv (S n)). #n lapply (one_bv_Sn_aux n) whd in match (one_bv ?) in ⊢ (? → (??%%)); elim (add_with_carries (S n) (zero (S n)) (zero (S n)) true) #bits #flags #H lapply (H bits flags (refl ??)) #H2 >H2 @refl qed. lemma increment_to_addition_n_aux : ∀n. ∀a : BitVector n. add_with_carries ? a (zero n) true = add_with_carries ? a (one_bv n) false. #n elim n [ 1: #a >(BitVector_O … a) normalize @refl | 2: #n' cases n' [ 1: #Hind #a elim (BitVector_Sn ? a) #xa * #tl #Heq destruct >(BitVector_O … tl) normalize cases xa @refl | 2: #n'' #Hind #a elim (BitVector_Sn ? a) #xa * #tl #Heq destruct >one_bv_Sn >zero_Sn lapply (Hind tl) >add_with_carries_Sn >add_with_carries_Sn #Hind >Hind elim (add_with_carries (S n'') tl (one_bv (S n'')) false) #bits #flags normalize nodelta elim (BitVector_Sn … flags) #flags_hd * #flags_tl #Hflags_eq >Hflags_eq normalize nodelta @refl ] qed. (* In order to use associativity on increment, we hide it under addition_n. *) lemma increment_to_addition_n : ∀n. ∀a : BitVector n. increment ? a = addition_n ? a (one_bv n). #n whd in match (increment ??) in ⊢ (∀_.??%?); whd in match (addition_n ???) in ⊢ (∀_.???%); #a lapply (increment_to_addition_n_aux n a) #Heq >Heq cases (add_with_carries n a (one_bv n) false) #bits #flags @refl qed. (* Explicit formulation of addition *) (* Explicit formulation of the last carry bit *) let rec ith_carry (n : nat) (a,b : BitVector n) (init : bool) on n : bool ≝ match n return λx. BitVector x → BitVector x → bool with [ O ⇒ λ_,_. init | S x ⇒ λa',b'. let hd_a ≝ head' … a' in let hd_b ≝ head' … b' in let tl_a ≝ tail … a' in let tl_b ≝ tail … b' in carry_of hd_a hd_b (ith_carry x tl_a tl_b init) ] a b. lemma ith_carry_unfold : ∀n. ∀init. ∀a,b : BitVector (S n). ith_carry ? a b init = (carry_of (head' … a) (head' … b) (ith_carry ? (tail … a) (tail … b) init)). #n #init #a #b @refl qed. lemma ith_carry_Sn : ∀n. ∀init. ∀xa,xb. ∀a,b : BitVector n. ith_carry ? (xa ::: a) (xb ::: b) init = (carry_of xa xb (ith_carry ? a b init)). // qed. (* correction of [ith_carry] *) lemma ith_carry_ok : ∀n. ∀init. ∀a,b,res_ab,flags_ab : BitVector (S n). 〈res_ab,flags_ab〉 = add_with_carries ? a b init → head' … flags_ab = ith_carry ? a b init. #n elim n [ 1: #init #a #b #res_ab #flags_ab elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags destruct >(BitVector_O … tl_a) >(BitVector_O … tl_b) cases hd_a cases hd_b cases init normalize #Heq destruct (Heq) @refl | 2: #n' #Hind #init #a #b #res_ab #flags_ab elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags destruct lapply (Hind … init tl_a tl_b tl_res tl_flags) >add_with_carries_Sn >(ith_carry_Sn (S n')) elim (add_with_carries (S n') tl_a tl_b init) #res_ab #flags_ab elim (BitVector_Sn … flags_ab) #hd_flags_ab * #tl_flags_ab #Heq_flags_ab >Heq_flags_ab normalize nodelta cases hd_flags_ab normalize nodelta whd in match (head' ? (S n') ?); #H1 #H2 destruct (H2) lapply (H1 (refl ??)) whd in match (head' ???); #Heq (BitVector_O … tl_a) >(BitVector_O … tl_b) >(BitVector_O … tl_flags) >(BitVector_O … tl_res) normalize cases init cases hd_a cases hd_b normalize #Heq destruct @refl | 2: #n' #init #a #b #res_ab #flags_ab elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b elim (BitVector_Sn … res_ab) #hd_res * #tl_res #Heq_res elim (BitVector_Sn … flags_ab) #hd_flags * #tl_flags #Heq_flags destruct lapply (ith_carry_ok … init tl_a tl_b tl_res tl_flags) #Hcarry >add_with_carries_Sn elim (add_with_carries ? tl_a tl_b init) in Hcarry; #res #flags normalize nodelta elim (BitVector_Sn … flags) #hd_flags' * #tl_flags' #Heq_flags' >Heq_flags' normalize nodelta cases hd_flags' normalize nodelta #H1 #H2 destruct (H2) cases hd_a cases hd_b >ith_bit_Sn whd in match (head' ???) in H1 ⊢ %; <(H1 (refl ??)) @refl ] qed. (* Transform a function from bit-vectors to bits into a vector by folding *) let rec bitvector_fold (n : nat) (v : BitVector n) (f : ∀sz. BitVector sz → bool) on v : BitVector n ≝ match v with [ VEmpty ⇒ VEmpty ? | VCons sz elt tl ⇒ let bit ≝ f ? v in bit ::: (bitvector_fold ? tl f) ]. (* Two-arguments version *) let rec bitvector_fold2 (n : nat) (v1, v2 : BitVector n) (f : ∀sz. BitVector sz → BitVector sz → bool) on v1 : BitVector n ≝ match v1 with [ VEmpty ⇒ λ_. VEmpty ? | VCons sz elt tl ⇒ λv2'. let bit ≝ f ? v1 v2 in bit ::: (bitvector_fold2 ? tl (tail … v2') f) ] v2. lemma bitvector_fold2_Sn : ∀n,x1,x2,v1,v2,f. bitvector_fold2 (S n) (x1 ::: v1) (x2 ::: v2) f = (f ? (x1 ::: v1) (x2 ::: v2)) ::: (bitvector_fold2 … v1 v2 f). // qed. (* These functions pack all the relevant information (including carries) directly. *) definition addition_n_direct ≝ λn,v1,v2,init. bitvector_fold2 n v1 v2 (λn,v1,v2. ith_bit n v1 v2 init). lemma addition_n_direct_Sn : ∀n,x1,x2,v1,v2,init. addition_n_direct (S n) (x1 ::: v1) (x2 ::: v2) init = (ith_bit ? (x1 ::: v1) (x2 ::: v2) init) ::: (addition_n_direct … v1 v2 init). // qed. lemma tail_Sn : ∀n. ∀x. ∀a : BitVector n. tail … (x ::: a) = a. // qed. (* Prove the equivalence of addition_n_direct with add_with_carries *) lemma addition_n_direct_ok : ∀n,carry,v1,v2. (\fst (add_with_carries n v1 v2 carry)) = addition_n_direct n v1 v2 carry. #n elim n [ 1: #carry #v1 #v2 >(BitVector_O … v1) >(BitVector_O … v2) normalize @refl | 2: #n' #Hind #carry #v1 #v2 elim (BitVector_Sn … v1) #hd1 * #tl1 #Heq1 elim (BitVector_Sn … v2) #hd2 * #tl2 #Heq2 lapply (Hind carry tl1 tl2) lapply (ith_bit_ok ? carry v1 v2) lapply (ith_carry_ok ? carry v1 v2) destruct #Hind >addition_n_direct_Sn >ith_bit_Sn >add_with_carries_Sn elim (add_with_carries n' tl1 tl2 carry) #bits #flags normalize nodelta cases (match flags in Vector return λsz:ℕ.(λfoo:Vector bool sz.bool) with  [VEmpty⇒carry|VCons (sz:ℕ)   (cy:bool)   (tl:(Vector bool sz))⇒cy]) normalize nodelta #Hcarry' lapply (Hcarry' ?? (refl ??)) whd in match head'; normalize nodelta #H1 #H2 >H1 >H2 @refl ] qed. lemma addition_n_direct_ok2 : ∀n,carry,v1,v2. (let 〈a,b〉 ≝ add_with_carries n v1 v2 carry in a) = addition_n_direct n v1 v2 carry. #n #carry #v1 #v2 Heq >(BitVector_O … tl) cases hd normalize @refl | 2: #n' #Hind #a elim (BitVector_Sn … a) #hd * #tl #Heq >Heq lapply (Hind … tl) #Hind >one_bv_Sn >ith_carry_Sn whd in match (andb_fold ??); cases hd >Hind @refl ] qed. lemma ith_increment_bit : ∀n. ∀a : BitVector (S n). ith_bit … a (one_bv ?) false = xorb (head' … a) (andb_fold … (tail … a)). #n #a elim (BitVector_Sn … a) #hd * #tl #Heq >Heq whd in match (head' ???); -a cases n in tl; [ 1: #tl >(BitVector_O … tl) cases hd normalize try // | 2: #n' #tl >one_bv_Sn >ith_bit_Sn >ith_increment_carry >tail_Sn cases hd try // ] qed. (* Lemma used to prove involutivity of two-complement negation *) lemma twocomp_neg_involutive_aux : ∀n. ∀v : BitVector (S n). (andb_fold (S n) (negation_bv (S n) v) = andb_fold (S n) (negation_bv (S n) (addition_n_direct (S n) (negation_bv (S n) v) (one_bv (S n)) false))). #n elim n [ 1: #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq >(BitVector_O … tl) cases hd @refl | 2: #n' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq lapply (Hind tl) -Hind #Hind >negation_bv_Sn >one_bv_Sn >addition_n_direct_Sn >andb_fold_Sn >ith_bit_Sn >negation_bv_Sn >andb_fold_Sn xorb_false >(xorb_comm false ?) >xorb_false | 2: >xorb_false >(xorb_comm true ?) >xorb_true ] >ith_increment_carry cases (andb_fold (S n') (negation_bv (S n') tl)) @refl ] qed. (* Test of the 'direct' approach: proof of the involutivity of two-complement negation. *) lemma twocomp_neg_involutive : ∀n. ∀v : BitVector n. twocomp_neg_direct ? (twocomp_neg_direct ? v) = v. #n elim n [ 1: #v >(BitVector_O v) @refl | 2: #n' cases n' [ 1: #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq >(BitVector_O … tl) normalize cases hd @refl | 2: #n'' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq lapply (Hind tl) -Hind #Hind (negation_bv_Sn ? hd tl) >one_bv_Sn >(addition_n_direct_Sn ? (¬hd) false ??) >ith_bit_Sn >negation_bv_Sn >addition_n_direct_Sn >ith_bit_Sn generalize in match (addition_n_direct (S n'') (negation_bv (S n'') (addition_n_direct (S n'') (negation_bv (S n'') tl) (one_bv (S n'')) false)) (one_bv (S n'')) false); #tail >ith_increment_carry >ith_increment_carry cases hd normalize nodelta [ 1: normalize in match (xorb false false); >(xorb_comm false ?) >xorb_false >xorb_false | 2: normalize in match (xorb true false); >(xorb_comm true ?) >xorb_true >xorb_false ] (BitVector_O … a) >(BitVector_O … b) normalize #_ @conj // | 2: #n' cases n' [ 1: #_ #a #b elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a >Heq_a elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b >Heq_b >(BitVector_O … tl_a) >(BitVector_O … tl_b) cases hd_a cases hd_b normalize #H @conj try // | 2: #n'' #Hind #a #b elim (BitVector_Sn … a) #hd_a * #tl_a #Heq_a >Heq_a elim (BitVector_Sn … b) #hd_b * #tl_b #Heq_b >Heq_b lapply (Hind … tl_a tl_b) -Hind #Hind >one_bv_Sn >addition_n_direct_Sn >ith_bit_Sn >addition_n_direct_Sn >ith_bit_Sn >xorb_false >xorb_false #H elim (bitvector_cons_inj_inv … H) #Heq1 #Heq2 lapply (Hind Heq2) * #Heq3 #Heq4 cut (hd_a = hd_b) [ 1: >Heq4 in Heq1; #Heq5 lapply (xorb_inj (ith_carry ? tl_b (one_bv ?) false) hd_a hd_b) * #Heq6 #_ >xorb_comm in Heq6; >(xorb_comm ? hd_b) #Heq6 >(Heq6 Heq5) @refl ] #Heq5 @conj [ 1: >Heq3 >Heq5 @refl ] >ith_carry_Sn >ith_carry_Sn >Heq4 >Heq5 @refl ] qed. (* Inverse of injecivity of increment, does not lose information (cf increment_inj) *) lemma increment_inj_inv : ∀n. ∀a,b : BitVector n. a = b → increment_direct ? a = increment_direct ? b. // qed. (* A more general result. *) lemma addition_n_direct_inj : ∀n. ∀x,y,delta: BitVector n. addition_n_direct ? x delta false = addition_n_direct ? y delta false → x = y ∧ (ith_carry n x delta false = ith_carry n y delta false). #n elim n [ 1: #x #y #delta >(BitVector_O … x) >(BitVector_O … y) >(BitVector_O … delta) * @conj @refl | 2: #n' #Hind #x #y #delta elim (BitVector_Sn … x) #hdx * #tlx #Heqx >Heqx elim (BitVector_Sn … y) #hdy * #tly #Heqy >Heqy elim (BitVector_Sn … delta) #hdd * #tld #Heqd >Heqd >addition_n_direct_Sn >ith_bit_Sn >addition_n_direct_Sn >ith_bit_Sn >ith_carry_Sn >ith_carry_Sn lapply (Hind … tlx tly tld) -Hind #Hind #Heq elim (bitvector_cons_inj_inv … Heq) #Heq_hd #Heq_tl lapply (Hind Heq_tl) -Hind * #HindA #HindB >HindA >HindB >HindB in Heq_hd; #Heq_hd cut (hdx = hdy) [ 1: cases hdd in Heq_hd; cases (ith_carry n' tly tld false) cases hdx cases hdy normalize #H try @H try @refl >H try @refl ] #Heq_hd >Heq_hd @conj @refl ] qed. (* We also need it the other way around. *) lemma addition_n_direct_inj_inv : ∀n. ∀x,y,delta: BitVector n. x ≠ y → (* ∧ (ith_carry n x delta false = ith_carry n y delta false). *) addition_n_direct ? x delta false ≠ addition_n_direct ? y delta false. #n elim n [ 1: #x #y #delta >(BitVector_O … x) >(BitVector_O … y) >(BitVector_O … delta) * #H @(False_ind … (H (refl ??))) | 2: #n' #Hind #x #y #delta elim (BitVector_Sn … x) #hdx * #tlx #Heqx >Heqx elim (BitVector_Sn … y) #hdy * #tly #Heqy >Heqy elim (BitVector_Sn … delta) #hdd * #tld #Heqd >Heqd #Hneq cut (hdx ≠ hdy ∨ tlx ≠ tly) [ @(eq_bv_elim … tlx tly) [ #Heq_tl >Heq_tl >Heq_tl in Hneq; #Hneq cut (hdx ≠ hdy) [ % #Heq_hd >Heq_hd in Hneq; * #H @H @refl ] #H %1 @H | #H %2 @H ] ] -Hneq #Hneq >addition_n_direct_Sn >addition_n_direct_Sn >ith_bit_Sn >ith_bit_Sn cases Hneq [ 1: #Hneq_hd lapply (addition_n_direct_inj … tlx tly tld) @(eq_bv_elim … (addition_n_direct ? tlx tld false) (addition_n_direct ? tly tld false)) [ 1: #Heq -Hind #Hind elim (Hind Heq) #Heq_tl >Heq_tl #Heq_carry >Heq_carry % #Habsurd elim (bitvector_cons_inj_inv … Habsurd) -Habsurd lapply Hneq_hd cases hdx cases hdd cases hdy cases (ith_carry ? tly tld false) normalize in ⊢ (? → % → ?); #Hneq_hd #Heq_hd #_ try @(absurd … Heq_hd Hneq_hd) elim Hneq_hd -Hneq_hd #Hneq_hd @Hneq_hd try @refl try assumption try @(sym_eq … Heq_hd) | 2: #Htl_not_eq #_ % #Habsurd elim (bitvector_cons_inj_inv … Habsurd) #_ elim Htl_not_eq -Htl_not_eq #HA #HB @HA @HB ] | 2: #Htl_not_eq lapply (Hind tlx tly tld Htl_not_eq) -Hind #Hind % #Habsurd elim (bitvector_cons_inj_inv … Habsurd) #_ elim Hind -Hind #HA #HB @HA @HB ] ] qed. lemma carry_notb : ∀a,b,c. notb (carry_of a b c) = carry_of (notb a) (notb b) (notb c). * * * @refl qed. lemma increment_to_carry_aux : ∀n. ∀a : BitVector (S n). ith_carry (S n) a (one_bv (S n)) false = ith_carry (S n) a (zero (S n)) true. #n elim n [ 1: #a elim (BitVector_Sn ? a) #hd_a * #tl_a #Heq >Heq >(BitVector_O … tl_a) @refl | 2: #n' #Hind #a elim (BitVector_Sn ? a) #hd_a * #tl_a #Heq >Heq lapply (Hind tl_a) #Hind >one_bv_Sn >zero_Sn >ith_carry_Sn >ith_carry_Sn >Hind @refl ] qed. lemma neutral_addition_n_direct_aux : ∀n. ∀v. ith_carry n v (zero n) false = false. #n elim n // #n' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq >zero_Sn >ith_carry_Sn >(Hind tl) cases hd @refl. qed. lemma neutral_addition_n_direct : ∀n. ∀v : BitVector n. addition_n_direct ? v (zero ?) false = v. #n elim n [ 1: #v >(BitVector_O … v) normalize @refl | 2: #n' #Hind #v elim (BitVector_Sn … v) #hd * #tl #Heq >Heq lapply (Hind … tl) #H >zero_Sn >addition_n_direct_Sn >ith_bit_Sn >H >xorb_false >neutral_addition_n_direct_aux >xorb_false @refl ] qed. lemma increment_to_carry_zero : ∀n. ∀a : BitVector n. addition_n_direct ? a (one_bv ?) false = addition_n_direct ? a (zero ?) true. #n elim n [ 1: #a >(BitVector_O … a) normalize @refl | 2: #n' cases n' [ 1: #_ #a elim (BitVector_Sn … a) #hd_a * #tl_a #Heq >Heq >(BitVector_O … tl_a) cases hd_a @refl | 2: #n'' #Hind #a elim (BitVector_Sn … a) #hd_a * #tl_a #Heq >Heq lapply (Hind tl_a) -Hind #Hind >one_bv_Sn >zero_Sn >addition_n_direct_Sn >ith_bit_Sn >addition_n_direct_Sn >ith_bit_Sn >xorb_false >Hind @bitvector_cons_eq >increment_to_carry_aux @refl ] ] qed. lemma increment_to_carry : ∀n. ∀a,b : BitVector n. addition_n_direct ? a (addition_n_direct ? b (one_bv ?) false) false = addition_n_direct ? a b true. #n #a #b >increment_to_carry_zero neutral_addition_n_direct @refl qed. lemma increment_direct_ok : ∀n,v. increment_direct n v = increment n v. #n #v whd in match (increment ??); >addition_n_direct_ok associative_addition_n_direct >(commutative_addition_n_direct ? (one_bv n)) >increment_to_carry -H lapply b lapply a -b -a cases n [ 1: #a #b >(BitVector_O … a) >(BitVector_O … b) @refl | 2: #n' #a #b cut (negation_bv ? (addition_n_direct ? a b false) = addition_n_direct ? (negation_bv ? a) (negation_bv ? b) true ∧ notb (ith_carry ? a b false) = (ith_carry ? (negation_bv ? a) (negation_bv ? b) true)) [ -n lapply b lapply a elim n' [ 1: #a #b elim (BitVector_Sn … a) #hd_a * #tl_a #Heqa >Heqa >(BitVector_O … tl_a) elim (BitVector_Sn … b) #hd_b * #tl_b #Heqb >Heqb >(BitVector_O … tl_b) cases hd_a cases hd_b normalize @conj @refl | 2: #n #Hind #a #b elim (BitVector_Sn … a) #hd_a * #tl_a #Heqa >Heqa elim (BitVector_Sn … b) #hd_b * #tl_b #Heqb >Heqb lapply (Hind tl_a tl_b) * #H1 #H2 @conj [ 2: >ith_carry_Sn >negation_bv_Sn >negation_bv_Sn >ith_carry_Sn >carry_notb >H2 @refl | 1: >addition_n_direct_Sn >ith_bit_Sn >negation_bv_Sn >negation_bv_Sn >negation_bv_Sn >addition_n_direct_Sn >ith_bit_Sn >H1 @bitvector_cons_eq >xorb_lneg >xorb_rneg >notb_notb H2 @refl ] ] ] * #H1 #H2 @H1 ] qed. lemma addition_n_direct_neg : ∀n. ∀a. (addition_n_direct n a (negation_bv n a) false) = replicate ?? true ∧ (ith_carry n a (negation_bv n a) false = false). #n elim n [ 1: #a >(BitVector_O … a) @conj @refl | 2: #n' #Hind #a elim (BitVector_Sn … a) #hd * #tl #Heq >Heq lapply (Hind … tl) -Hind * #HA #HB @conj [ 2: >negation_bv_Sn >ith_carry_Sn >HB cases hd @refl | 1: >negation_bv_Sn >addition_n_direct_Sn >ith_bit_Sn >HB >xorb_false >HA @bitvector_cons_eq elim hd @refl ] ] qed. (* -a + a = 0 *) lemma bitvector_opp_direct : ∀n. ∀a : BitVector n. addition_n_direct ? a (twocomp_neg_direct ? a) false = (zero ?). whd in match twocomp_neg_direct; whd in match increment_direct; normalize nodelta #n #a H -H -a cases n try // #n' cut ((addition_n_direct (S n') (replicate bool ? true) (one_bv ?) false = (zero (S n'))) ∧ (ith_carry ? (replicate bool (S n') true) (one_bv (S n')) false = true)) [ elim n' [ 1: @conj @refl | 2: #n' * #HA #HB @conj [ 1: >replicate_Sn >one_bv_Sn >addition_n_direct_Sn >ith_bit_Sn >HA >zero_Sn @bitvector_cons_eq >HB @refl | 2: >replicate_Sn >one_bv_Sn >ith_carry_Sn >HB @refl ] ] ] * #H1 #H2 @H1 qed. (* Lift back the previous result to standard operations. *) lemma twocomp_neg_direct_ok : ∀n. ∀v. twocomp_neg_direct ? v = two_complement_negation n v. #n #v whd in match twocomp_neg_direct; normalize nodelta whd in match increment_direct; normalize nodelta whd in match two_complement_negation; normalize nodelta >increment_to_addition_n twocomp_neg_direct_ok >twocomp_neg_direct_ok >twocomp_neg_direct_ok twocomp_neg_direct_ok Hclassify normalize nodelta [ 1,2: #H destruct ] cases v2 normalize nodelta [ | #sz' #i' | | #p' | | #sz' #i' | | #p' ] #H2 destruct cases target in H2; [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id | | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] normalize nodelta #H destruct [ 2,4,5,6,7,8,9: cases (eq_block (pblock p) (pblock p')) in H; normalize nodelta #H destruct cases (eqb (sizeof ty1) O) in H; normalize nodelta #H destruct ] %{sz} %{sg} @conj try @refl try /4 by or_introl, or_intror, conj, refl/ cases (if_opt_inversion ???? H) #Hblocks_eq -H @(eqb_elim … (sizeof ty1) 0) normalize nodelta [ #Heq_sizeof #Habsurd destruct ] #_ #Hdiv %1 %{p} %{p'} cases (division_u ???) in Hdiv; normalize nodelta [ #Habsurd destruct ] #i #Heq destruct %{i} try @conj try @conj try @conj try @conj try @refl try @(eq_block_to_refl … Hblocks_eq) qed. lemma sem_sub_pi_inversion : ∀sz,sg,ty',optlen,target. ∀v1,v2,res. sem_sub v1 (ptr_type ty' optlen) v2 (Tint sz sg) target = Some ? res → ∃sz',i. v2 = Vint sz' i ∧ ((∃p. v1 = Vptr p ∧ res = Vptr (neg_shift_pointer_n ? p (sizeof ty') sg i)) ∨ (v1 = Vnull ∧ i = (zero ?) ∧ res = Vnull)). #tsz #tsg #ty' * [ | #n ] #target * [ | #sz' #i' | | #p' | | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' | | #sz2' #i2' | | #p2' | | #sz2' #i2' | | #p2' | | #sz2' #i2' | | #p2' ] #H destruct [ 2,4: %{sz2'} %{i2'} @conj try @refl %1 %{p'} @conj try @refl | *: lapply H -H @(eq_bv_elim … i2' (zero ?)) normalize nodelta #Hi #Heq destruct (Heq) %{sz2'} %{(zero ?)} @conj destruct try @refl %2 @conj try @conj try @refl ] qed. lemma sem_sub_ii_inversion : ∀sz,sg,ty. ∀v1,v2,res. sem_sub v1 (Tint sz sg) v2 (Tint sz sg) ty = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ res = Vint sz' (subtraction (bitsize_of_intsize sz') i1 i2). #sz #sg #ty * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); whd in match (classify_sub ??); cases sz cases sg normalize nodelta #H destruct cases v2 in H; normalize nodelta #H1 destruct #H2 destruct #Heq %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases H1 in H2; #j' normalize nodelta #Heq destruct (Heq) %{i'} %{j'} @conj try @conj try @conj try @refl qed. lemma sem_mul_inversion : ∀sz,sg. ∀v1,v2,res. sem_mul v1 (Tint sz sg) v2 (Tint sz sg) = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ res = Vint sz' (short_multiplication ? i1 i2). #sz #sg * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); whd in match (classify_aop ??); cases sz cases sg normalize nodelta #H destruct cases v2 in H; normalize nodelta #H1 destruct #H2 destruct #Heq %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases H1 in H2; #j' normalize nodelta #Heq destruct (Heq) %{i'} %{j'} @conj try @conj try @conj try @refl qed. lemma sem_div_inversion_s : ∀sz. ∀v1,v2,res. sem_div v1 (Tint sz Signed) v2 (Tint sz Signed) = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ match division_s ? i1 i2 with [ Some q ⇒ res = (Vint sz' q) | None ⇒ False ]. #sz * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); whd in match (classify_aop ??); >type_eq_dec_true normalize nodelta #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta whd in match (option_map ????); #Hdiv destruct %{i'} %{i2'} @conj try @conj try @conj try @refl cases (division_s ???) in Hdiv; [ 2,4,6: #bv ] normalize #H destruct (H) try @refl qed. lemma sem_div_inversion_u : ∀sz. ∀v1,v2,res. sem_div v1 (Tint sz Unsigned) v2 (Tint sz Unsigned) = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ match division_u ? i1 i2 with [ Some q ⇒ res = (Vint sz' q) | None ⇒ False ]. #sz * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); whd in match (classify_aop ??); >type_eq_dec_true normalize nodelta #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta whd in match (option_map ????); #Hdiv destruct %{i'} %{i2'} @conj try @conj try @conj try @refl cases (division_u ???) in Hdiv; [ 2,4,6: #bv ] normalize #H destruct (H) try @refl qed. lemma sem_mod_inversion_s : ∀sz. ∀v1,v2,res. sem_mod v1 (Tint sz Signed) v2 (Tint sz Signed) = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ match modulus_s ? i1 i2 with [ Some q ⇒ res = (Vint sz' q) | None ⇒ False ]. #sz * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); whd in match (classify_aop ??); >type_eq_dec_true normalize nodelta #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta whd in match (option_map ????); #Hdiv destruct %{i'} %{i2'} @conj try @conj try @conj try @refl cases (modulus_s ???) in Hdiv; [ 2,4,6: #bv ] normalize #H destruct (H) try @refl qed. lemma sem_mod_inversion_u : ∀sz. ∀v1,v2,res. sem_mod v1 (Tint sz Unsigned) v2 (Tint sz Unsigned) = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ match modulus_u ? i1 i2 with [ Some q ⇒ res = (Vint sz' q) | None ⇒ False ]. #sz * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); whd in match (classify_aop ??); >type_eq_dec_true normalize nodelta #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta whd in match (option_map ????); #Hdiv destruct %{i'} %{i2'} @conj try @conj try @conj try @refl cases (modulus_u ???) in Hdiv; [ 2,4,6: #bv ] normalize #H destruct (H) try @refl qed. lemma sem_and_inversion : ∀v1,v2,res. sem_and v1 v2 = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ res = Vint sz' (conjunction_bv ? i1 i2). * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta #H destruct %{i'} %{i2'} @conj try @conj try @conj try @refl qed. lemma sem_or_inversion : ∀v1,v2,res. sem_or v1 v2 = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ res = Vint sz' (inclusive_disjunction_bv ? i1 i2). * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta #H destruct %{i'} %{i2'} @conj try @conj try @conj try @refl qed. lemma sem_xor_inversion : ∀v1,v2,res. sem_xor v1 v2 = Some ? res → ∃sz',i1,i2. v1 = Vint sz' i1 ∧ v2 = Vint sz' i2 ∧ res = Vint sz' (exclusive_disjunction_bv ? i1 i2). * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} lapply Heq -Heq cases sz' in i'; #i' whd in match (intsize_eq_elim ???????); cases sz2' in i2'; #i2' normalize nodelta #H destruct %{i'} %{i2'} @conj try @conj try @conj try @refl qed. lemma sem_shl_inversion : ∀v1,v2,res. sem_shl v1 v2 = Some ? res → ∃sz1,sz2,i1,i2. v1 = Vint sz1 i1 ∧ v2 = Vint sz2 i2 ∧ res = Vint sz1 (shift_left bool (bitsize_of_intsize sz1) (nat_of_bitvector (bitsize_of_intsize sz2) i2) i1 false) ∧ lt_u (bitsize_of_intsize sz2) i2 (bitvector_of_nat (bitsize_of_intsize sz2) (bitsize_of_intsize sz1)) = true. * [ | #sz' #i' | | #p' ] #v2 #res whd in ⊢ ((??%?) → ?); #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} %{sz2'} lapply (if_opt_inversion ???? Heq) * #Hlt_u #Hres %{i'} %{i2'} >Hlt_u destruct (Hres) @conj try @conj try @conj try @refl qed. lemma sem_shr_inversion : ∀v1,v2,sz,sg,res. sem_shr v1 (Tint sz sg) v2 (Tint sz sg) = Some ? res → ∃sz1,sz2,i1,i2. v1 = Vint sz1 i1 ∧ v2 = Vint sz2 i2 ∧ lt_u (bitsize_of_intsize sz2) i2 (bitvector_of_nat (bitsize_of_intsize sz2) (bitsize_of_intsize sz1)) = true ∧ match sg with [ Signed ⇒ res = (Vint sz1 (shift_right bool (7+pred_size_intsize sz1*8) (nat_of_bitvector (bitsize_of_intsize sz2) i2) i1 (head' bool (7+pred_size_intsize sz1*8) i1))) | Unsigned ⇒ res = (Vint sz1 (shift_right bool (7+pred_size_intsize sz1*8) (nat_of_bitvector (bitsize_of_intsize sz2) i2) i1 false)) ]. * [ | #sz' #i' | | #p' ] #v2 #sz * #res whd in ⊢ ((??%?) → ?); whd in match (classify_aop ??); >type_eq_dec_true normalize nodelta #H destruct cases v2 in H; normalize nodelta [ | #sz2' #i2' | | #p2' | | #sz2' #i2' | | #p2' ] #Heq destruct %{sz'} %{sz2'} lapply (if_opt_inversion ???? Heq) * #Hlt_u #Hres %{i'} %{i2'} >Hlt_u destruct (Hres) @conj try @conj try @conj try @refl qed. lemma sem_cmp_int_inversion : ∀v1,v2,sz,sg,op,m,res. sem_cmp op v1 (Tint sz sg) v2 (Tint sz sg) m = Some ? res → ∃sz,i1,i2. v1 = Vint sz i1 ∧ v2 = Vint sz i2 ∧ match sg with [ Unsigned ⇒ of_bool (cmpu_int ? op i1 i2) = res | Signed ⇒ of_bool (cmp_int ? op i1 i2) = res ]. #v1 #v2 #sz0 #sg #op * #contents #next #Hnext #res whd in ⊢ ((??%?) → ?); whd in match (classify_cmp ??); >type_eq_dec_true normalize nodelta cases v1 [ | #sz #i | | #p ] normalize nodelta #H destruct cases v2 in H; [ | #sz' #i' | | #p' ] normalize nodelta #H destruct lapply H -H cases sz in i; #i cases sz' in i'; #i' cases sg normalize nodelta whd in match (intsize_eq_elim ???????); #H destruct [ 1,2: %{I8} | 3,4: %{I16} | 5,6: %{I32} ] %{i} %{i'} @conj try @conj try @refl qed. lemma sem_cmp_ptr_inversion : ∀v1,v2,ty',n,op,m,res. sem_cmp op v1 (ptr_type ty' n) v2 (ptr_type ty' n) m = Some ? res → (∃p1,p2. v1 = Vptr p1 ∧ v2 = Vptr p2 ∧ valid_pointer m p1 = true ∧ valid_pointer m p2 = true ∧ (if eq_block (pblock p1) (pblock p2) then Some ? (of_bool (cmp_offset op (poff p1) (poff p2))) else sem_cmp_mismatch op) = Some ? res) ∨ (∃p1. v1 = Vptr p1 ∧ v2 = Vnull ∧ sem_cmp_mismatch op = Some ? res) ∨ (∃p2. v1 = Vnull ∧ v2 = Vptr p2 ∧ sem_cmp_mismatch op = Some ? res) ∨ (v1 = Vnull ∧ v2 = Vnull ∧ sem_cmp_match op = Some ? res). * [ | #sz #i | | #p ] normalize nodelta #v2 #ty' #n #op * #contents #nextblock #Hnextblock #res whd in ⊢ ((??%?) → ?); whd in match (classify_cmp ??); cases n normalize nodelta [ 2,4,6,8: #array_len ] whd in match (ptr_type ty' ?); whd in match (if_type_eq ?????); >type_eq_dec_true normalize nodelta #H destruct cases v2 in H; normalize nodelta [ | #sz' #i' | | #p' | | #sz' #i' | | #p' | | #sz' #i' | | #p' | | #sz' #i' | | #p' ] #H destruct try /6 by or_introl, or_intror, ex_intro, conj/ [ 1: %1 %1 %2 %{p} @conj try @conj // | 3: %1 %1 %2 %{p} @conj try @conj // | *: %1 %1 %1 %{p} %{p'} cases (valid_pointer (mk_mem ???) p) in H; normalize nodelta cases (valid_pointer (mk_mem contents nextblock Hnextblock) p') normalize nodelta #H try @conj try @conj try @conj try @conj try @conj try @refl try @H destruct ] qed.