include "Clight/Cexec.ma". include "Clight/MemProperties.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). *) (* --------------------------------------------------------------------------- *) (* Aux lemmas that are likely to be found elsewhere. *) (* --------------------------------------------------------------------------- *) lemma zlt_succ : ∀a,b. Zltb a b = true → Zltb a (Zsucc b) = true. #a #b #HA lapply (Zltb_true_to_Zlt … HA) #HA_prop @Zlt_to_Zltb_true /2/ qed. lemma zlt_succ_refl : ∀a. Zltb a (Zsucc a) = true. #a @Zlt_to_Zltb_true /2/ qed. (* definition le_offset : offset → offset → bool. λx,y. Zleb (offv x) (offv y). *) lemma not_refl_absurd : ∀A:Type[0].∀x:A. x ≠ x → False. /2/. qed. lemma eqZb_reflexive : ∀x:Z. eqZb x x = true. #x /2/. qed. (* --------------------------------------------------------------------------- *) (* Some shorthands for conversion functions from BitVectorZ. *) (* --------------------------------------------------------------------------- *) definition Zoub ≝ Z_of_unsigned_bitvector. definition boZ ≝ bitvector_of_Z. (* Offsets are just bitvectors packed inside some useless and annoying constructor. *) definition Zoo ≝ λx. Zoub ? (offv x). definition boo ≝ λx. mk_offset (boZ ? x). (* --------------------------------------------------------------------------- *) (* In the definition of injections below, we maintain a list of blocks that are writeable. We need some lemmas to reason on the fact that a block cannot be both writeable and not writeable, etc. *) (* --------------------------------------------------------------------------- *) (* When equality on A is decidable, [mem A elt l] is too. *) lemma mem_dec : ∀A:Type[0]. ∀eq:(∀a,b:A. a = b ∨ a ≠ b). ∀elt,l. mem A elt l ∨ ¬ (mem A elt l). #A #dec #elt #l elim l [ 1: normalize %2 /2/ | 2: #hd #tl #Hind elim (dec elt hd) [ 1: #Heq >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 : value_eq E Vundef Vundef | vint_eq : ∀sz,i. value_eq E (Vint sz i) (Vint sz i) | 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 block_is_empty : mem → block → Prop ≝ λm,b. high_bound m b ≤ low_bound m b. (* Adapted from Compcert's Memory *) definition non_aliasing : embedding → mem → Prop ≝ λE,m. ∀b1,b2,b1',b2',delta1,delta2. b1 ≠ b2 → block_region b1 = block_region b2 → E b1 = Some ? 〈b1',delta1〉 → E b2 = Some ? 〈b2',delta2〉 → (b1' ≠ b2') ∨ block_is_empty m b1' ∨ block_is_empty m b2' ∨ high_bound m b1 + (Zoo delta1) ≤ low_bound m b2 + (Zoo delta2) ∨ high_bound m b2 + (Zoo delta2) ≤ low_bound m b1 + (Zoo delta1). (* Definition of a memory injection *) record memory_inj (E : embedding) (m1 : mem) (m2 : mem) : Prop ≝ { (* 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'; (* sub-blocks do not overlap (non-aliasing property) *) mi_nonalias : non_aliasing E m1 }. (* ---------------------------------------------------------------------------- *) (* Setting up an empty injection *) (* ---------------------------------------------------------------------------- *) definition empty_injection : memory_inj (λx. None ?) empty empty. % [ 1: whd #b1 #off1 #b2 #off2 #ty #v1 #Hvalid #Hptr_trans #Hload normalize in Hptr_trans; destruct | 2: * #b #o #p' #_ normalize #Habsurd destruct | 3: #b #b' #o' #Habsurd destruct | 4: normalize #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9 destruct ] qed. (* ---------------------------------------------------------------------------- *) (* 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: #Habsurd destruct (Habsurd) | 2: #sz #i #Habsurd destruct (Habsurd) | 3: #Habsurd destruct (Habsurd) | 4: #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 → (P Vundef Vundef) → (∀sz,i. P (Vint sz i) (Vint sz i)) → (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 inversion Heq [ 1: #Hv1 #Hv2 #_ destruct @H1 | 2: #sz #i #Hv1 #Hv2 #_ destruct @H2 | 3: #Hv1 #Hv2 #_ destruct @H3 | 4: #p1 #p2 #Hembed #Hv1 #Hv2 #_ destruct @H4 // ] 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 [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] whd in match (load_value_of_type ????); [ 1,6,7: #_ #Habsurd destruct (Habsurd) ] #Hmode [ 1,2,5: [ 1: elim sz ] 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 [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] whd in match (load_value_of_type ????); [ 1,6,7: #_ #Habsurd destruct (Habsurd) ] #Hmode [ 1,2,5: [ 1: elim sz ] 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 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 @Zltb_true_to_Zlt 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 >H3 @I qed. (* Making explicit the contents of memory_inj for load_value_of_type. * Equivalent to Lemma 59 of Leroy & Blazy *) 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 [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] normalize in ⊢ ((???%) → ?); #Hmode #Hyp [ 1,6,7: normalize in Hyp; destruct (Hyp) | 4,5: 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 (Zlt_to_Zltb_true … 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 by not_to_not/ ] #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 in m2 preserves [memory_inj]. * This is lemma 43 from Leroy&Blazy. *) lemma alloc_memory_inj_m2 : ∀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 % [ 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; [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] #Hload normalize in ⊢ ((???%) → ?); #Haccess [ 1,6,7: normalize in Hload; destruct (Hload) | 2,3,8: 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 | 4,5: whd in match (load_value_of_type ????) in Hload ⊢ %; @Hload ] | #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 | @(mi_region … Hmemory_inj) | @(mi_nonalias … Hmemory_inj) qed. (* Allocating in m1 such that the resulting block has no image in m2 preserves the injection. This is lemma 44 for Leroy&Blazy. The proof should proceed as in Blazy & Leroy. Careful of using the mi_embedding hypothesis to rule out absurd cases. *) axiom alloc_memory_inj_m1 : ∀E:embedding. ∀m1,m2,m1',z1,z2,r,new_block. ∀H:memory_inj E m1 m2. alloc m1 z1 z2 r = 〈m1', new_block〉 → E (pi1 … new_block) = None ? → memory_inj E m1' m2. (* Embed a fresh block inside an unmapped portion of the target block. *) axiom alloc_memory_inj_m1_to_m2 : ∀E:embedding. ∀m1,m2,m1':mem. ∀z1,z2:Z. ∀b2,new_block. ∀delta:offset. ∀H:memory_inj E m1 m2. alloc m1 z1 z2 (block_region b2) = 〈m1', new_block〉 → E (pi1 … new_block) = Some ? 〈b2, delta〉 → low_bound m2 b2 ≤ z1 + Zoo delta → z2 + Zoo delta ≤ high_bound m2 b2 → (∀b,b',delta'. b ≠ (pi1 … new_block) → E b = Some ? 〈b', delta'〉 → high_bound m1 b + (Zoo delta') ≤ z1 + (Zoo delta) ∨ z2 + (Zoo delta) ≤ low_bound m1 b + (Zoo delta')) → memory_inj E m1' m2. (* -------------------------------------- *) (* Lemmas pertaining to [free] *) (* Lemma 46 by Leroy&Blazy *) lemma free_load_sim_ptr_m1 : ∀E:embedding. ∀m1,m2,m1'. load_sim_ptr E m1 m2 → ∀b. free m1 b = m1' → load_sim_ptr E m1' m2. #E #m1 #m2 #m1' #Hinj #b #Hfree whd #b1 #off1 #b2 #off2 #ty #v1 #Hvalid #Htrans cases ty [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] [ 1,4,5,6,7: whd in ⊢ ((??%?) → ?); #Heq destruct (Heq) normalize %{(Vptr (mk_pointer b2 off2))} @conj try @refl %4 @Htrans ] #Hload Hregions_eq cases (block_region b') normalize nodelta @(eqZb_elim … (block_id b) (block_id b')) [ 1,3: #Heq normalize nodelta // | 2,4: #Hneq normalize nodelta @HA ] qed. (* Lemma 49 in Leroy&Blazy *) axiom free_non_aliasing_m1 : ∀E:embedding. ∀m1,m2,m1'. memory_inj E m1 m2 → ∀b. free m1 b = m1' → non_aliasing E m1'. (* The proof proceeds by a first case analysis to see wether b lives in the same world as the non-aliasing blocks (if not : trivial), in this case we proceed by a case analysis on the non-aliasing hypothesis in memory_inj. *) (* Extending lemma 46 and 49 to memory injections *) lemma free_memory_inj_m1 : ∀E:embedding. ∀m1,m2,m1'. memory_inj E m1 m2 → ∀b. free m1 b = m1' → memory_inj E m1' m2. #E #m1 #m2 #m1' #Hinj #b #Hfree cases Hinj #Hload_sim #Hvalid #Hregion #Hnonalias % try assumption [ @(free_load_sim_ptr_m1 … Hload_sim … Hfree) | #p #p' #Hvalid_m1' #Hptr_translation @(Hvalid p p' ? Hptr_translation) (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 [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] // [ 1: elim sz ] whd in ⊢ (??%%); >(bestorev_loadn_conservation … Hstore … Hneq) @refl qed. (* lift [bestorev_load_value_of_type_conservation] to storen *) lemma storen_load_value_of_type_conservation : ∀l,mA,mB,wb,wo. storen mA (mk_pointer wb wo) l = 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. #l elim l [ 1: #mA #mB #wb #wo normalize #Heq destruct // | 2: #hd #tl #Hind #mA #mB #wb #wo #Hstoren cases (some_inversion ????? Hstoren) #mA' * #Hbestorev #Hstoren' whd in match (shift_pointer ???) in Hstoren':(??%?); #rb #ro #Hneq_block #ty lapply (Hind ???? Hstoren' … ro … Hneq_block ty) #Heq1 lapply (bestorev_load_value_of_type_conservation … Hbestorev … ro … Hneq_block ty) #Heq2 // ] qed. lemma store_value_of_type_load_value_of_type_conservation : ∀ty,mA,mB,wb,wo,v. store_value_of_type ty mA 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. #ty #mA #mB #wb #wo #v #Hstore #rb #ro #Heq_block cases ty in Hstore; [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] [ 1,4,5,6,7: whd in ⊢ ((??%?) → ?); #Heq destruct ] whd in ⊢ ((??%?) → ?); #Hstoren @(storen_load_value_of_type_conservation … Hstoren … Heq_block) qed. definition typesize' ≝ λty. typesize (typ_of_type ty). lemma storen_load_value_of_type_conservation_in_block_high : ∀E,mA,mB,mC,wo,l. memory_inj E mA mB → ∀blo. storen mB (mk_pointer blo wo) l = Some ? mC → ∀b1,delta. E b1 = Some ? 〈blo,delta〉 → high_bound mA b1 + Zoo delta < Zoo wo → ∀ty,off. Zoo off + Z_of_nat (typesize' ty) < high_bound mA b1 → low_bound … mA b1 ≤ Zoo off → Zoo off < high_bound … mA b1 → load_value_of_type ty mB blo (mk_offset (addition_n ? (offv off) (offv delta))) = load_value_of_type ty mC blo (mk_offset (addition_n ? (offv off) (offv delta))). #E #mA #mB #mC #wo #data #Hinj #blo #Hstoren #b1 #delta #Hembed #Hhigh #ty (* need some stuff asserting that if a ptr is valid at the start of a write it's valid at the end. *) cases data in Hstoren; [ 1: normalize in ⊢ (% → ?); #Heq destruct // | 2: #xd #data ] #Hstoren cases ty [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ]#off #Hofflt #Hlow_load #Hhigh_load try @refl whd in match (load_value_of_type ????) in ⊢ (??%%); [ 1: lapply (storen_to_valid_pointer … Hstoren) * * #Hbounds #Hbefore #Hafter lapply Hofflt -Hofflt lapply Hlow_load -Hlow_load lapply Hhigh_load -Hhigh_load lapply off -off whd in match typesize'; normalize nodelta generalize in match (typesize ?); #n elim n try // #n' #Hind #o #Hhigh_load #Hlow_load #Hlt whd in match (loadn ???) in ⊢ (??%%); whd in match (beloadv ??) in ⊢ (??%%); cases (valid_pointer_to_Prop … Hbefore) * #HltB_store #HlowB_store #HhighB_store cases (valid_pointer_to_Prop … Hafter) * #HltC_store #HlowC_store #HhighC_store >(Zlt_to_Zltb_true … HltB_store) >(Zlt_to_Zltb_true … HltC_store) normalize nodelta cut (Zleb (low (blocks mB blo)) (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta))) = true) [ 1: (* Notice that: low mA b1 ≤ o < high mA b1 hence, since E b1 = 〈blo, delta〉 with delta >= 0 low mB blo ≤ (low mA b1 + delta) ≤ o+delta < (high mA b1 + delta) ≤ high mB blo *) @cthulhu ] #HA >HA >andb_lsimpl_true -HA cut (Zltb (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta))) (high (blocks mB blo)) = true) [ 1: (* Same argument as above *) @cthulhu ] #HA >HA normalize nodelta -HA cut (Zleb (low (blocks mC blo)) (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta))) = true) [ 1: (* Notice that storen does not modify the bounds of a block. Related lemma present in [MemProperties]. This cut follows from this lemma (which needs some info on the size of the data written, which we have but must make explicit) and from the first cut. *) @cthulhu ] #HA >HA >andb_lsimpl_true -HA cut (Zltb (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta))) (high (blocks mC blo)) = true) [ 1: (* Same argument as above *) @cthulhu ] #HA >HA normalize nodelta -HA normalize in match (bitvector_of_nat ??); whd in match (shift_pointer ???); whd in match (shift_offset ???); >commutative_addition_n >associative_addition_n lapply (Hind (mk_offset (addition_n offset_size (sign_ext 2 ? [[false; true]]) (offv o))) ???) [ 1: (* Provable from Hlt *) @cthulhu | 2: (* Provable from Hlow_load, need to make a "succ" commute from bitvector to Z *) @cthulhu | 3: (* Provable from Hlt, again *) @cthulhu ] cases (loadn mB (mk_pointer blo (mk_offset (addition_n ? (addition_n ? (sign_ext 2 offset_size [[false; true]]) (offv o)) (offv delta)))) n') normalize nodelta cases (loadn mC (mk_pointer blo (mk_offset (addition_n ? (addition_n ? (sign_ext 2 offset_size [[false; true]]) (offv o)) (offv delta)))) n') normalize nodelta try // [ 1,2: #l #Habsurd destruct ] #l1 #l2 #Heq cut (contents (blocks mB blo) (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta))) = contents (blocks mC blo) (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta)))) [ 1: (* Follows from Hhigh, indirectly *) @cthulhu ] #Hcontents_eq >Hcontents_eq whd in match (be_to_fe_value ??) in ⊢ (??%%); cases (contents (blocks mC blo) (Z_of_unsigned_bitvector ? (addition_n ? (offv o) (offv delta)))) normalize nodelta try // (* Ok this is going to be more painful than what I thought. *) @cthulhu | *: @cthulhu ] qed. lemma storen_load_value_of_type_conservation_in_block_low : ∀E,mA,mB,mC,wo,l. memory_inj E mA mB → ∀blo. storen mB (mk_pointer blo wo) l = Some ? mC → ∀b1,delta. E b1 = Some ? 〈blo,delta〉 → Zoo wo < low_bound mA b1 + Zoo delta → ∀ty,off. Zoo off + Z_of_nat (typesize' ty) < high_bound mA b1 → low_bound … mA b1 ≤ Zoo off → Zoo off < high_bound … mA b1 → load_value_of_type ty mB blo (mk_offset (addition_n ? (offv off) (offv delta))) = load_value_of_type ty mC blo (mk_offset (addition_n ? (offv off) (offv delta))). @cthulhu qed. (* Writing in the "extended" part of a memory preserves the underlying injection. *) lemma bestorev_memory_inj_to_load_sim : ∀E,mA,mB,mC. ∀Hext:memory_inj E mA mB. ∀block2. ∀i : offset. ∀ty : type. (∀block1,delta. E block1 = Some ? 〈block2, delta〉 → (high_bound mA block1 + (Zoo delta) < (Zoo i)) ∨ (Zoo i + (sizeof ty) ≤ (low_bound mA block1 + (Zoo delta)))) → ∀v.store_value_of_type ty mB block2 i v = Some ? mC → load_sim_ptr E mA mC. #E #mA #mB #mC #Hinj #block2 #i #storety cases storety [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ]#Hout #storeval #Hstore whd #b1 #off1 #b2 #off2 #ty #readval #Hvalid #Hptr_trans #Hload_value whd in Hstore:(??%?); [ 1,5,6,7,8: destruct ] [ 1: lapply (mi_inj … Hinj … Hvalid … Hptr_trans … Hload_value) lapply Hload_value -Hload_value cases ty [ | #sz' #sg' | #ptr_ty' | #array_ty' #array_sz' | #domain' #codomain' | #structname' #fieldspec' | #unionname' #fieldspec' | #id' ] #Hload_value (* Prove that the contents of mB where v1 was were untouched. *) * #readval' * #Hload_value2 #Hvalue_eq %{readval'} @conj try assumption cases (some_inversion ????? Hptr_trans) * #b2' #delta' * #Hembed -Hptr_trans normalize nodelta #Heq destruct (Heq) @(eq_block_elim … b2 block2) [ 2,4,6,8: #Hblocks_neq 0 *) @cthulhu | 2,4,6,8: (* deductible from Hload_value *) @cthulhu ] ] ] | *: (* exactly the same proof as above *) @cthulhu ] qed. (* Lift the above result to memory_inj * This is Lemma 40 of Leroy & Blazy *) lemma bestorev_memory_inj_to_memory_inj : ∀E,mA,mB,mC. ∀Hext:memory_inj E mA mB. ∀block2. ∀i : offset. ∀ty : type. (∀block1,delta. E block1 = Some ? 〈block2, delta〉 → (high_bound mA block1 + (Zoo delta) < (Zoo i)) ∨ (Zoo i + (sizeof ty) ≤ (low_bound mA block1 + (Zoo delta)))) → ∀v.store_value_of_type ty mB block2 i v = Some ? mC → memory_inj E mA mC. #E #mA #mB #mC #Hinj #block2 #i #ty #Hout #v #Hstore % lapply (bestorev_memory_inj_to_load_sim … Hinj … Hout … Hstore) #Hsim try // cases Hinj #Hsim' #Hvalid #Hregion #Hnonalias try assumption #p #p' #Hptr #Hptr_trans lapply (Hvalid p p' Hptr Hptr_trans) #Hvalid cases ty in Hstore; [ | #sz' #sg' | #ptr_ty' | #array_ty' #array_sz' | #domain' #codomain' | #structname' #fieldspec' | #unionname' #fieldspec' | #id' ] whd in ⊢ ((??%?) → ?); [ 1,4,5,6,7: #Habsurd destruct ] cases (fe_to_be_values ??) [ 1,3,5,7: whd in ⊢ ((??%?) → ?); #Heq Hnext try // cases (Hbounds (pblock p')) #HA #HB whd in match (low_bound ??); whd in match (high_bound ??); >HA >HB try assumption ] qed. (* ---------------------------------------------------------- *) (* Lemma 41 of the paper of Leroy & Blazy on the memory model * and related lemmas *) (* The back-end might contain something similar to this lemma. *) lemma be_to_fe_value_ptr : ∀b,o. (be_to_fe_value ASTptr (fe_to_be_values ASTptr (Vptr (mk_pointer b o))) = Vptr (mk_pointer b o)). #b * #o whd in ⊢ (??%%); normalize cases b #br #bid normalize nodelta cases br normalize nodelta >eqZb_z_z normalize nodelta cases (vsplit_eq bool 7 8 … o) #lhs * #rhs #Heq <(vsplit_prod … Heq) >eq_v_true normalize nodelta try @refl * // qed. lemma fe_to_be_values_length : ∀E,v1,v2,ty. value_eq E v1 v2 → |fe_to_be_values ty v1| = |fe_to_be_values ty v2|. #E #v1 #v2 #ty #Hvalue_eq @(value_eq_inversion … Hvalue_eq) // qed. lemma value_eq_to_be_and_back_again : ∀E,ty,v1,v2. value_eq E v1 v2 → ast_typ_consistent_with_value ty v1 → value_eq E (be_to_fe_value ty (fe_to_be_values ty v1 )) (be_to_fe_value ty (fe_to_be_values ty v2)). #E #ty #v1 #v2 #Hvalue_eq @(value_eq_inversion … Hvalue_eq) try // [ 1: cases ty // | 2: #sz #i cases ty [ 2: @False_ind | 1: #sz' #sg #H lapply H whd in ⊢ (% → ?); #Heq lapply (fe_to_be_to_fe_value_int … H) #H >H // ] | 3: #p1 #p2 #Hembed cases ty [ 1: #sz #sg @False_ind | 2: #_ cases p1 in Hembed; #b1 #o1 cases p2 #b2 #o2 whd in ⊢ ((??%%) → ?); #H cases (some_inversion ????? H) * #b3 #o3 * #Hembed normalize nodelta #Heq >be_to_fe_value_ptr >be_to_fe_value_ptr destruct %4 whd in match (pointer_translation ??); >Hembed normalize nodelta @refl ] ] qed. (* ------------------------------------------------------------ *) (* Lemma ?? of the paper of Leroy & Blazy on the memory model ? *) lemma storen_parallel_memory_exists : ∀E,m1,m2. memory_inj E m1 m2 → ∀b1,b2,delta. E b1 = Some ? 〈b2,delta〉 → ∀v1,v2. value_eq E v1 v2 → ∀i,m1',ty. storen m1 (mk_pointer b1 i) (fe_to_be_values ty v1) = Some ? m1'→ ∃m2'. storen m2 (mk_pointer b2 (offset_plus i delta)) (fe_to_be_values ty v2) = Some ? m2'. (* ∧ loadn m2' (mk_pointer b2 (offset_plus i delta)) (|fe_to_be_values ty v2|) = Some ? (fe_to_be_values ty v2).*) #E #m1 #m2 #Hinj #b1 #b2 #delta #Hembed #v1 #v2 #Hvalue_eq #i #m1' #ty lapply (mi_valid_pointers … Hinj) generalize in match m2; generalize in match m1; generalize in match i; lapply (fe_to_be_values_length … ty Hvalue_eq) generalize in match (fe_to_be_values ty v2); generalize in match (fe_to_be_values ty v1); #l1 elim l1 [ 1: #l2 #Hl2 cut (l2 = []) [ cases l2 in Hl2; // #hd' #tl' normalize #Habsurd destruct ] #Hl2_empty >Hl2_empty #o #m1 #m2 #_ normalize /2 by ex_intro/ | 2: #hd1 #tl1 #Hind #l2 #Hlength cut (∃hd2,tl2.l2 = hd2::tl2 ∧ |tl1| = |tl2|) [ cases l2 in Hlength; [ normalize #Habsurd destruct | #hd2 #tl2 normalize #H %{hd2} %{tl2} @conj try @refl destruct (H) assumption ] ] * #hd2 * #tl2 * #Heq2 destruct (Heq2) #Hlen_tl #o #m1 #m2 #Hvalid_embed #Hstoren cases (some_inversion ????? Hstoren) #m1_tmp * #Hbestorev #Hstoren_tl lapply (Hvalid_embed (mk_pointer b1 o) (mk_pointer b2 (offset_plus o delta)) ??) [ 1: whd in match (pointer_translation ??); >Hembed normalize nodelta @refl | 2: @(bestorev_to_valid_pointer … Hbestorev) ] #Hvalid_ptr_m2 whd in match (storen ???); lapply (valid_pointer_to_bestorev_ok m2 (mk_pointer b2 (offset_plus o delta)) hd2 Hvalid_ptr_m2) * #m2_tmp #Hbestorev2 >Hbestorev2 normalize nodelta whd in match (shift_pointer ???) in Hstoren_tl ⊢ %; whd in match (shift_offset ???) in Hstoren_tl ⊢ %; (* normalize in match (sign_ext ???) in Hstoren_tl ⊢ %;*) cut (mk_offset (addition_n ? (offv (offset_plus o delta)) (sign_ext 2 offset_size (bitvector_of_nat 2 1))) = offset_plus (offset_plus o (mk_offset (sign_ext 2 offset_size (bitvector_of_nat 2 1)))) delta) [ cases o #o' normalize nodelta whd in match (offset_plus ??) in ⊢ (??%%); whd in match (offset_plus ??) in ⊢ (???(?%)); /2 by associative_addition_n, commutative_addition_n, refl/ ] #Heq_cleanup >Heq_cleanup >Heq_cleanup in Hind; #Hind @(Hind … Hstoren_tl) try assumption #p #p' #Hvalid_in_m1_tmp #Hp_embed @valid_pointer_of_Prop lapply (mem_bounds_invariant_after_bestorev … Hbestorev) * * * * #Hnextblock_eq #Hbounds_eq #Hoffset_in_bounds #Hcontents_at_eq #Hcontents_else_eq lapply (mem_bounds_invariant_after_bestorev … Hbestorev2) * * * * #Hnextblock_eq2 #Hbounds_eq2 #Hoffset_in_bounds2 #Hcontents_at_eq2 #Hcontents_else_eq2 cut (valid_pointer m1 p = true) [ @valid_pointer_of_Prop cases (valid_pointer_to_Prop … Hvalid_in_m1_tmp) >Hnextblock_eq cases (Hbounds_eq (pblock p)) #Hlow' #Hhigh' whd in match (low_bound ??); whd in match (high_bound ??); >Hlow' >Hhigh' * /3 by conj/ ] #Hvalid_in_m1 lapply (Hvalid_embed p p' Hvalid_in_m1 Hp_embed) #Hvalid_in_m2 cases (valid_pointer_to_Prop … Hvalid_in_m2) * #Hnextblock2' #Hlow2' #Hhigh2' @conj try @conj >Hnextblock_eq2 try assumption cases (Hbounds_eq2 … (pblock p')) #Hlow2'' #Hhigh2'' whd in match (low_bound ??); whd in match (high_bound ??); >Hlow2'' >Hhigh2'' assumption ] qed. lemma storen_parallel_memory_exists_and_preserves_loads : ∀E,m1,m2. memory_inj E m1 m2 → ∀b1,b2,delta. E b1 = Some ? 〈b2,delta〉 → ∀v1,v2. value_eq E v1 v2 → ∀i,m1',ty. storen m1 (mk_pointer b1 i) (fe_to_be_values ty v1) = Some ? m1'→ ∃m2'. storen m2 (mk_pointer b2 (offset_plus i delta)) (fe_to_be_values ty v2) = Some ? m2' ∧ loadn m2' (mk_pointer b2 (offset_plus i delta)) (|fe_to_be_values ty v2|) = Some ? (fe_to_be_values ty v2). #E #m1 #m2 #Hinj #b1 #b2 #delta #Hembed #v1 #v2 #Hvalue_eq #i #m1' #ty #Hstoren cases (storen_parallel_memory_exists … Hinj … Hembed … Hvalue_eq i m1' ty Hstoren) #m2' #Hstoren' %{m2'} @conj try assumption @(storen_loadn_ok … Hstoren') // qed. (* If E ⊢ m1 ↦ m2 *and* if E b1 = 〈b2,delta〉, *and* if we write /A PROPER VALUE/ at 〈b1,i〉 successfuly, *then* we can write /something value_eq-compatible/ in m2 at 〈b2,i+delta〉 successfuly yielding m2' *and* preserve the injection : E ⊢ m1' ↦ m2' N.B.: this result cannot be given at the back-end level : we could overwrite a pointer and break the value_eq correspondence between the memories. *) axiom storen_parallel_aux : ∀E,m1,m2. ∀Hinj:memory_inj E m1 m2. ∀v1,v2. value_eq E v1 v2 → ∀b1,b2,delta. E b1 = Some ? 〈b2, delta〉 → ∀ty,i,m1'. ast_typ_consistent_with_value ty v1 → storen m1 (mk_pointer b1 i) (fe_to_be_values ty v1) = Some ? m1' → ∃m2'. storen m2 (mk_pointer b2 (offset_plus i delta)) (fe_to_be_values ty v2) = Some ? m2' ∧ memory_inj E m1' m2'. (* This lemma is not provable as in CompCert. In the following chunk of text, I'll try to * explain why, and how we might still be able to prove it given enough time. I'll be refering to paper by Leroy & Blazy in the J.A.R. In CompCert, when storing some data of size S in some location 〈block, offset〉, what happens is that 1) the memory is extended with a map from 〈block,offset〉 to the actual data 2) the memory inteval from 〈block,offset+1〉 to 〈block,offset+S-1〉 is invalidated, meaning that every load in this interval will fail. This behaviour is documented at page 9 in the aforementioned paper. The memory model of Cerco is in a sense much more realistic. When storing a front-end value fv, the story is the following : 1) the value fv is marshelled into a list of back-end values (byte-sized) which correspond to the actual size of the data in-memory. This list is then stored as-is at the location 〈block, offset〉. Now on to the lemma we want to prove. The difficult part is to prove [load_sim E m1' m2'], i.e. we want to prove that ∀c1,off1,c2,off2,delta. s.t. E c1 = 〈c2, off2〉 load_value_of_type m1' c1 off1 = ⌊vA⌋ → load_value_of_type m2' c2 (off1+off2) = ⌊vB⌋ ∧ value_eq E vA vB, where m1' and m2' are the result of storing some values v1 and v2 in resp. m1 and m2 at some resp. locations 〈b1,i〉 and (pointer_translation E b1 i) (cf statement). In CompCert, the proof of this statement goes by case analysis on 〈c1,off1〉. Case 1: 〈b1,i〉 = 〈c1,off1〉 → we read the freshly stored data. Case 2: b1 = c1 but [i; i+|v1|] and [c1, c1+|vA|] describe /disjoint/ areas of the same block. In this case, we can use the lemma store_value_load_disjoint on the hypothesis (load_value_of_type m1' c1 off1 = ⌊vA⌋) yielding (load_value_of_type m1' c1 off1 = load_value_of_type m1 c1 off1) allowing us to use the load_sim hypothesis on (m1, m2) to obtain (load_value_of_type m2 c2 (off1+off2) = ⌊vB⌋) which we can lift back to (load_value_of_type m2' c2 (off1+off2) = ⌊vB⌋) using the disjointness hypothesis contained in the original memory injection [Hinj]. case 4: b1 ≠ c1, nothing difficult case 3: the intervals [i; i+|v1|] and [c1, c1+|vA|] /overalap/ ! in CompCert, the prof follows simply from the fact that the load (load_value_of_type m1' c1 off1) will fail because of invalidated memory, yielding a contradiction. We have no such possibility in Cerco. I see a possible, convoluted way out of this problem. In the case of an overlap, we need to proceed to an inversion on load_value_of_type m1' c1 off1 = ⌊vA⌋ and to actually look at the data beeing written. If we succeeded in proceeding to this load, this means that the back-end values stored are "consistent". Besides, we need to proceed to a case analysis on what we stored beforehand. A) If we stored an integer . of size 8: this means that the overlap actually includes all of the freshly stored area. This area contains one [BVByte]. If we succeeded in performing an overlapping load, we either overlapped from the beginning, from the end or from both ends of the stored area. In all cases, since this area contains a BVByte, all other offsets MUST contain BVBytes, otherwise we would have a contradiction ... (cumbersome to show but possible) Hence, we can splice the load in 2 or 3 sub-loads : . re-use the Case 2 above (disjoint areas) for the parts of the load that are before and after the stored area . re-use the Case 1 above for the stored area and show the whole successful load . of size 16,32: we have more freedom but only a finite number of overlap possibilities in each case. We can enumerate them and proceed along the same lines as in the 8 bit case. B) If we stored a pointer (of size 2 bytes, necessarily) . the bytes of a pointer are stored in order and (/!\ important /!\) are /numbered/. This means that any successful overlapping load has managed to read a pointer in the wrong order, which yields a contradiction. C) If we stored a Vnull, roughly the same reasoning as in the pointer case D) If we stored a Vundef ... This gets hairy. We must reason on the way fe_to_be_values and be_to_fe_value works. What we want is to show that the load [load_value_of_a type m1' c1 off1] yields a Vundef. This can happen in several ways. . If we try to load something of integer type, we will have a Vundef because we can show that some of the back-end values are BVundef (cf [build_integer]) . If we try to load a pointer, it will also fail for the same reason. I'll admit that I'm not too sure about the last case. Also, my reasoning on pointers might fail if we have "longer that 2 bytes" pointers. This was a high-level and sketchy proof, and in the interests of time I decide to axiomatize it. *) (* This is lemma 60 from Leroy&Blazy *) lemma store_value_of_type_parallel : ∀E,m1,m2. ∀Hinj:memory_inj E m1 m2. ∀v1,v2. value_eq E v1 v2 → ∀b1,b2,delta. E b1 = Some ? 〈b2, delta〉 → ∀ty,i,m1'. type_consistent_with_value ty v1 → store_value_of_type ty m1 b1 i v1 = Some ? m1' → ∃m2'. store_value_of_type ty m2 b2 (offset_plus i delta) v2 = Some ? m2' ∧ memory_inj E m1' m2'. #E #m1 #m2 #Hinj #v1 #v2 #Hvalue_eq #b1 #b2 #delta #Hembed #ty #i #m1' cases ty [ | #sz' #sg' | #ptr_ty' | #array_ty' #array_sz' | #domain' #codomain' | #structname' #fieldspec' | #unionname' #fieldspec' | #id' ] #Hconsistent whd in ⊢ ((??%?) → ?); [ 1,4,5,6,7: #Habsurd destruct ] whd in match (store_value_of_type ?????); @(storen_parallel_aux … Hinj … Hembed … Hconsistent) // qed. lemma store_value_of_type_load_sim : ∀E,m1,m2,m1'. ∀Hinj:memory_inj E m1 m2. ∀ty,b,i,v. E b = None ? → store_value_of_type ty m1 b i v = Some ? m1' → load_sim_ptr E m1' m2. #E #m1 #m2 #m1' #Hinj #ty #b #i #v #Hembed #Hstore cases Hinj #Hsim #Hvalid #Hregion_eq #Hnonalias cases ty in Hstore; [ | #sz #sg | #ptr_ty | #array_ty #array_sz | #domain #codomain | #structname #fieldspec | #unionname #fieldspec | #id ] [ 1,4,5,6,7: #Heq normalize in Heq; destruct ] #Hstore whd #b1 #off1 #b2 #off2 #ty' #v1 #Hvalid #Hembed' #Hload lapply (store_value_of_type_load_value_of_type_conservation … Hstore b1 off1 ? ty') [ 1,3,5: @(eq_block_elim … b b1) try // #Heq >Heq in Hembed; #Hembed whd in Hembed':(??%?); >Hembed in Hembed'; #H normalize in H; destruct ] #Heq 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: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 1,3,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: | 2: #sz' #i'| 3: | 4: #p1' #p2' #Hembed' ] [ 1,3,4,5,7: #Habsurd destruct (Habsurd) ] #Heq >Heq %{r1} @conj // /3 by ex_intro, conj, vfloat_eq/ *) | 3: whd in match (sem_add ????); whd in match (sem_add ????); normalize nodelta cases (classify_add ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 1,3,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: | 2: #sz' #i' | 3: | 4: #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) ] | 4: whd in match (sem_add ????); whd in match (sem_add ????); normalize nodelta cases (classify_add ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 1,3,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1: | 2: #sz' #i' | 3: | 4: #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) [ | #sz #i | 3: | 4: #p1 #p2 #Hembed ] [ 1: whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ #sz #sg | #n #ty #sz #sg | #n #sz #sg #ty |#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: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 2,3,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ | #sz' #i' | | #p1' #p2' #Hembed' ] [ 1,3,4: #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: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 1,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ | #sz' #i' | | #p1' #p2' #Hembed' ] [ 1,2,4,5: #Habsurd destruct (Habsurd) ] #Heq destruct (Heq) /3 by ex_intro, conj, vfloat_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: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 1,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1,5: | 2,6: #sz' #i' | 3,7: | 4,8: #p1' #p2' #Hembed' ] [ 1,2,4,5,7,8: #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/ ] | 4: whd in match (sem_sub ????); whd in match (sem_sub ????); normalize nodelta cases (classify_sub ty1 ty2) normalize nodelta [ 1: #tsz #tsg | 2: #tn #ty #tsz #tsg | 3: #tn #tsz #tsg #ty | 4: #ty1' #ty2' ] [ 1,4: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ 1,5: | 2,6: #sz' #i' | 3,7: | 4,8: #p1' #p2' #Hembed' ] [ 1,2,4,5,6,7: #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) [ | #sz #i | 3: | 4: #p1 #p2 #Hembed ] [ 1: whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #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: #ty1' #ty2' ] [ 2: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ | #sz' #i' | | #p1' #p2' #Hembed' ] [ 1,3,4: #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: #ty1' #ty2' ] [ 1,2: #Habsurd destruct (Habsurd) ] | 4: whd in match (sem_mul ????); whd in match (sem_mul ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #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) [ | #sz #i | 3: | 4: #p1 #p2 #Hembed ] [ 1: whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #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: #ty1' #ty2' ] [ 2: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ | #sz' #i' | | #p1' #p2' #Hembed' ] [ 1,3,4: #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: #ty1' #ty2' ] [ 1,2: #Habsurd destruct (Habsurd) ] | 4: whd in match (sem_div ????); whd in match (sem_div ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #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) [ | #sz #i | 3: | 4: #p1 #p2 #Hembed ] [ 1: whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta [ 1: #sz #sg | 2: #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: #ty1' #ty2' ] [ 2: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ | #sz' #i' | | #p1' #p2' #Hembed' ] [ 1,3,4: #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/ ] ] | *: whd in match (sem_mod ????); whd in match (sem_mod ????); normalize nodelta cases (classify_aop ty1 ty2) normalize nodelta #foo #bar #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) [ | #sz #i | 3: | 4: #p1 #p2 #Hembed ] [ 2: @(value_eq_inversion E … Hvalue_eq2) [ | #sz' #i' | | #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) [ | #sz #i | 3: | 4: #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: #ty1' #ty2' ] #Habsurd destruct (Habsurd) | 2: cases (classify_aop ty ty') normalize nodelta [ 1: #sz #sg | 2: #ty1' #ty2' ] [ 2: #Habsurd destruct (Habsurd) ] @(value_eq_inversion E … Hvalue_eq2) normalize nodelta [ | #sz' #i' | | #p1' #p2' #Hembed' ] [ 1,3,4: #Habsurd destruct (Habsurd) ] elim sg normalize nodelta cases (lt_u ???) normalize nodelta #Heq destruct (Heq) /3 by ex_intro, conj, refl, vint_eq/ | *: cases (classify_aop ty ty') normalize nodelta #foo #bar #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). (* 5 Nov. 2012 : the rest of the stuff doesn't checks anymore. Not a big deal, it was here * for historical purposes, and for potential inspiration for Clight to Cminor. I leave the * code commented for future reference. *) (* 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. *)