include "Clight/Cexec.ma". include "Clight/frontend_misc.ma". (* This file contains some definitions and lemmas related to memory injections. Could be useful for the Clight → Cminor. Needs revision: the definitions are too lax and allow inconsistent embeddings (more precisely, these embeddings do not allow to prove that the semantics for pointer less-than comparisons is conserved). *) (* --------------------------------------------------------------------------- *) (* 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. (* trivially lift associativity to our new setting *) lemma associative_addition_n_direct : ∀n. ∀carry1,carry2. ∀v1,v2,v3 : BitVector n. addition_n_direct ? (addition_n_direct ? v1 v2 carry1) v3 carry2 = addition_n_direct ? v1 (addition_n_direct ? v2 v3 carry1) carry2. #n #carry1 #carry2 #v1 #v2 #v3 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. (* Prove -(a + b) = -a + -b *) lemma twocomp_neg_plus : ∀n. ∀a,b : BitVector n. twocomp_neg_direct ? (addition_n_direct ? a b false) = addition_n_direct ? (twocomp_neg_direct … a) (twocomp_neg_direct … b) false. whd in match twocomp_neg_direct; normalize nodelta lapply increment_inj_inv whd in match increment_direct; normalize nodelta #H #n #a #b 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 Heq %1 /2/ | 2: #Hneq cases Hind [ 1: #Hmem %1 /2/ | 2: #Hnmem %2 normalize % #H cases H [ 1: lapply Hneq * #Hl #Eq @(Hl Eq) | 2: lapply Hnmem * #Hr #Hmem @(Hr Hmem) ] ] ] ] qed. lemma block_eq_dec : ∀a,b:block. a = b ∨ a ≠ b. #a #b @(eq_block_elim … a b) /2/ qed. lemma mem_not_mem_neq : ∀l,elt1,elt2. (mem block elt1 l) → ¬ (mem block elt2 l) → elt1 ≠ elt2. #l #elt1 #elt2 elim l [ 1: normalize #Habsurd @(False_ind … Habsurd) | 2: #hd #tl #Hind normalize #Hl #Hr cases Hl [ 1: #Heq >Heq @(eq_block_elim … hd elt2) [ 1: #Heq >Heq /2 by not_to_not/ | 2: #x @x ] | 2: #Hmem1 cases (mem_dec … block_eq_dec elt2 tl) [ 1: #Hmem2 % #Helt_eq cases Hr #H @H %2 @Hmem2 | 2: #Hmem2 @Hind // ] ] ] qed. lemma neq_block_eq_block_false : ∀b1,b2:block. b1 ≠ b2 → eq_block b1 b2 = false. #b1 #b2 * #Hb @(eq_block_elim … b1 b2) [ 1: #Habsurd @(False_ind … (Hb Habsurd)) | 2: // ] qed. (* --------------------------------------------------------------------------- *) (* Lemmas related to freeing memory and pointer validity. *) (* --------------------------------------------------------------------------- *) (* lemma nextblock_free_ok : ∀m,b. nextblock m = nextblock (free m b). * #contents #next #posnext * #rg #id normalize @refl qed. lemma low_bound_free_ok : ∀m,b,ptr. b ≠ (pblock ptr) → Zleb (low_bound (free m b) (pblock ptr)) (Z_of_unsigned_bitvector offset_size (offv (poff ptr))) = Zleb (low_bound m (pblock ptr)) (Z_of_unsigned_bitvector offset_size (offv (poff ptr))). * #contents #next #nextpos * #brg #bid * * #prg #pid #poff normalize cases prg cases brg normalize nodelta try // @(eqZb_elim pid bid) [ 1,3: #Heq >Heq * #Habsurd @(False_ind … (Habsurd (refl ??))) | 2,4: #_ #_ @refl ] qed. lemma high_bound_free_ok : ∀m,b,ptr. b ≠ (pblock ptr) → Zleb (Z_of_unsigned_bitvector offset_size (offv (poff ptr))) (high_bound (free m b) (pblock ptr)) = Zleb (Z_of_unsigned_bitvector offset_size (offv (poff ptr))) (high_bound m (pblock ptr)). * #contents #next #nextpos * #brg #bid * * #prg #pid #poff normalize cases prg cases brg normalize nodelta try // @(eqZb_elim pid bid) [ 1,3: #Heq >Heq * #Habsurd @(False_ind … (Habsurd (refl ??))) | 2,4: #_ #_ @refl ] qed. lemma valid_pointer_free_ok : ∀b : block. ∀m,ptr. valid_pointer m ptr = true → pblock ptr ≠ b → valid_pointer (free m b) ptr = true. * #br #bid * #contents #next #posnext * * #pbr #pbid #poff normalize @(eqZb_elim pbid bid) [ 1: #Heqid >Heqid cases pbr cases br normalize [ 1,4: #_ * #Habsurd @(False_ind … (Habsurd (refl ??))) ] cases (Zltb bid next) normalize nodelta [ 2,4: #Habsurd destruct (Habsurd) ] // | 2: #Hneqid cases pbr cases br normalize try // ] qed. lemma valid_pointer_free_list_ok : ∀blocks : list block. ∀m,ptr. valid_pointer m ptr = true → ¬ mem ? (pblock ptr) blocks → valid_pointer (free_list m blocks) ptr = true. #blocks elim blocks [ 1: #m #ptr normalize #H #_ @H | 2: #b #tl #Hind #m #ptr #Hvalid whd in match (mem block (pblock ptr) ?); >free_list_cons * #Hguard @valid_pointer_free_ok [ 2: % #Heq @Hguard %1 @Heq ] @Hind [ 1: @Hvalid | 2: % #Hmem @Hguard %2 @Hmem ] ] qed. *) lemma valid_pointer_ok_free : ∀b : block. ∀m,ptr. valid_pointer m ptr = true → pblock ptr ≠ b → valid_pointer (free m b) ptr = true. * #br #bid * #contents #next #posnext * * #pbr #pbid #poff normalize @(eqZb_elim pbid bid) [ 1: #Heqid >Heqid cases pbr cases br normalize [ 1,4: #_ * #Habsurd @(False_ind … (Habsurd (refl ??))) ] cases (Zltb bid next) normalize nodelta [ 2,4: #Habsurd destruct (Habsurd) ] // | 2: #Hneqid cases pbr cases br normalize try // ] qed. lemma free_list_cons : ∀m,b,tl. free_list m (b :: tl) = (free (free_list m tl) b). #m #b #tl whd in match (free_list ??) in ⊢ (??%%); whd in match (free ??); @refl qed. lemma valid_pointer_free_ok : ∀b : block. ∀m,ptr. valid_pointer (free m b) ptr = true → pblock ptr ≠ b → (* This hypothesis is necessary to handle the cse of "valid" pointers to empty *) valid_pointer m ptr = true. * #br #bid * #contents #next #posnext * * #pbr #pbid #poff @(eqZb_elim pbid bid) [ 1: #Heqid >Heqid cases pbr cases br normalize [ 1,4: #_ * #Habsurd @(False_ind … (Habsurd (refl ??))) ] cases (Zltb bid next) normalize nodelta [ 2,4: #Habsurd destruct (Habsurd) ] // | 2: #Hneqid cases pbr cases br normalize try // >(eqZb_false … Hneqid) normalize nodelta try // qed. lemma valid_pointer_free_list_ok : ∀blocks : list block. ∀m,ptr. valid_pointer (free_list m blocks) ptr = true → ¬ mem ? (pblock ptr) blocks → valid_pointer m ptr = true. #blocks elim blocks [ 1: #m #ptr normalize #H #_ @H | 2: #b #tl #Hind #m #ptr #Hvalid whd in match (mem block (pblock ptr) ?); >free_list_cons in Hvalid; #Hvalid * #Hguard lapply (valid_pointer_free_ok … Hvalid) #H @Hind [ 2: % #Heq @Hguard %2 @Heq | 1: @H % #Heq @Hguard %1 @Heq ] ] qed. (* An alternative version without the gard on the equality of blocks. *) lemma valid_pointer_free_ok_alt : ∀b : block. ∀m,ptr. valid_pointer (free m b) ptr = true → (pblock ptr) = b ∨ (pblock ptr ≠ b ∧ valid_pointer m ptr = true). * #br #bid * #contents #next #posnext * * #pbr #pbid #poff @(eq_block_elim … (mk_block br bid) (mk_block pbr pbid)) [ 1: #Heq #_ %1 @(sym_eq … Heq) | 2: cases pbr cases br normalize nodelta [ 1,4: @(eqZb_elim pbid bid) [ 1,3: #Heq >Heq * #H @(False_ind … (H (refl ??))) | 2,4: #Hneq #_ normalize >(eqZb_false … Hneq) normalize nodelta #H %2 @conj try assumption % #Habsurd destruct (Habsurd) elim Hneq #Hneq @Hneq try @refl ] | *: #_ #H %2 @conj try @H % #Habsurd destruct (Habsurd) ] ] qed. (* Well in fact a valid pointer can be valid even after a free. Ah. *) (* lemma pointer_in_free_list_not_valid : ∀blocks,m,ptr. mem ? (pblock ptr) blocks → valid_pointer (free_list m blocks) ptr = false. *) (* #blocks elim blocks [ 1: #m #ptr normalize #Habsurd @(False_ind … Habsurd) | 2: #b #tl #Hind #m #ptr whd in match (mem block (pblock ptr) ?); >free_list_cons * [ 1: #Hptr_match whd in match (free ??); whd in match (update_block ????); whd in match (valid_pointer ??); whd in match (low_bound ??); whd in match (high_bound ??); >Hptr_match >eq_block_identity normalize nodelta whd in match (low ?); cut (Zleb OZ (Z_of_unsigned_bitvector offset_size (offv (poff ptr))) = true) [ 1: elim (offv (poff ptr)) // #n' #hd #tl cases hd normalize [ 1: #_ try @refl | 2: #H @H ] ] #H >H cut (Zleb (Z_of_unsigned_bitvector offset_size (offv (poff ptr))) OZ = false) [ 1: elim (offv (poff ptr)) normalize #n' #hd #tl cases hd normalize [ 1: normalize #_ try @refl | 2: #H @H ] ] valid_pointer (free_list m blocks) ptr = false. *) (* --------------------------------------------------------------------------- *) (* Memory injections *) (* --------------------------------------------------------------------------- *) (* An embedding is a function from blocks to (blocks × offset). *) definition embedding ≝ block → option (block × offset). definition offset_plus : offset → offset → offset ≝ λo1,o2. mk_offset (addition_n ? (offv o1) (offv o2)). lemma offset_plus_0 : ∀o. offset_plus o (mk_offset (zero ?)) = o. * #o whd in match (offset_plus ??); >addition_n_0 @refl qed. (* Translates a pointer through an embedding. *) definition pointer_translation : ∀p:pointer. ∀E:embedding. option pointer ≝ λp,E. match p with [ mk_pointer pblock poff ⇒ match E pblock with [ None ⇒ None ? | Some loc ⇒ let 〈dest_block,dest_off〉 ≝ loc in let ptr ≝ (mk_pointer dest_block (offset_plus poff dest_off)) in Some ? ptr ] ]. (* We parameterise the "equivalence" relation on values with an embedding. *) (* Front-end values. *) inductive value_eq (E : embedding) : val → val → Prop ≝ | undef_eq : ∀v. value_eq E Vundef v | vint_eq : ∀sz,i. value_eq E (Vint sz i) (Vint sz i) | vfloat_eq : ∀f. value_eq E (Vfloat f) (Vfloat f) | vnull_eq : value_eq E Vnull Vnull | vptr_eq : ∀p1,p2. pointer_translation p1 E = Some ? p2 → value_eq E (Vptr p1) (Vptr p2). (* [load_sim] states the fact that each successful load in [m1] is matched by a load in [m2] s.t. * the values are equivalent. *) definition load_sim ≝ λ(E : embedding).λ(m1 : mem).λ(m2 : mem). ∀b1,off1,b2,off2,ty,v1. valid_block m1 b1 → (* We need this because it is not provable from [load_value_of_type ty m1 b1 off1] when loading by-ref *) E b1 = Some ? 〈b2,off2〉 → load_value_of_type ty m1 b1 off1 = Some ? v1 → (∃v2. load_value_of_type ty m2 b2 (offset_plus off1 off2) = Some ? v2 ∧ value_eq E v1 v2). definition load_sim_ptr ≝ λ(E : embedding).λ(m1 : mem).λ(m2 : mem). ∀b1,off1,b2,off2,ty,v1. valid_pointer m1 (mk_pointer b1 off1) = true → (* We need this because it is not provable from [load_value_of_type ty m1 b1 off1] when loading by-ref *) pointer_translation (mk_pointer b1 off1) E = Some ? (mk_pointer b2 off2) → load_value_of_type ty m1 b1 off1 = Some ? v1 → (∃v2. load_value_of_type ty m2 b2 off2 = Some ? v2 ∧ value_eq E v1 v2). (* Definition of a memory injection *) record memory_inj (E : embedding) (m1 : mem) (m2 : mem) : Type[0] ≝ { (* Load simulation *) mi_inj : load_sim_ptr E m1 m2; (* Invalid blocks are not mapped *) mi_freeblock : ∀b. ¬ (valid_block m1 b) → E b = None ?; (* Valid pointers are mapped to valid pointers*) mi_valid_pointers : ∀p,p'. valid_pointer m1 p = true → pointer_translation p E = Some ? p' → valid_pointer m2 p' = true; (* Blocks in the codomain are valid *) (* mi_incl : ∀b,b',o'. E b = Some ? 〈b',o'〉 → valid_block m2 b'; *) (* Regions are preserved *) mi_region : ∀b,b',o'. E b = Some ? 〈b',o'〉 → block_region b = block_region b'; (* Disjoint blocks are mapped to disjoint blocks. Note that our condition is stronger than compcert's. * This may cause some problems if we reuse this def for the translation from Clight to Cminor, where * all variables are allocated in the same block. *) mi_disjoint : ∀b1,b2,b1',b2',o1',o2'. b1 ≠ b2 → E b1 = Some ? 〈b1',o1'〉 → E b2 = Some ? 〈b2',o2'〉 → b1' ≠ b2' }. (* Definition of a memory extension. /!\ Not equivalent to the compcert concept. /!\ * A memory extension is a [memory_inj] with some particular blocks designated as * being writeable. *) alias id "meml" = "cic:/matita/basics/lists/list/mem.fix(0,2,1)". record memory_ext (E : embedding) (m1 : mem) (m2 : mem) : Type[0] ≝ { me_inj : memory_inj E m1 m2; (* A list of blocks where we can write freely *) me_writeable : list block; (* These blocks are valid *) me_writeable_valid : ∀b. meml ? b me_writeable → valid_block m2 b; (* And pointers to m1 are oblivious to these blocks *) me_writeable_ok : ∀p,p'. valid_pointer m1 p = true → pointer_translation p E = Some ? p' → ¬ (meml ? (pblock p') me_writeable) }. (* ---------------------------------------------------------------------------- *) (* End of the definitions on memory injections. Now, on to proving some lemmas. *) (* A particular inversion. *) lemma value_eq_ptr_inversion : ∀E,p1,v. value_eq E (Vptr p1) v → ∃p2. v = Vptr p2 ∧ pointer_translation p1 E = Some ? p2. #E #p1 #v #Heq inversion Heq [ 1: #v #Habsurd destruct (Habsurd) | 2: #sz #i #Habsurd destruct (Habsurd) | 3: #f #Habsurd destruct (Habsurd) | 4: #Habsurd destruct (Habsurd) | 5: #p1' #p2 #Heq #Heqv #Heqv2 #_ destruct %{p2} @conj try @refl try assumption ] qed. (* A cleaner version of inversion for [value_eq] *) lemma value_eq_inversion : ∀E,v1,v2. ∀P : val → val → Prop. value_eq E v1 v2 → (∀v. P Vundef v) → (∀sz,i. P (Vint sz i) (Vint sz i)) → (∀f. P (Vfloat f) (Vfloat f)) → (P Vnull Vnull) → (∀p1,p2. pointer_translation p1 E = Some ? p2 → P (Vptr p1) (Vptr p2)) → P v1 v2. #E #v1 #v2 #P #Heq #H1 #H2 #H3 #H4 #H5 inversion Heq [ 1: #v #Hv1 #Hv2 #_ destruct @H1 | 2: #sz #i #Hv1 #Hv2 #_ destruct @H2 | 3: #f #Hv1 #Hv2 #_ destruct @H3 | 4: #Hv1 #Hv2 #_ destruct @H4 | 5: #p1 #p2 #Hembed #Hv1 #Hv2 #_ destruct @H5 // ] qed. (* If we succeed to load something by value from a 〈b,off〉 location, [b] is a valid block. *) lemma load_by_value_success_valid_block_aux : ∀ty,m,b,off,v. access_mode ty = By_value (typ_of_type ty) → load_value_of_type ty m b off = Some ? v → Zltb (block_id b) (nextblock m) = true. #ty #m * #brg #bid #off #v cases ty [ 1: | 2: #sz #sg | 3: #fsz | 4: #ptr_ty | 5: #array_ty #array_sz | 6: #domain #codomain | 7: #structname #fieldspec | 8: #unionname #fieldspec | 9: #id ] whd in match (load_value_of_type ????); [ 1,7,8: #_ #Habsurd destruct (Habsurd) ] #Hmode [ 1,2,3,6: [ 1: elim sz | 2: elim fsz ] normalize in match (typesize ?); whd in match (loadn ???); whd in match (beloadv ??); cases (Zltb bid (nextblock m)) normalize nodelta try // #Habsurd destruct (Habsurd) | *: normalize in Hmode; destruct (Hmode) ] qed. (* If we succeed in loading some data from a location, then this loc is a valid pointer. *) lemma load_by_value_success_valid_ptr_aux : ∀ty,m,b,off,v. access_mode ty = By_value (typ_of_type ty) → load_value_of_type ty m b off = Some ? v → Zltb (block_id b) (nextblock m) = true ∧ Zleb (low_bound m b) (Z_of_unsigned_bitvector ? (offv off)) = true ∧ Zltb (Z_of_unsigned_bitvector ? (offv off)) (high_bound m b) = true. #ty #m * #brg #bid #off #v cases ty [ 1: | 2: #sz #sg | 3: #fsz | 4: #ptr_ty | 5: #array_ty #array_sz | 6: #domain #codomain | 7: #structname #fieldspec | 8: #unionname #fieldspec | 9: #id ] whd in match (load_value_of_type ????); [ 1,7,8: #_ #Habsurd destruct (Habsurd) ] #Hmode [ 1,2,3,6: [ 1: elim sz | 2: elim fsz ] normalize in match (typesize ?); whd in match (loadn ???); whd in match (beloadv ??); cases (Zltb bid (nextblock m)) normalize nodelta cases (Zleb (low (blocks m (mk_block brg bid))) (Z_of_unsigned_bitvector offset_size (offv off))) cases (Zltb (Z_of_unsigned_bitvector offset_size (offv off)) (high (blocks m (mk_block brg bid)))) normalize nodelta #Heq destruct (Heq) try /3 by conj, refl/ | *: normalize in Hmode; destruct (Hmode) ] qed. lemma valid_block_from_bool : ∀b,m. Zltb (block_id b) (nextblock m) = true → valid_block m b. * #rg #id #m normalize elim id /2/ qed. lemma valid_block_to_bool : ∀b,m. valid_block m b → Zltb (block_id b) (nextblock m) = true. * #rg #id #m normalize elim id /2/ qed. lemma load_by_value_success_valid_block : ∀ty,m,b,off,v. access_mode ty = By_value (typ_of_type ty) → load_value_of_type ty m b off = Some ? v → valid_block m b. #ty #m #b #off #v #Haccess_mode #Hload @valid_block_from_bool elim (load_by_value_success_valid_ptr_aux ty m b off v Haccess_mode Hload) * // qed. lemma load_by_value_success_valid_pointer : ∀ty,m,b,off,v. access_mode ty = By_value (typ_of_type ty) → load_value_of_type ty m b off = Some ? v → valid_pointer m (mk_pointer b off). #ty #m #b #off #v #Haccess_mode #Hload normalize elim (load_by_value_success_valid_ptr_aux ty m b off v Haccess_mode Hload) * #H1 #H2 #H3 >H1 >H2 normalize nodelta >Zle_to_Zleb_true try // lapply (Zlt_to_Zle … (Zltb_true_to_Zlt … H3)) /2/ qed. (* Making explicit the contents of memory_inj for load_value_of_type *) lemma load_value_of_type_inj : ∀E,m1,m2,b1,off1,v1,b2,off2,ty. memory_inj E m1 m2 → value_eq E (Vptr (mk_pointer b1 off1)) (Vptr (mk_pointer b2 off2)) → load_value_of_type ty m1 b1 off1 = Some ? v1 → ∃v2. load_value_of_type ty m2 b2 off2 = Some ? v2 ∧ value_eq E v1 v2. #E #m1 #m2 #b1 #off1 #v1 #b2 #off2 #ty #Hinj #Hvalue_eq lapply (value_eq_ptr_inversion … Hvalue_eq) * #p2 * #Hp2_eq #Hembed destruct lapply (refl ? (access_mode ty)) cases ty [ 1: | 2: #sz #sg | 3: #fsz | 4: #ptr_ty | 5: #array_ty #array_sz | 6: #domain #codomain | 7: #structname #fieldspec | 8: #unionname #fieldspec | 9: #id ] normalize in ⊢ ((???%) → ?); #Hmode #Hyp [ 1,7,8: normalize in Hyp; destruct (Hyp) | 5,6: normalize in Hyp ⊢ %; destruct (Hyp) /3 by ex_intro, conj, vptr_eq/ ] lapply (load_by_value_success_valid_pointer … Hmode … Hyp) #Hvalid_pointer lapply (mi_inj … Hinj b1 off1 b2 off2 … Hvalid_pointer Hembed Hyp) #H @H qed. (* -------------------------------------- *) (* Lemmas pertaining to memory allocation *) (* A valid block stays valid after an alloc. *) lemma alloc_valid_block_conservation : ∀m,m',z1,z2,r,new_block. alloc m z1 z2 r = 〈m', new_block〉 → ∀b. valid_block m b → valid_block m' b. #m #m' #z1 #z2 #r * #b' #Hregion_eq elim m #contents #nextblock #Hcorrect whd in match (alloc ????); #Halloc destruct (Halloc) #b whd in match (valid_block ??) in ⊢ (% → %); /2/ qed. (* Allocating a new zone produces a valid block. *) lemma alloc_valid_new_block : ∀m,m',z1,z2,r,new_block. alloc m z1 z2 r = 〈m', new_block〉 → valid_block m' new_block. * #contents #nextblock #Hcorrect #m' #z1 #z2 #r * #new_block #Hregion_eq whd in match (alloc ????); whd in match (valid_block ??); #Halloc destruct (Halloc) /2/ qed. (* All data in a valid block is unchanged after an alloc. *) lemma alloc_beloadv_conservation : ∀m,m',block,offset,z1,z2,r,new_block. valid_block m block → alloc m z1 z2 r = 〈m', new_block〉 → beloadv m (mk_pointer block offset) = beloadv m' (mk_pointer block offset). * #contents #next #Hcorrect #m' #block #offset #z1 #z2 #r #new_block #Hvalid #Halloc whd in Halloc:(??%?); destruct (Halloc) whd in match (beloadv ??) in ⊢ (??%%); lapply (valid_block_to_bool … Hvalid) #Hlt >Hlt >(zlt_succ … Hlt) normalize nodelta whd in match (update_block ?????); whd in match (eq_block ??); cut (eqZb (block_id block) next = false) [ lapply (Zltb_true_to_Zlt … Hlt) #Hlt' @eqZb_false /2/ ] #Hneq >Hneq cases (eq_region ??) normalize nodelta @refl qed. (* Lift [alloc_beloadv_conservation] to loadn *) lemma alloc_loadn_conservation : ∀m,m',z1,z2,r,new_block. alloc m z1 z2 r = 〈m', new_block〉 → ∀n. ∀block,offset. valid_block m block → loadn m (mk_pointer block offset) n = loadn m' (mk_pointer block offset) n. #m #m' #z1 #z2 #r #new_block #Halloc #n elim n try // #n' #Hind #block #offset #Hvalid_block whd in ⊢ (??%%); >(alloc_beloadv_conservation … Hvalid_block Halloc) cases (beloadv m' (mk_pointer block offset)) // #bev normalize nodelta whd in match (shift_pointer ???); >Hind try // qed. (* Memory allocation preserves [memory_inj] *) lemma alloc_memory_inj : ∀E:embedding.∀m1,m2,m2',z1,z2,r,new_block. ∀H:memory_inj E m1 m2. alloc m2 z1 z2 r = 〈m2', new_block〉 → memory_inj E m1 m2'. #E #m1 #m2 #m2' #z1 #z2 #r * #new_block #Hregion #Hmemory_inj #Halloc % [ 1: whd #b1 #off1 #b2 #off2 #ty #v1 #Hvalid #Hembed #Hload elim (mi_inj E m1 m2 Hmemory_inj b1 off1 b2 off2 … ty v1 Hvalid Hembed Hload) #v2 * #Hload_eq #Hvalue_eq %{v2} @conj try // lapply (refl ? (access_mode ty)) cases ty in Hload_eq; [ 1: | 2: #sz #sg | 3: #fsz | 4: #ptr_ty | 5: #array_ty #array_sz | 6: #domain #codomain | 7: #structname #fieldspec | 8: #unionname #fieldspec | 9: #id ] #Hload normalize in ⊢ ((???%) → ?); #Haccess [ 1,7,8: normalize in Hload; destruct (Hload) | 2,3,4,9: whd in match (load_value_of_type ????); whd in match (load_value_of_type ????); lapply (load_by_value_success_valid_block … Haccess Hload) #Hvalid_block whd in match (load_value_of_type ????) in Hload; <(alloc_loadn_conservation … Halloc … Hvalid_block) @Hload | 5,6: whd in match (load_value_of_type ????) in Hload ⊢ %; @Hload ] | 2: @(mi_freeblock … Hmemory_inj) | 3: #p #p' #Hvalid #Hptr_trans lapply (mi_valid_pointers … Hmemory_inj p p' Hvalid Hptr_trans) elim m2 in Halloc; #contentmap #nextblock #Hnextblock elim p' * #br' #bid' #o' #Halloc whd in Halloc:(??%?) ⊢ ?; destruct (Halloc) whd in match (valid_pointer ??) in ⊢ (% → %); @Zltb_elim_Type0 [ 2: normalize #_ #Habsurd destruct (Habsurd) ] #Hbid' cut (Zltb bid' (Zsucc nextblock) = true) [ lapply (Zlt_to_Zltb_true … Hbid') @zlt_succ ] #Hbid_succ >Hbid_succ whd in match (low_bound ??) in ⊢ (% → %); whd in match (high_bound ??) in ⊢ (% → %); whd in match (update_block ?????); whd in match (eq_block ??); cut (eqZb bid' nextblock = false) [ 1: @eqZb_false /2 by not_to_not/ ] #Hbid_neq >Hbid_neq cases (eq_region br' r) normalize #H @H | 4: @(mi_region … Hmemory_inj) | 5: @(mi_disjoint … Hmemory_inj) ] qed. (* Memory allocation induces a memory extension. *) lemma alloc_memory_ext : ∀E:embedding.∀m1,m2,m2',z1,z2,r,new_block. ∀H:memory_inj E m1 m2. alloc m2 z1 z2 r = 〈m2', new_block〉 → memory_ext E m1 m2'. #E #m1 #m2 #m2' #z1 #z2 #r * #new_block #Hblock_region_eq #Hmemory_inj #Halloc lapply (alloc_memory_inj … Hmemory_inj Halloc) #Hinj' % [ 1: @Hinj' | 2: @[new_block] | 3: #b normalize in ⊢ (%→ ?); * [ 2: #H @(False_ind … H) ] #Heq destruct (Heq) whd elim m2 in Halloc; #Hcontents #nextblock #Hnextblock whd in ⊢ ((??%?) → ?); #Heq destruct (Heq) /2/ | 4: * #b #o * #b' #o' #Hvalid_ptr #Hembed % normalize in ⊢ (% → ?); * [ 2: #H @H ] #Heq destruct (Heq) lapply (mi_valid_pointers … Hmemory_inj … Hvalid_ptr Hembed) whd in ⊢ (% → ?); (* contradiction because ¬ (valid_block m2 new_block) *) elim m2 in Halloc; #contents_m2 #nextblock_m2 #Hnextblock_m2 whd in ⊢ ((??%?) → ?); #Heq_alloc destruct (Heq_alloc) normalize lapply (irreflexive_Zlt nextblock_m2) @Zltb_elim_Type0 [ #H * #Habsurd @(False_ind … (Habsurd H)) ] #_ #_ normalize #Habsurd destruct (Habsurd) ] qed. lemma bestorev_beloadv_conservation : ∀mA,mB,wb,wo,v. bestorev mA (mk_pointer wb wo) v = Some ? mB → ∀rb,ro. eq_block wb rb = false → beloadv mA (mk_pointer rb ro) = beloadv mB (mk_pointer rb ro). #mA #mB #wb #wo #v #Hstore #rb #ro #Hblock_neq whd in ⊢ (??%%); elim mB in Hstore; #contentsB #nextblockB #HnextblockB normalize in ⊢ (% → ?); cases (Zltb (block_id wb) (nextblock mA)) normalize nodelta [ 2: #Habsurd destruct (Habsurd) ] cases (if Zleb (low (blocks mA wb)) (Z_of_unsigned_bitvector 16 (offv wo)) then Zltb (Z_of_unsigned_bitvector 16 (offv wo)) (high (blocks mA wb))  else false) normalize nodelta [ 2: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) elim rb in Hblock_neq; #rr #rid elim wb #wr #wid cases rr cases wr normalize try // @(eqZb_elim … rid wid) [ 1,3: #Heq destruct (Heq) >(eqZb_reflexive wid) #Habsurd destruct (Habsurd) ] try // qed. (* lift [bestorev_beloadv_conservation to [loadn] *) lemma bestorev_loadn_conservation : ∀mA,mB,n,wb,wo,v. bestorev mA (mk_pointer wb wo) v = Some ? mB → ∀rb,ro. eq_block wb rb = false → loadn mA (mk_pointer rb ro) n = loadn mB (mk_pointer rb ro) n. #mA #mB #n elim n [ 1: #wb #wo #v #Hstore #rb #ro #Hblock_neq normalize @refl | 2: #n' #Hind #wb #wo #v #Hstore #rb #ro #Hneq whd in ⊢ (??%%); >(bestorev_beloadv_conservation … Hstore … Hneq) >(Hind … Hstore … Hneq) @refl ] qed. (* lift [bestorev_loadn_conservation] to [load_value_of_type] *) lemma bestorev_load_value_of_type_conservation : ∀mA,mB,wb,wo,v. bestorev mA (mk_pointer wb wo) v = Some ? mB → ∀rb,ro. eq_block wb rb = false → ∀ty.load_value_of_type ty mA rb ro = load_value_of_type ty mB rb ro. #mA #mB #wb #wo #v #Hstore #rb #ro #Hneq #ty cases ty [ 1: | 2: #sz #sg | 3: #fsz | 4: #ptr_ty | 5: #array_ty #array_sz | 6: #domain #codomain | 7: #structname #fieldspec | 8: #unionname #fieldspec | 9: #id ] try // [ 1: elim sz | 2: elim fsz | 3: | 4: ] whd in ⊢ (??%%); >(bestorev_loadn_conservation … Hstore … Hneq) @refl qed. (* Writing in the "extended" part of a memory preserves the underlying injection *) lemma bestorev_memory_ext_to_load_sim : ∀E,mA,mB,mC. ∀Hext:memory_ext E mA mB. ∀wb,wo,v. meml ? wb (me_writeable … Hext) → bestorev mB (mk_pointer wb wo) v = Some ? mC → load_sim_ptr E mA mC. #E #mA #mB #mC #Hext #wb #wo #v #Hwb #Hstore whd #b1 #off1 #b2 #off2 #ty #v1 #Hvalid #Hembed #Hload (* Show that (wb ≠ b2) by showing b2 ∉ (me_writeable ...) *) lapply (me_writeable_ok … Hext (mk_pointer b1 off1) (mk_pointer b2 off2) Hvalid Hembed) #Hb2 lapply (mem_not_mem_neq … Hwb Hb2) #Hwb_neq_b2 cut ((eq_block wb b2) = false) [ @neq_block_eq_block_false @Hwb_neq_b2 ] #Heq_block_false <(bestorev_load_value_of_type_conservation … Hstore … Heq_block_false) @(mi_inj … (me_inj … Hext) … Hvalid … Hembed … Hload) qed. (* Writing in the "extended" part of a memory preserves the whole memory injection *) lemma bestorev_memory_ext_to_memory_inj : ∀E,mA,mB,mC. ∀Hext:memory_ext E mA mB. ∀wb,wo,v. meml ? wb (me_writeable … Hext) → bestorev mB (mk_pointer wb wo) v = Some ? mC → memory_inj E mA mC. #E #mA * #contentsB #nextblockB #HnextblockB #mC #Hext #wb #wo #v #Hwb #Hstore % [ 1: @(bestorev_memory_ext_to_load_sim … Hext … Hwb Hstore) ] elim Hext in Hwb; * #Hinj #Hvalid #Hcodomain #Hregion #Hdisjoint #writeable #Hwriteable_valid #Hwriteable_ok #Hmem [ 1: @Hvalid | 3: @Hregion | 4: @Hdisjoint ] -Hvalid -Hregion -Hdisjoint whd in Hstore:(??%?); lapply (Hwriteable_valid … Hmem) normalize in ⊢ (% → ?); #Hlt_wb #p #p' #HvalidA #Hembed lapply (Hcodomain … HvalidA Hembed) -Hcodomain normalize in match (valid_pointer ??) in ⊢ (% → %); >(Zlt_to_Zltb_true … Hlt_wb) in Hstore; normalize nodelta cases (Zleb (low (contentsB wb)) (Z_of_unsigned_bitvector offset_size (offv wo)) ∧Zltb (Z_of_unsigned_bitvector offset_size (offv wo)) (high (contentsB wb))) normalize nodelta [ 2: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) cases (Zltb (block_id (pblock p')) nextblockB) normalize nodelta [ 2: #H @H ] whd in match (update_block ?????); cut (eq_block (pblock p') wb = false) [ 2: #Heq >Heq normalize nodelta #H @H ] @neq_block_eq_block_false @sym_neq @(mem_not_mem_neq writeable … Hmem) @(Hwriteable_ok … HvalidA Hembed) qed. (* It even preserves memory extensions, with the same writeable blocks. *) lemma bestorev_memory_ext_to_memory_ext : ∀E,mA,mB. ∀Hext:memory_ext E mA mB. ∀wb,wo,v. meml ? wb (me_writeable … Hext) → ∀mC.bestorev mB (mk_pointer wb wo) v = Some ? mC → ΣExt:memory_ext E mA mC.(me_writeable … Hext = me_writeable … Ext). #E #mA #mB #Hext #wb #wo #v #Hmem #mC #Hstore %{(mk_memory_ext … (bestorev_memory_ext_to_memory_inj … Hext … Hmem … Hstore) (me_writeable … Hext) ? (me_writeable_ok … Hext) )} try @refl #b #Hmemb lapply (me_writeable_valid … Hext b Hmemb) lapply (me_writeable_valid … Hext wb Hmem) elim mB in Hstore; #contentsB #nextblockB #HnextblockB #Hstore #Hwb_valid #Hb_valid lapply Hstore normalize in Hwb_valid Hb_valid ⊢ (% → ?); >(Zlt_to_Zltb_true … Hwb_valid) normalize nodelta cases (if Zleb (low (contentsB wb)) (Z_of_unsigned_bitvector 16 (offv wo)) then Zltb (Z_of_unsigned_bitvector 16 (offv wo)) (high (contentsB wb))  else false) normalize [ 2: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) @Hb_valid qed. (* Lift [bestorev_memory_ext_to_memory_ext] to storen *) lemma storen_memory_ext_to_memory_ext : ∀E,mA,l,mB,mC. ∀Hext:memory_ext E mA mB. ∀wb,wo. meml ? wb (me_writeable … Hext) → storen mB (mk_pointer wb wo) l = Some ? mC → memory_ext E mA mC. #E #mA #l elim l [ 1: #mB #mC #Hext #wb #wo #Hmem normalize in ⊢ (% → ?); #Heq destruct (Heq) @Hext | 2: #hd #tl #Hind #mB #mC #Hext #wb #wo #Hmem whd in ⊢ ((??%?) → ?); lapply (bestorev_memory_ext_to_memory_ext … Hext … wb wo hd Hmem) cases (bestorev mB (mk_pointer wb wo) hd) normalize nodelta [ 1: #H #Habsurd destruct (Habsurd) ] #mD #H lapply (H mD (refl ??)) * #HextD #Heq #Hstore @(Hind … HextD … Hstore) commutative_addition_n >associative_addition_n <(commutative_addition_n … (offv offset) (offv o1')) @refl ] #Heq >Heq @refl ] ] | 3: whd in match (sem_add ????); whd in match (sem_add ????); normalize nodelta cases (classify_add ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) /3 by ex_intro, conj, vfloat_eq/ | 4: whd in match (sem_add ????); whd in match (sem_add ????); normalize nodelta cases (classify_add ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] @eq_bv_elim [ 1: normalize nodelta #Heq1 #Heq2 destruct /3 by ex_intro, conj, vnull_eq/ | 2: #_ normalize nodelta #Habsurd destruct (Habsurd) ] | 5: whd in match (sem_add ????); whd in match (sem_add ????); normalize nodelta cases (classify_add ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) %{(Vptr (shift_pointer_n (bitsize_of_intsize sz') p2 (sizeof ty) i'))} @conj try @refl @vptr_eq whd in match (pointer_translation ??) in Hembed ⊢ %; elim p1 in Hembed; #b1 #o1 normalize nodelta cases (E b1) [ 1: normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) | 2: * #block #offset normalize nodelta #Heq destruct (Heq) whd in match (shift_pointer_n ????); whd in match (shift_offset_n ????) in ⊢ (??%%); whd in match (offset_plus ??); whd in match (offset_plus ??); >commutative_addition_n >(associative_addition_n … offset_size ?) >(commutative_addition_n ? (offv offset) ?) @refl ] ] qed. lemma subtraction_delta : ∀x,y,delta. subtraction offset_size (addition_n offset_size x delta) (addition_n offset_size y delta) = subtraction offset_size x y. #x #y #delta whd in match subtraction; normalize nodelta (* Remove all the equal parts on each side of the equation. *) two_complement_negation_plus <(commutative_addition_n … (two_complement_negation ? delta)) >(associative_addition_n ? delta) >bitvector_opp_addition_n >(commutative_addition_n ? (zero ?)) >neutral_addition_n @refl qed. (* Offset subtraction is invariant by translation *) lemma sub_offset_translation : ∀n,x,y,delta. sub_offset n x y = sub_offset n (offset_plus x delta) (offset_plus y delta). #n #x #y #delta whd in match (sub_offset ???) in ⊢ (??%%); elim x #x' elim y #y' elim delta #delta' whd in match (offset_plus ??); whd in match (offset_plus ??); >subtraction_delta @refl qed. (* value_eq lifts to subtraction *) lemma sub_value_eq : ∀E,v1,v2,v1',v2',ty1,ty2. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_sub v1 ty1 v1' ty2=Some val r1→ ∃r2:val.sem_sub v2 ty1 v2' ty2=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #ty1 #ty2 #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 1: #v | 2: #sz #i | 3: #f | 4: | 5: #p1 #p2 #Hembed ] [ 1: whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #n #ty #sz #sg | 4: #n #sz #sg #ty | 5: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 2: whd in match (sem_sub ????); whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 2,3,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1,6: #v' | 2,7: #sz' #i' | 3,8: #f' | 4,9: | 5,10: #p1' #p2' #Hembed' ] [ 1,2,4,5,6,7,8,9,10: #Habsurd destruct (Habsurd) ] @intsize_eq_elim_elim [ 1: #_ #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) normalize nodelta #Heq destruct (Heq) /3 by ex_intro, conj, vint_eq/ ] | 3: whd in match (sem_sub ????); whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) /3 by ex_intro, conj, vfloat_eq/ | 4: whd in match (sem_sub ????); whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 1,2,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1,6: #v' | 2,7: #sz' #i' | 3,8: #f' | 4,9: | 5,10: #p1' #p2' #Hembed' ] [ 1,2,4,5,6,7,9,10: #Habsurd destruct (Habsurd) ] [ 1: @eq_bv_elim [ 1: normalize nodelta #Heq1 #Heq2 destruct /3 by ex_intro, conj, vnull_eq/ | 2: #_ normalize nodelta #Habsurd destruct (Habsurd) ] | 2: #Heq destruct (Heq) /3 by ex_intro, conj, vnull_eq/ ] | 5: whd in match (sem_sub ????); whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tfsz | 3: #tn #ty #tsz #tsg | 4: #tn #tsz #tsg #ty | 5: #ty1' #ty2' ] [ 1,2,5: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1,6: #v' | 2,7: #sz' #i' | 3,8: #f' | 4,9: | 5,10: #p1' #p2' #Hembed' ] [ 1,2,4,5,6,7,8,9: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) [ 1: %{(Vptr (neg_shift_pointer_n (bitsize_of_intsize sz') p2 (sizeof ty) i'))} @conj try @refl @vptr_eq whd in match (pointer_translation ??) in Hembed ⊢ %; elim p1 in Hembed; #b1 #o1 normalize nodelta cases (E b1) [ 1: normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) | 2: * #block #offset normalize nodelta #Heq destruct (Heq) whd in match (offset_plus ??) in ⊢ (??%%); whd in match (neg_shift_pointer_n ????) in ⊢ (??%%); whd in match (neg_shift_offset_n ????) in ⊢ (??%%); whd in match (subtraction) in ⊢ (??%%); normalize nodelta generalize in match (short_multiplication ???); #mult /3 by associative_addition_n, commutative_addition_n, refl/ ] | 2: lapply Heq @eq_block_elim [ 2: #_ normalize nodelta #Habsurd destruct (Habsurd) | 1: #Hpblock1_eq normalize nodelta elim p1 in Hpblock1_eq Hembed Hembed'; #b1 #off1 elim p1' #b1' #off1' whd in ⊢ (% → % → ?); #Hpblock1_eq destruct (Hpblock1_eq) whd in ⊢ ((??%?) → (??%?) → ?); cases (E b1') normalize nodelta [ 1: #Habsurd destruct (Habsurd) ] * #dest_block #dest_off normalize nodelta #Heq_ptr1 #Heq_ptr2 destruct >eq_block_b_b normalize nodelta cases (eqb (sizeof tsg) O) normalize nodelta [ 1: #Habsurd destruct (Habsurd) | 2: >(sub_offset_translation 32 off1 off1' dest_off) cases (division_u 31 (sub_offset 32 (offset_plus off1 dest_off) (offset_plus off1' dest_off)) (repr (sizeof tsg))) [ 1: normalize nodelta #Habsurd destruct (Habsurd) | 2: #r1' normalize nodelta #Heq2 destruct (Heq2) /3 by ex_intro, conj, vint_eq/ ] ] ] ] ] qed. lemma mul_value_eq : ∀E,v1,v2,v1',v2',ty1,ty2. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_mul v1 ty1 v1' ty2=Some val r1→ ∃r2:val.sem_mul v2 ty1 v2' ty2=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #ty1 #ty2 #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 1: #v | 2: #sz #i | 3: #f | 4: | 5: #p1 #p2 #Hembed ] [ 1: whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 2: whd in match (sem_mul ????); whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] [ 2,3: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] @intsize_eq_elim_elim [ 1: #_ #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) normalize nodelta #Heq destruct (Heq) /3 by ex_intro, conj, vint_eq/ ] | 3: whd in match (sem_mul ????); whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] [ 1,3: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) /3 by ex_intro, conj, vfloat_eq/ | 4: whd in match (sem_mul ????); whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 5: whd in match (sem_mul ????); whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) ] qed. lemma div_value_eq : ∀E,v1,v2,v1',v2',ty1,ty2. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_div v1 ty1 v1' ty2=Some val r1→ ∃r2:val.sem_div v2 ty1 v2' ty2=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #ty1 #ty2 #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 1: #v | 2: #sz #i | 3: #f | 4: | 5: #p1 #p2 #Hembed ] [ 1: whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 2: whd in match (sem_div ????); whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] [ 2,3: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] elim sg normalize nodelta @intsize_eq_elim_elim [ 1,3: #_ #Habsurd destruct (Habsurd) | 2,4: #Heq destruct (Heq) normalize nodelta @(match (division_s (bitsize_of_intsize sz') i i') with [ None ⇒ ? | Some bv' ⇒ ? ]) [ 1: normalize #Habsurd destruct (Habsurd) | 2: normalize #Heq destruct (Heq) /3 by ex_intro, conj, vint_eq/ | 3,4: elim sz' in i' i; #i' #i normalize in match (pred_size_intsize ?); generalize in match division_u; #division_u normalize @(match (division_u ???) with [ None ⇒ ? | Some bv ⇒ ?]) normalize nodelta #H destruct (H) /3 by ex_intro, conj, vint_eq/ ] ] | 3: whd in match (sem_div ????); whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] [ 1,3: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) /3 by ex_intro, conj, vfloat_eq/ | 4: whd in match (sem_div ????); whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 5: whd in match (sem_div ????); whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) ] qed. lemma mod_value_eq : ∀E,v1,v2,v1',v2',ty1,ty2. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_mod v1 ty1 v1' ty2=Some val r1→ ∃r2:val.sem_mod v2 ty1 v2' ty2=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #ty1 #ty2 #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 1: #v | 2: #sz #i | 3: #f | 4: | 5: #p1 #p2 #Hembed ] [ 1: whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 2: whd in match (sem_mod ????); whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] [ 2,3: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] elim sg normalize nodelta @intsize_eq_elim_elim [ 1,3: #_ #Habsurd destruct (Habsurd) | 2,4: #Heq destruct (Heq) normalize nodelta @(match (modulus_s (bitsize_of_intsize sz') i i') with [ None ⇒ ? | Some bv' ⇒ ? ]) [ 1: normalize #Habsurd destruct (Habsurd) | 2: normalize #Heq destruct (Heq) /3 by ex_intro, conj, vint_eq/ | 3,4: elim sz' in i' i; #i' #i generalize in match modulus_u; #modulus_u normalize @(match (modulus_u ???) with [ None ⇒ ? | Some bv ⇒ ?]) normalize nodelta #H destruct (H) /3 by ex_intro, conj, vint_eq/ ] ] | 3: whd in match (sem_mod ????); whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 4: whd in match (sem_mod ????); whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 5: whd in match (sem_mod ????); whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) ] qed. (* boolean ops *) lemma and_value_eq : ∀E,v1,v2,v1',v2'. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_and v1 v1'=Some val r1→ ∃r2:val.sem_and v2 v2'=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 2: #sz #i @(value_eq_inversion E … Hvalue_eq2) [ 2: #sz' #i' whd in match (sem_and ??); @intsize_eq_elim_elim [ 1: #_ #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) normalize nodelta #Heq destruct (Heq) /3 by ex_intro,conj,vint_eq/ ] ] ] normalize in match (sem_and ??); #arg1 destruct normalize in match (sem_and ??); #arg2 destruct normalize in match (sem_and ??); #arg3 destruct normalize in match (sem_and ??); #arg4 destruct qed. lemma or_value_eq : ∀E,v1,v2,v1',v2'. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_or v1 v1'=Some val r1→ ∃r2:val.sem_or v2 v2'=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 2: #sz #i @(value_eq_inversion E … Hvalue_eq2) [ 2: #sz' #i' whd in match (sem_or ??); @intsize_eq_elim_elim [ 1: #_ #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) normalize nodelta #Heq destruct (Heq) /3 by ex_intro,conj,vint_eq/ ] ] ] normalize in match (sem_or ??); #arg1 destruct normalize in match (sem_or ??); #arg2 destruct normalize in match (sem_or ??); #arg3 destruct normalize in match (sem_or ??); #arg4 destruct qed. lemma xor_value_eq : ∀E,v1,v2,v1',v2'. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_xor v1 v1'=Some val r1→ ∃r2:val.sem_xor v2 v2'=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 2: #sz #i @(value_eq_inversion E … Hvalue_eq2) [ 2: #sz' #i' whd in match (sem_xor ??); @intsize_eq_elim_elim [ 1: #_ #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) normalize nodelta #Heq destruct (Heq) /3 by ex_intro,conj,vint_eq/ ] ] ] normalize in match (sem_xor ??); #arg1 destruct normalize in match (sem_xor ??); #arg2 destruct normalize in match (sem_xor ??); #arg3 destruct normalize in match (sem_xor ??); #arg4 destruct qed. lemma shl_value_eq : ∀E,v1,v2,v1',v2'. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_shl v1 v1'=Some val r1→ ∃r2:val.sem_shl v2 v2'=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 1: #v | 2: #sz #i | 3: #f | 4: | 5: #p1 #p2 #Hembed ] [ 2: @(value_eq_inversion E … Hvalue_eq2) [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 2: whd in match (sem_shl ??); cases (lt_u ???) normalize nodelta [ 1: #Heq destruct (Heq) /3 by ex_intro,conj,vint_eq/ | 2: #Habsurd destruct (Habsurd) ] | *: whd in match (sem_shl ??); #Habsurd destruct (Habsurd) ] | *: whd in match (sem_shl ??); #Habsurd destruct (Habsurd) ] qed. lemma shr_value_eq : ∀E,v1,v2,v1',v2',ty,ty'. value_eq E v1 v2 → value_eq E v1' v2' → ∀r1. (sem_shr v1 ty v1' ty'=Some val r1→ ∃r2:val.sem_shr v2 ty v2' ty'=Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #ty #ty' #Hvalue_eq1 #Hvalue_eq2 #r1 @(value_eq_inversion E … Hvalue_eq1) [ 1: #v | 2: #sz #i | 3: #f | 4: | 5: #p1 #p2 #Hembed ] whd in match (sem_shr ????); whd in match (sem_shr ????); [ 1: cases (classify_aop ty ty') normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 2: cases (classify_aop ty ty') normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] [ 2,3: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v' | 2: #sz' #i' | 3: #f' | 4: | 5: #p1' #p2' #Hembed' ] [ 1,3,4,5: #Habsurd destruct (Habsurd) ] elim sg normalize nodelta cases (lt_u ???) normalize nodelta #Heq destruct (Heq) /3 by ex_intro, conj, refl, vint_eq/ | 3: cases (classify_aop ty ty') normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 4: cases (classify_aop ty ty') normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 5: cases (classify_aop ty ty') normalize nodelta [ 1: #sz #sg | 2: #fsz | 3: #ty1' #ty2' ] #Habsurd destruct (Habsurd) ] qed. lemma eq_offset_translation : ∀delta,x,y. cmp_offset Ceq (offset_plus x delta) (offset_plus y delta) = cmp_offset Ceq x y. * #delta * #x * #y whd in match (offset_plus ??); whd in match (offset_plus ??); whd in match cmp_offset; normalize nodelta whd in match eq_offset; normalize nodelta @(eq_bv_elim … x y) [ 1: #Heq >Heq >eq_bv_true @refl | 2: #Hneq lapply (injective_inv_addition_n … x y delta Hneq) @(eq_bv_elim … (addition_n offset_size x delta) (addition_n offset_size y delta)) [ 1: #H * #H' @(False_ind … (H' H)) | 2: #_ #_ @refl ] ] qed. lemma neq_offset_translation : ∀delta,x,y. cmp_offset Cne (offset_plus x delta) (offset_plus y delta) = cmp_offset Cne x y. * #delta * #x * #y whd in match (offset_plus ??); whd in match (offset_plus ??); whd in match cmp_offset; normalize nodelta whd in match eq_offset; normalize nodelta @(eq_bv_elim … x y) [ 1: #Heq >Heq >eq_bv_true @refl | 2: #Hneq lapply (injective_inv_addition_n … x y delta Hneq) @(eq_bv_elim … (addition_n offset_size x delta) (addition_n offset_size y delta)) [ 1: #H * #H' @(False_ind … (H' H)) | 2: #_ #_ @refl ] ] qed. (* /!\ Here is the source of all sadness (op = lt) /!\ *) axiom cmp_offset_translation : ∀op,delta,x,y. cmp_offset op x y = cmp_offset op (offset_plus x delta) (offset_plus y delta). lemma cmp_value_eq : ∀E,v1,v2,v1',v2',ty,ty',m1,m2. value_eq E v1 v2 → value_eq E v1' v2' → memory_inj E m1 m2 → ∀op,r1. (sem_cmp op v1 ty v1' ty' m1 = Some val r1→ ∃r2:val.sem_cmp op v2 ty v2' ty' m2 = Some val r2∧value_eq E r1 r2). #E #v1 #v2 #v1' #v2' #ty #ty' #m1 #m2 #Hvalue_eq1 #Hvalue_eq2 #Hinj #op #r1 elim m1 in Hinj; #contmap1 #nextblock1 #Hnextblock1 elim m2 #contmap2 #nextblock2 #Hnextblock2 #Hinj whd in match (sem_cmp ??????) in ⊢ ((??%?) → %); cases (classify_cmp ty ty') normalize nodelta [ 1: #tsz #tsg @(value_eq_inversion E … Hvalue_eq1) normalize nodelta [ 1: #v #Habsurd destruct (Habsurd) | 3: #f #Habsurd destruct (Habsurd) | 4: #Habsurd destruct (Habsurd) | 5: #p1 #p2 #Hembed #Habsurd destruct (Habsurd) ] #sz #i @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v #Habsurd destruct (Habsurd) | 3: #f #Habsurd destruct (Habsurd) | 4: #Habsurd destruct (Habsurd) | 5: #p1 #p2 #Hembed #Habsurd destruct (Habsurd) ] #sz' #i' cases tsg normalize nodelta @intsize_eq_elim_elim [ 1,3: #Hneq #Habsurd destruct (Habsurd) | 2,4: #Heq destruct (Heq) normalize nodelta #Heq destruct (Heq) [ 1: cases (cmp_int ????) whd in match (of_bool ?); | 2: cases (cmpu_int ????) whd in match (of_bool ?); ] /3 by ex_intro, conj, vint_eq/ ] | 3: #fsz @(value_eq_inversion E … Hvalue_eq1) normalize nodelta [ 1: #v #Habsurd destruct (Habsurd) | 2: #sz #i #Habsurd destruct (Habsurd) | 4: #Habsurd destruct (Habsurd) | 5: #p1 #p2 #Hembed #Habsurd destruct (Habsurd) ] #f @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: #v #Habsurd destruct (Habsurd) | 2: #sz #i #Habsurd destruct (Habsurd) | 4: #Habsurd destruct (Habsurd) | 5: #p1 #p2 #Hembed #Habsurd destruct (Habsurd) ] #f' #Heq destruct (Heq) cases (Fcmp op f f') /3 by ex_intro, conj, vint_eq/ | 4: #ty1 #ty2 #Habsurd destruct (Habsurd) | 2: #optn #typ @(value_eq_inversion E … Hvalue_eq1) normalize nodelta [ 1: #v #Habsurd destruct (Habsurd) | 2: #sz #i #Habsurd destruct (Habsurd) | 3: #f #Habsurd destruct (Habsurd) | 5: #p1 #p2 #Hembed ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1,6: #v #Habsurd destruct (Habsurd) | 2,7: #sz #i #Habsurd destruct (Habsurd) | 3,8: #f #Habsurd destruct (Habsurd) | 5,10: #p1' #p2' #Hembed' ] [ 2,3: cases op whd in match (sem_cmp_mismatch ?); #Heq destruct (Heq) [ 1,3: %{Vfalse} @conj try @refl @vint_eq | 2,4: %{Vtrue} @conj try @refl @vint_eq ] | 4: cases op whd in match (sem_cmp_match ?); #Heq destruct (Heq) [ 2,4: %{Vfalse} @conj try @refl @vint_eq | 1,3: %{Vtrue} @conj try @refl @vint_eq ] ] lapply (mi_valid_pointers … Hinj p1' p2') lapply (mi_valid_pointers … Hinj p1 p2) cases (valid_pointer (mk_mem ???) p1') [ 2: #_ #_ >commutative_andb normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) ] cases (valid_pointer (mk_mem ???) p1) [ 2: #_ #_ normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) ] #H1 #H2 lapply (H1 (refl ??) Hembed) #Hvalid1 lapply (H2 (refl ??) Hembed') #Hvalid2 >Hvalid1 >Hvalid2 normalize nodelta -H1 -H2 elim p1 in Hembed; #b1 #o1 elim p1' in Hembed'; #b1' #o1' whd in match (pointer_translation ??); whd in match (pointer_translation ??); @(eq_block_elim … b1 b1') [ 1: #Heq destruct (Heq) cases (E b1') normalize nodelta [ 1: #Habsurd destruct (Habsurd) ] * #eb1' #eo1' normalize nodelta #Heq1 #Heq2 #Heq3 destruct >eq_block_b_b normalize nodelta (neq_block_eq_block_false … Hneq_block) normalize nodelta elim op whd in match (sem_cmp_mismatch ?); #Heq destruct (Heq) /3 by ex_intro, conj, vint_eq/ ] ] qed. lemma binary_operation_value_eq : ∀E,op,v1,v2,v1',v2',ty1,ty2,m1,m2. value_eq E v1 v2 → value_eq E v1' v2' → memory_inj E m1 m2 → ∀r1. sem_binary_operation op v1 ty1 v1' ty2 m1 = Some ? r1 → ∃r2.sem_binary_operation op v2 ty1 v2' ty2 m2 = Some ? r2 ∧ value_eq E r1 r2. #E #op #v1 #v2 #v1' #v2' #ty1 #ty2 #m1 #m2 #Hvalue_eq1 #Hvalue_eq2 #Hinj #r1 cases op whd in match (sem_binary_operation ??????); whd in match (sem_binary_operation ??????); [ 1: @add_value_eq try assumption | 2: @sub_value_eq try assumption | 3: @mul_value_eq try assumption | 4: @div_value_eq try assumption | 5: @mod_value_eq try assumption | 6: @and_value_eq try assumption | 7: @or_value_eq try assumption | 8: @xor_value_eq try assumption | 9: @shl_value_eq try assumption | 10: @shr_value_eq try assumption | *: @cmp_value_eq try assumption ] qed. lemma cast_value_eq : ∀E,m1,m2,v1,v2. (* memory_inj E m1 m2 → *) value_eq E v1 v2 → ∀ty,cast_ty,res. exec_cast m1 v1 ty cast_ty = OK ? res → ∃res'. exec_cast m2 v2 ty cast_ty = OK ? res' ∧ value_eq E res res'. #E #m1 #m2 #v1 #v2 (* #Hmemory_inj *) #Hvalue_eq #ty #cast_ty #res @(value_eq_inversion … Hvalue_eq) [ 1: #v normalize #Habsurd destruct (Habsurd) | 2: #vsz #vi whd in match (exec_cast ????); cases ty [ 1: | 2: #sz #sg | 3: #fl | 4: #ptrty | 5: #arrayty #n | 6: #tl #retty | 7: #id #fl | 8: #id #fl | 9: #comptrty ] normalize nodelta [ 1,3,7,8,9: #Habsurd destruct (Habsurd) | 2: @intsize_eq_elim_elim [ 1: #Hneq #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) normalize nodelta cases cast_ty [ 1: | 2: #csz #csg | 3: #cfl | 4: #cptrty | 5: #carrayty #cn | 6: #ctl #cretty | 7: #cid #cfl | 8: #cid #cfl | 9: #ccomptrty ] normalize nodelta [ 1,7,8,9: #Habsurd destruct (Habsurd) | 2: #Heq destruct (Heq) /3 by ex_intro, conj, vint_eq/ | 3: #Heq destruct (Heq) /3 by ex_intro, conj, vfloat_eq/ | 4,5,6: whd in match (try_cast_null ?????); normalize nodelta @eq_bv_elim [ 1,3,5: #Heq destruct (Heq) >eq_intsize_identity normalize nodelta whd in match (m_bind ?????); #Heq destruct (Heq) /3 by ex_intro, conj, vnull_eq/ | 2,4,6: #Hneq >eq_intsize_identity normalize nodelta whd in match (m_bind ?????); #Habsurd destruct (Habsurd) ] ] ] | 4,5,6: whd in match (try_cast_null ?????); normalize nodelta @eq_bv_elim [ 1,3,5: #Heq destruct (Heq) normalize nodelta whd in match (m_bind ?????); #Habsurd destruct (Habsurd) | 2,4,6: #Hneq normalize nodelta whd in match (m_bind ?????); #Habsurd destruct (Habsurd) ] ] | 3: #f whd in match (exec_cast ????); cases ty [ 1: | 2: #sz #sg | 3: #fl | 4: #ptrty | 5: #arrayty #n | 6: #tl #retty | 7: #id #fl | 8: #id #fl | 9: #comptrty ] normalize nodelta [ 1,2,4,5,6,7,8,9: #Habsurd destruct (Habsurd) ] cases cast_ty [ 1: | 2: #csz #csg | 3: #cfl | 4: #cptrty | 5: #carrayty #cn | 6: #ctl #cretty | 7: #cid #cfl | 8: #cid #cfl | 9: #ccomptrty ] normalize nodelta [ 1,4,5,6,7,8,9: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) [ 1: /3 by ex_intro, conj, vint_eq/ | 2: /3 by ex_intro, conj, vfloat_eq/ ] | 4: whd in match (exec_cast ????); cases ty [ 1: | 2: #sz #sg | 3: #fl | 4: #ptrty | 5: #arrayty #n | 6: #tl #retty | 7: #id #fl | 8: #id #fl | 9: #comptrty ] normalize [ 1,2,3,7,8,9: #Habsurd destruct (Habsurd) ] cases cast_ty normalize nodelta [ 1,10,19: #Habsurd destruct (Habsurd) | 2,11,20: #csz #csg #Habsurd destruct (Habsurd) | 3,12,21: #cfl #Habsurd destruct (Habsurd) | 4,13,22: #cptrty #Heq destruct (Heq) /3 by ex_intro, conj, vnull_eq/ | 5,14,23: #carrayty #cn #Heq destruct (Heq) /3 by ex_intro, conj, vnull_eq/ | 6,15,24: #ctl #cretty #Heq destruct (Heq) /3 by ex_intro, conj, vnull_eq/ | 7,16,25: #cid #cfl #Habsurd destruct (Habsurd) | 8,17,26: #cid #cfl #Habsurd destruct (Habsurd) | 9,18,27: #ccomptrty #Habsurd destruct (Habsurd) ] | 5: #p1 #p2 #Hembed whd in match (exec_cast ????); cases ty [ 1: | 2: #sz #sg | 3: #fl | 4: #ptrty | 5: #arrayty #n | 6: #tl #retty | 7: #id #fl | 8: #id #fl | 9: #comptrty ] normalize [ 1,2,3,7,8,9: #Habsurd destruct (Habsurd) ] cases cast_ty normalize nodelta [ 1,10,19: #Habsurd destruct (Habsurd) | 2,11,20: #csz #csg #Habsurd destruct (Habsurd) | 3,12,21: #cfl #Habsurd destruct (Habsurd) | 4,13,22: #cptrty #Heq destruct (Heq) %{(Vptr p2)} @conj try @refl @vptr_eq assumption | 5,14,23: #carrayty #cn #Heq destruct (Heq) %{(Vptr p2)} @conj try @refl @vptr_eq assumption | 6,15,24: #ctl #cretty #Heq destruct (Heq) %{(Vptr p2)} @conj try @refl @vptr_eq assumption | 7,16,25: #cid #cfl #Habsurd destruct (Habsurd) | 8,17,26: #cid #cfl #Habsurd destruct (Habsurd) | 9,18,27: #ccomptrty #Habsurd destruct (Habsurd) ] qed. lemma bool_of_val_value_eq : ∀E,v1,v2. value_eq E v1 v2 → ∀ty,b.exec_bool_of_val v1 ty = OK ? b → exec_bool_of_val v2 ty = OK ? b. #E #v1 #v2 #Hvalue_eq #ty #b @(value_eq_inversion … Hvalue_eq) // [ 1: #v #H normalize in H; destruct (H) | 2: #p1 #p2 #Hembed #H @H ] qed. (* lemma sim_related_globals : ∀ge,ge',en1,m1,en2,m2,ext. ∀E:embedding. ∀Hext:memory_ext E m1 m2. switch_removal_globals E ? fundef_switch_removal ge ge' → disjoint_extension en1 m1 en2 m2 ext E Hext → ext_fresh_for_genv ext ge → (∀e. exec_expr_sim E (exec_expr ge en1 m1 e) (exec_expr ge' en2 m2 e)) ∧ (∀ed, ty. exec_lvalue_sim E (exec_lvalue' ge en1 m1 ed ty) (exec_lvalue' ge' en2 m2 ed ty)). #ge #ge' #en1 #m1 #en2 #m2 #ext #E #Hext #Hrelated #Hdisjoint #Hext_fresh_for_genv @expr_lvalue_ind_combined [ 1: #csz #cty #i #a1 whd in match (exec_expr ????); elim cty [ 1: | 2: #sz #sg | 3: #fl | 4: #ty | 5: #ty #n | 6: #tl #ty | 7: #id #fl | 8: #id #fl | 9: #ty ] normalize nodelta [ 2: cases (eq_intsize csz sz) normalize nodelta [ 1: #H destruct (H) /4 by ex_intro, conj, vint_eq/ | 2: #Habsurd destruct (Habsurd) ] | 4,5,6: #_ #H destruct (H) | *: #H destruct (H) ] | 2: #ty #fl #a1 whd in match (exec_expr ????); #H1 destruct (H1) /4 by ex_intro, conj, vint_eq/ | 3: * [ 1: #sz #i | 2: #fl | 3: #var_id | 4: #e1 | 5: #e1 | 6: #op #e1 | 7: #op #e1 #e2 | 8: #cast_ty #e1 | 9: #cond #iftrue #iffalse | 10: #e1 #e2 | 11: #e1 #e2 | 12: #sizeofty | 13: #e1 #field | 14: #cost #e1 ] #ty whd in ⊢ (% → ?); #Hind try @I whd in match (Plvalue ???); [ 1,2,3: whd in match (exec_expr ????); whd in match (exec_expr ????); #a1 cases (exec_lvalue' ge en1 m1 ? ty) in Hind; [ 2,4,6: #error #_ normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) | 1,3,5: #b1 #H elim (H b1 (refl ??)) #b2 * elim b1 * #bl1 #o1 #tr1 elim b2 * #bl2 #o2 #tr2 #Heq >Heq normalize nodelta * #Hvalue_eq #Htrace_eq whd in match (load_value_of_type' ???); whd in match (load_value_of_type' ???); lapply (load_value_of_type_inj E … (\fst a1) … ty (me_inj … Hext) Hvalue_eq) cases (load_value_of_type ty m1 bl1 o1) [ 1,3,5: #_ #Habsurd normalize in Habsurd; destruct (Habsurd) | 2,4,6: #v #Hyp normalize in ⊢ (% → ?); #Heq destruct (Heq) elim (Hyp (refl ??)) #v2 * #Hload #Hvalue_eq >Hload normalize /4 by ex_intro, conj/ ] ] ] | 4: #v #ty whd * * #b1 #o1 #tr1 whd in match (exec_lvalue' ?????); whd in match (exec_lvalue' ?????); lapply (Hdisjoint v) lapply (Hext_fresh_for_genv v) cases (mem_assoc_env v ext) #Hglobal [ 1: * #vblock * * #Hlookup_en2 #Hwriteable #Hnot_in_en1 >Hnot_in_en1 normalize in Hglobal ⊢ (% → ?); >(Hglobal (refl ??)) normalize #Habsurd destruct (Habsurd) | 2: normalize nodelta cases (lookup ?? en1 v) normalize nodelta [ 1: #Hlookup2 >Hlookup2 normalize nodelta lapply (rg_find_symbol … Hrelated v) cases (find_symbol ???) normalize [ 1: #_ #Habsurd destruct (Habsurd) | 2: #block cases (lookup ?? (symbols clight_fundef ge') v) [ 1: normalize nodelta #Hfalse @(False_ind … Hfalse) | 2: #block' normalize #Hvalue_eq #Heq destruct (Heq) %{〈block',mk_offset (zero offset_size),[]〉} @conj try @refl normalize /2/ ] ] | 2: #block cases (lookup ?? en2 v) normalize nodelta [ 1: #Hfalse @(False_ind … Hfalse) | 2: #b * #Hvalid_block #Hvalue_eq #Heq destruct (Heq) %{〈b, zero_offset, E0〉} @conj try @refl normalize /2/ ] ] ] | 5: #e #ty whd in ⊢ (% → %); whd in match (exec_lvalue' ?????); whd in match (exec_lvalue' ?????); cases (exec_expr ge en1 m1 e) [ 1: * #v1 #tr1 #H elim (H 〈v1,tr1〉 (refl ??)) * #v1' #tr1' * #Heq >Heq normalize nodelta * elim v1 normalize nodelta [ 1: #_ #_ #a1 #Habsurd destruct (Habsurd) | 2: #sz #i #_ #_ #a1 #Habsurd destruct (Habsurd) | 3: #fl #_ #_ #a1 #Habsurd destruct (Habsurd) | 4: #_ #_ #a1 #Habsurd destruct (Habsurd) | 5: #ptr #Hvalue_eq lapply (value_eq_ptr_inversion … Hvalue_eq) * #p2 * #Hp2_eq >Hp2_eq in Hvalue_eq; elim ptr #b1 #o1 elim p2 #b2 #o2 #Hvalue_eq normalize cases (E b1) [ 1: normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) ] * #b2' #o2' normalize #Heq destruct (Heq) #Htrace destruct (Htrace) * * #b1' #o1' #tr1'' #Heq2 destruct (Heq2) normalize %{〈b2,mk_offset (addition_n ? (offv o1') (offv o2')),tr1''〉} @conj try @refl normalize @conj // ] | 2: #error #_ normalize #a1 #Habsurd destruct (Habsurd) ] | 6: #ty #e #ty' #Hsim @(exec_lvalue_expr_elim … Hsim) cases ty [ 1: | 2: #sz #sg | 3: #fl | 4: #ty | 5: #ty #n | 6: #tl #ty | 7: #id #fl | 8: #id #fl | 9: #ty ] * #b1 #o1 * #b2 #o2 normalize nodelta try /2 by I/ #tr #H @conj try @refl try assumption | 7: #ty #op #e #Hsim @(exec_expr_expr_elim … Hsim) #v1 #v2 #trace #Hvalue_eq lapply (unary_operation_value_eq E op v1 v2 (typeof e) Hvalue_eq) cases (sem_unary_operation op v1 (typeof e)) normalize nodelta [ 1: #_ @I | 2: #r1 #H elim (H r1 (refl ??)) #r1' * #Heq >Heq normalize /2/ ] | 8: #ty #op #e1 #e2 #Hsim1 #Hsim2 @(exec_expr_expr_elim … Hsim1) #v1 #v2 #trace #Hvalue_eq cases (exec_expr ge en1 m1 e2) in Hsim2; [ 2: #error // ] * #val #trace normalize in ⊢ (% → ?); #Hsim2 elim (Hsim2 ? (refl ??)) * #val2 #trace2 * #Hexec2 * #Hvalue_eq2 #Htrace >Hexec2 whd in match (m_bind ?????); whd in match (m_bind ?????); lapply (binary_operation_value_eq E op … (typeof e1) (typeof e2) ?? Hvalue_eq Hvalue_eq2 (me_inj … Hext)) cases (sem_binary_operation op v1 (typeof e1) val (typeof e2) m1) [ 1: #_ // ] #opval #Hop elim (Hop ? (refl ??)) #opval' * #Hopval_eq #Hvalue_eq_opval >Hopval_eq normalize destruct /2 by conj/ | 9: #ty #cast_ty #e #Hsim @(exec_expr_expr_elim … Hsim) #v1 #v2 #trace #Hvalue_eq lapply (cast_value_eq E m1 m2 … Hvalue_eq (typeof e) cast_ty) cases (exec_cast m1 v1 (typeof e) cast_ty) [ 2: #error #_ normalize @I | 1: #res #H lapply (H res (refl ??)) whd in match (m_bind ?????); * #res' * #Hexec_cast >Hexec_cast #Hvalue_eq normalize nodelta @conj // ] | 10: #ty #e1 #e2 #e3 #Hsim1 #Hsim2 #Hsim3 @(exec_expr_expr_elim … Hsim1) #v1 #v2 #trace #Hvalue_eq lapply (bool_of_val_value_eq E v1 v2 Hvalue_eq (typeof e1)) cases (exec_bool_of_val ? (typeof e1)) #b [ 2: #_ normalize @I ] #H lapply (H ? (refl ??)) #Hexec >Hexec normalize cases b normalize nodelta [ 1: (* true branch *) cases (exec_expr ge en1 m1 e2) in Hsim2; [ 2: #error normalize #_ @I | 1: * #e2v #e2tr normalize #H elim (H ? (refl ??)) * #e2v' #e2tr' * #Hexec2 >Hexec2 * #Hvalue_eq2 #Htrace_eq2 normalize destruct @conj try // ] | 2: (* false branch *) cases (exec_expr ge en1 m1 e3) in Hsim3; [ 2: #error normalize #_ @I | 1: * #e3v #e3tr normalize #H elim (H ? (refl ??)) * #e3v' #e3tr' * #Hexec3 >Hexec3 * #Hvalue_eq3 #Htrace_eq3 normalize destruct @conj // ] ] | 11,12: #ty #e1 #e2 #Hsim1 #Hsim2 @(exec_expr_expr_elim … Hsim1) #v1 #v1' #trace #Hvalue_eq lapply (bool_of_val_value_eq E v1 v1' Hvalue_eq (typeof e1)) cases (exec_bool_of_val v1 (typeof e1)) [ 2,4: #error #_ normalize @I ] #b cases b #H lapply (H ? (refl ??)) #Heq >Heq whd in match (m_bind ?????); whd in match (m_bind ?????); [ 2,3: normalize @conj try @refl try @vint_eq ] cases (exec_expr ge en1 m1 e2) in Hsim2; [ 2,4: #error #_ normalize @I ] * #v2 #tr2 whd in ⊢ (% → %); #H2 normalize nodelta elim (H2 ? (refl ??)) * #v2' #tr2' * #Heq2 * #Hvalue_eq2 #Htrace2 >Heq2 normalize nodelta lapply (bool_of_val_value_eq E v2 v2' Hvalue_eq2 (typeof e2)) cases (exec_bool_of_val v2 (typeof e2)) [ 2,4: #error #_ normalize @I ] #b2 #H3 lapply (H3 ? (refl ??)) #Heq3 >Heq3 normalize nodelta destruct @conj try @conj // cases b2 whd in match (of_bool ?); @vint_eq | 13: #ty #ty' cases ty [ 1: | 2: #sz #sg | 3: #fl | 4: #ty | 5: #ty #n | 6: #tl #ty | 7: #id #fl | 8: #id #fl | 9: #ty ] whd in match (exec_expr ????); whd * #v #trace #Heq destruct %{〈Vint sz (repr sz (sizeof ty')), E0〉} @conj try @refl @conj // | 14: #ty #ed #aggregty #i #Hsim whd * * #b #o #tr normalize nodelta whd in match (exec_lvalue' ?????); whd in match (exec_lvalue' ge' en2 m2 (Efield (Expr ed aggregty) i) ty); whd in match (typeof ?); cases aggregty in Hsim; [ 1: | 2: #sz' #sg' | 3: #fl' | 4: #ty' | 5: #ty' #n' | 6: #tl' #ty' | 7: #id' #fl' | 8: #id' #fl' | 9: #ty' ] normalize nodelta #Hsim [ 1,2,3,4,5,6,9: #Habsurd destruct (Habsurd) ] whd in match (m_bind ?????); whd in match (m_bind ?????); whd in match (exec_lvalue ge en1 m1 (Expr ed ?)); cases (exec_lvalue' ge en1 m1 ed ?) in Hsim; [ 2,4: #error #_ normalize in ⊢ (% → ?); #Habsurd destruct (Habsurd) ] * * #b1 #o1 #tr1 whd in ⊢ (% → ?); #H elim (H ? (refl ??)) * * #b1' #o1' #tr1' * #Hexec normalize nodelta * #Hvalue_eq #Htrace_eq whd in match (exec_lvalue ????); >Hexec normalize nodelta [ 2: #Heq destruct (Heq) %{〈 b1',o1',tr1'〉} @conj // normalize @conj // ] cases (field_offset i fl') [ 2: #error normalize #Habsurd destruct (Habsurd) ] #offset whd in match (m_bind ?????); #Heq destruct (Heq) whd in match (m_bind ?????); %{〈b1',shift_offset (bitsize_of_intsize I32) o1' (repr I32 offset),tr1'〉} @conj destruct // normalize nodelta @conj try @refl @vptr_eq -H lapply (value_eq_ptr_inversion … Hvalue_eq) * #p2 * #Hptr_eq whd in match (pointer_translation ??); whd in match (pointer_translation ??); cases (E b) [ 1: normalize nodelta #Habsurd destruct (Habsurd) ] * #b' #o' normalize nodelta #Heq destruct (Heq) destruct (Hptr_eq) cut (offset_plus (mk_offset (addition_n offset_size (offv o1) (sign_ext (bitsize_of_intsize I32) offset_size (repr I32 offset)))) o' = (shift_offset (bitsize_of_intsize I32) (offset_plus o1 o') (repr I32 offset))) [ whd in match (shift_offset ???) in ⊢ (???%); whd in match (offset_plus ??) in ⊢ (??%%); /3 by associative_addition_n, commutative_addition_n, refl/ ] #Heq >Heq @refl | 15: #ty #l #e #Hsim @(exec_expr_expr_elim … Hsim) #v1 #v2 #trace #Hvalue_eq normalize nodelta @conj // | 16: * [ 1: #sz #i | 2: #fl | 3: #var_id | 4: #e1 | 5: #e1 | 6: #op #e1 | 7: #op #e1 #e2 | 8: #cast_ty #e1 | 9: #cond #iftrue #iffalse | 10: #e1 #e2 | 11: #e1 #e2 | 12: #sizeofty | 13: #e1 #field | 14: #cost #e1 ] #ty normalize in ⊢ (% → ?); [ 3,4,13: @False_ind | *: #_ normalize #a1 #Habsurd destruct (Habsurd) ] ] qed. *)