diff --git a/floyd/efield_lemmas.v.crashcoqide b/floyd/efield_lemmas.v.crashcoqide deleted file mode 100644 index 2245a88f3..000000000 --- a/floyd/efield_lemmas.v.crashcoqide +++ /dev/null @@ -1,1459 +0,0 @@ -Require Import VST.floyd.base2. -Require Import VST.floyd.client_lemmas. -Require Import VST.floyd.nested_pred_lemmas. -Require Import VST.floyd.nested_field_lemmas. -Require Import VST.floyd.fieldlist. -Import LiftNotation. -Import -(notations) compcert.lib.Maps. -(* Local Open Scope logic. *) - -Inductive efield : Type := - | eArraySubsc: forall i: expr, efield - | eStructField: forall i: ident, efield - | eUnionField: forall i: ident, efield. - -Section CENV. - -Context `{!VSTGS OK_ty Σ} {cs: compspecs}. - -Fixpoint nested_efield (e: expr) (efs: list efield) (tts: list type) : expr := - match efs, tts with - | nil, _ => e - | _, nil => e - | cons ef efs', cons t0 tts' => - match ef with - | eArraySubsc ei => Ederef (Ebinop Cop.Oadd (nested_efield e efs' tts') ei (tptr t0)) t0 - | eStructField i => Efield (nested_efield e efs' tts') i t0 - | eUnionField i => Efield (nested_efield e efs' tts') i t0 - end - end. - -Inductive array_subsc_denote {cs: compspecs}: expr -> Z -> environ -> Prop := - | array_subsc_denote_intro_int: - forall e i rho, Vint (Int.repr i) = eval_expr e rho -> array_subsc_denote e i rho - | array_subsc_denote_intro_long: - forall e i rho, Vlong (Int64.repr i) = eval_expr e rho -> array_subsc_denote e i rho. - -Inductive efield_denote {cs: compspecs}: list efield -> list gfield -> environ -> Prop := - | efield_denote_nil: forall rho, efield_denote nil nil rho - | efield_denote_ArraySubsc_int: forall ei efs i gfs rho, - match typeconv (typeof ei) with - | Tint _ Signed _ => Int.min_signed <= i <= Int.max_signed - | Tint _ Unsigned _ => 0 <= i <= Int.max_unsigned - | _ => False - end -> - array_subsc_denote ei i rho -> - efield_denote efs gfs rho -> - efield_denote (eArraySubsc ei :: efs) (ArraySubsc i :: gfs) rho - | efield_denote_ArraySubsc_long: forall ei efs i gfs rho, - is_long_type (typeof ei) = true -> - array_subsc_denote ei i rho -> - efield_denote efs gfs rho -> - efield_denote (eArraySubsc ei :: efs) (ArraySubsc i :: gfs) rho - | efield_denote_ArraySubsc_ptrofs: forall ei efs i gfs rho, - is_ptrofs_type (typeof ei) = true -> - array_subsc_denote ei i rho -> - efield_denote efs gfs rho -> - efield_denote (eArraySubsc ei :: efs) (ArraySubsc i :: gfs) rho - | efield_denote_StructField: forall i efs gfs rho, - efield_denote efs gfs rho -> - efield_denote (eStructField i :: efs) (StructField i :: gfs) rho - | efield_denote_UnionField: forall i efs gfs rho, - efield_denote efs gfs rho -> - efield_denote (eUnionField i :: efs) (UnionField i :: gfs) rho. - -Fixpoint typecheck_efield {cs: compspecs} (Delta: tycontext) (efs: list efield) : tc_assert := - match efs with - | nil => tc_TT - | eArraySubsc ei :: efs' => - tc_andp (typecheck_expr Delta ei) (typecheck_efield Delta efs') - | eStructField i :: efs' => - typecheck_efield Delta efs' - | eUnionField i :: efs' => - typecheck_efield Delta efs' - end. - -Definition tc_efield {cs: compspecs} (Delta: tycontext) (efs: list efield) := denote_tc_assert (typecheck_efield Delta efs). - -Definition typeconv' (ty: type): type := -match ty with -| Tvoid => remove_attributes ty -| Tint I8 _ _ => Tint I32 Signed noattr -| Tint I16 _ _ => Tint I32 Signed noattr -| Tint I32 _ _ => remove_attributes ty -| Tint IBool _ _ => Tint I32 Signed noattr -| Tlong _ _ => remove_attributes ty -| Tfloat _ _ => remove_attributes ty -| Tpointer _ _ => if eqb_type ty int_or_ptr_type then ty else remove_attributes ty -| Tarray t _ _ => Tpointer t noattr -| Tfunction _ _ _ => Tpointer ty noattr -| Tstruct _ _ => remove_attributes ty -| Tunion _ _ => remove_attributes ty -end. - -(* Null Empty Path situation *) -Definition type_almost_match e t lr:= - match typeof e, t, lr with - | _, Tarray t1 _ a1, RRRR => eqb_type (typeconv' (typeof e)) (Tpointer t1 noattr) - | _, _, LLLL => eqb_type (typeof e) t - | _, _, _ => false - end. - -(* TODO: remove almost_match' and use "type_is_by_value" in proof for assistent. *) -(* Empty Path situation *) -Definition type_almost_match' e t lr:= - match typeof e, t, lr with - | _, _, LLLL => eqb_type (typeof e) t - | _, _, _ => false - end. - -Fixpoint legal_nested_efield_rec t_root (gfs: list gfield) (tts: list type): bool := - match gfs, tts with - | nil, nil => true - | nil, _ => false - | _ , nil => false - | gf :: gfs0, t0 :: tts0 => (legal_nested_efield_rec t_root gfs0 tts0 && eqb_type (nested_field_type t_root gfs) t0)%bool - end. - -Definition legal_nested_efield t_root e gfs tts lr := - (match gfs with - | nil => type_almost_match' e t_root lr - | _ => type_almost_match e t_root lr - end && - legal_nested_efield_rec t_root gfs tts)%bool. - -Lemma legal_nested_efield_rec_cons: forall t_root gf gfs t tts, - legal_nested_efield_rec t_root (gf :: gfs) (t :: tts) = true -> - legal_nested_efield_rec t_root gfs tts = true. -Proof. - intros. - simpl in H. - rewrite andb_true_iff in H. - tauto. -Qed. - -Lemma typeconv_typeconv'_eq: forall t1 t2, - typeconv' t1 = typeconv' t2 -> - typeconv t1 = typeconv t2. -Proof. - intros. - destruct t1 as [| [| | |] | [|] | [|] | | | | |], t2 as [| [| | |] | [|] | [|] | | | | |]; simpl in *; - do 2 try match type of H with context [if ?A then _ else _] => destruct A end; congruence. -Qed. - -Lemma tc_efield_ind: forall {cs: compspecs} (Delta: tycontext) (efs: list efield) (rho: environ), - tc_efield Delta efs rho ⊣⊢ - match efs with - | nil => True - | eArraySubsc ei :: efs' => - tc_expr Delta ei rho ∧ tc_efield Delta efs' rho - | eStructField i :: efs' => - tc_efield Delta efs' rho - | eUnionField i :: efs' => - tc_efield Delta efs' rho - end. -Proof. - intros. - destruct efs; auto. - destruct e; auto. - unfold tc_efield. - simpl typecheck_efield. - rewrite denote_tc_assert_andp. - constructor; intros; monPred.unseal. (* FIXME is this necessary? *) - reflexivity. -Qed. - -Lemma typeof_nested_efield': forall rho t_root e ef efs gf gfs t tts, - legal_nested_efield_rec t_root (gf :: gfs) (t :: tts) = true -> - efield_denote (ef :: efs) (gf :: gfs) rho -> - nested_field_type t_root (gf :: gfs) = typeof (nested_efield e (ef :: efs) (t :: tts)). -Proof. - intros. - simpl in H. - rewrite andb_true_iff in H; destruct H. - apply eqb_type_true in H1; subst. - destruct ef; reflexivity. -Qed. - -Lemma typeof_nested_efield: forall rho t_root e efs gfs tts lr, - legal_nested_efield t_root e gfs tts lr = true -> - efield_denote efs gfs rho -> - nested_field_type t_root gfs = typeof (nested_efield e efs tts). -Proof. - intros. - unfold legal_nested_efield in H. - rewrite andb_true_iff in H. - destruct H. - inversion H0; subst; destruct tts; - try solve [inversion H1 | simpl; auto | destruct e0; simpl; auto]. - + destruct lr; try discriminate. - apply eqb_type_true in H; subst. - reflexivity. - + eapply typeof_nested_efield'; eauto. - + eapply typeof_nested_efield'; eauto. - + eapply typeof_nested_efield'; eauto. - + eapply typeof_nested_efield'; eauto. - + eapply typeof_nested_efield'; eauto. -Qed. - -Lemma offset_val_sem_add_pi: forall ofs t0 si e rho i, - match si with - | Signed => Int.min_signed <= i <= Int.max_signed - | Unsigned => 0 <= i <= Int.max_unsigned - end -> - offset_val ofs - (force_val (Cop.sem_add_ptr_int _ t0 si (eval_expr e rho) (Vint (Int.repr i)))) = - offset_val ofs - (offset_val (sizeof t0 * i) (eval_expr e rho)). -Proof. - intros. - destruct (eval_expr e rho); try reflexivity. - rewrite sem_add_pi_ptr; auto. - apply I. -Qed. - -Lemma By_reference_eval_expr: forall Delta e rho, - access_mode (typeof e) = By_reference -> - tc_environ Delta rho -> - tc_lvalue Delta e rho ⊢ - ⌜ (eval_expr e rho = eval_lvalue e rho) ⌝. -Proof. - intros. - iIntros "H". - iPoseProof (typecheck_lvalue_sound with "[-]") as "%HH"; eauto. - iPureIntro. - destruct e; try contradiction; simpl in *; - reflexivity. -Qed. - -Lemma By_reference_tc_expr: forall Delta e rho, - access_mode (typeof e) = By_reference -> - tc_environ Delta rho -> - tc_lvalue Delta e rho ⊢ tc_expr Delta e rho. -Proof. - intros. - unfold tc_lvalue, tc_expr. - destruct e; ((iIntros (hyp); hnf in hyp; done) + - (constructor; intros; unfold typecheck_expr; rewrite H; done)). -Qed. - -Definition LR_of_type (t: type) := - match t with - | Tarray _ _ _ => RRRR - | _ => LLLL - end. - -Lemma legal_nested_efield_weaken: forall t_root e gfs tts, - legal_nested_efield t_root e gfs tts (LR_of_type t_root) = true -> - legal_nested_efield_rec t_root gfs tts = true /\ - type_almost_match e t_root (LR_of_type t_root) = true. -Proof. - intros. - unfold legal_nested_efield in H. - rewrite andb_true_iff in H. - split; [tauto |]. - destruct gfs; [| tauto]. - destruct H as [? _]. - unfold type_almost_match' in H. - unfold type_almost_match. - destruct (LR_of_type t_root), t_root, (typeof e); simpl in H |- *; - try inv H; auto. -Qed. - -Lemma weakened_legal_nested_efield_spec: forall t_root e gfs efs tts rho, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - efield_denote efs gfs rho -> - typeconv' (nested_field_type t_root gfs) = typeconv' (typeof (nested_efield e efs tts)). -Proof. - intros. - inversion H1; subst; destruct tts; try solve [inv H]. - + rewrite nested_field_type_ind. - simpl typeof. - unfold type_almost_match in H0. - destruct (LR_of_type t_root), t_root, (typeof e); try solve [inv H0]; auto; - try (apply eqb_type_spec in H0; rewrite H0; auto). - + f_equal. - eapply typeof_nested_efield'; eauto. - + f_equal. - eapply typeof_nested_efield'; eauto. - + f_equal. - eapply typeof_nested_efield'; eauto. - + f_equal. - eapply typeof_nested_efield'; eauto. - + f_equal. - eapply typeof_nested_efield'; eauto. -Qed. - - -Lemma classify_add_typeconv: forall t n a ty, - typeconv (Tarray t n a) = typeconv ty -> - Cop.classify_add ty = Cop.classify_add (Tpointer t a). -Proof. -intros. -simpl in H. -extensionality t2. -destruct ty; inv H. -destruct i; inv H1. -all: simpl; destruct (typeconv t2); auto. -Qed. - -Lemma isBinOpResultType_add_ptr_long: forall e t n a t0 ei, - is_long_type (typeof ei) = true -> - typeconv (Tarray t0 n a) = typeconv (typeof e) -> - complete_legal_cosu_type t0 = true -> - eqb_type (typeof e) int_or_ptr_type = false -> - isBinOpResultType Cop.Oadd e ei (tptr t) = tc_isptr e. -Proof. - intros. - unfold isBinOpResultType. - rewrite (classify_add_typeconv _ _ _ _ H0). - destruct (typeof ei); inv H. - apply complete_legal_cosu_type_complete_type in H1. - simpl. - try destruct i; rewrite H1; simpl tc_bool; cbv iota; - rewrite andb_false_r; simpl; rewrite tc_andp_TT2; - unfold tc_int_or_ptr_type; rewrite H2; simpl; auto. -Qed. - -Lemma isBinOpResultType_add_ptr_ptrofs: forall e t n a t0 ei, - is_ptrofs_type (typeof ei) = true -> - typeconv (Tarray t0 n a) = typeconv (typeof e) -> - complete_legal_cosu_type t0 = true -> - eqb_type (typeof e) int_or_ptr_type = false -> - isBinOpResultType Cop.Oadd e ei (tptr t) = tc_isptr e. -Proof. - intros. - unfold isBinOpResultType. - rewrite (classify_add_typeconv _ _ _ _ H0). - destruct (typeof ei); inv H. - apply complete_legal_cosu_type_complete_type in H1. - simpl. - try destruct i; rewrite H1; simpl tc_bool; cbv iota; - rewrite andb_false_r; simpl; rewrite tc_andp_TT2; - unfold tc_int_or_ptr_type; rewrite H2; simpl; auto. -Qed. - -Lemma isBinOpResultType_add_ptr: forall e t n a t0 ei, - is_int_type (typeof ei) = true -> - typeconv (Tarray t0 n a) = typeconv (typeof e) -> - complete_legal_cosu_type t0 = true -> - eqb_type (typeof e) int_or_ptr_type = false -> - isBinOpResultType Cop.Oadd e ei (tptr t) = tc_isptr e. -Proof. - intros. - unfold isBinOpResultType. - rewrite (classify_add_typeconv _ _ _ _ H0). - destruct (typeof ei); inv H. - apply complete_legal_cosu_type_complete_type in H1. - simpl. - destruct i; rewrite H1; simpl tc_bool; cbv iota; - rewrite andb_false_r; simpl; rewrite tc_andp_TT2; - unfold tc_int_or_ptr_type; rewrite H2; simpl; auto. -Qed. - -Definition add_case_pptrofs t si := - if Archi.ptr64 then Cop.add_case_pl t else Cop.add_case_pi t si. - -Lemma array_op_facts_long: forall ei rho t_root e efs gfs tts t n a t0 p, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - is_long_type (typeof ei) = true -> - nested_field_type t_root gfs = Tarray t n a -> - field_compatible t_root gfs p -> - efield_denote efs gfs rho -> - (Cop.classify_add (typeof (nested_efield e efs tts)) (typeof ei) = Cop.add_case_pl t) /\ - isBinOpResultType Cop.Oadd (nested_efield e efs tts) ei (tptr t0) = tc_isptr (nested_efield e efs tts). -Proof. - intros. - pose proof (weakened_legal_nested_efield_spec _ _ _ _ _ _ H H0 H4). - rewrite H2 in H5. - split. - + - erewrite classify_add_typeconv - by (apply typeconv_typeconv'_eq; eassumption). - destruct (typeof ei); inv H1. - reflexivity. - + - eapply isBinOpResultType_add_ptr_long; [auto | apply typeconv_typeconv'_eq; eassumption | |]. - - destruct H3 as [_ [? [_ [_ ?]]]]. - eapply @nested_field_type_complete_legal_cosu_type with (gfs := gfs) in H3; auto. - rewrite H2 in H3. - exact H3. - - destruct (typeof (nested_efield e efs tts)); try solve [inv H5]; - apply eqb_type_false; try (unfold int_or_ptr_type; congruence). - Opaque eqb_type. simpl in H5. Transparent eqb_type. - destruct (eqb_type (Tpointer t1 a0) int_or_ptr_type) eqn:?H. - * apply eqb_type_true in H6. - unfold int_or_ptr_type in *; inv H5; inv H6. - * apply eqb_type_false in H6; auto. -Qed. - -Lemma array_op_facts_ptrofs: forall ei rho t_root e efs gfs tts t n a t0 p, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - is_ptrofs_type (typeof ei) = true -> - nested_field_type t_root gfs = Tarray t n a -> - field_compatible t_root gfs p -> - efield_denote efs gfs rho -> - (exists si, Cop.classify_add (typeof (nested_efield e efs tts)) (typeof ei) = add_case_pptrofs t si) /\ - isBinOpResultType Cop.Oadd (nested_efield e efs tts) ei (tptr t0) = tc_isptr (nested_efield e efs tts). -Proof. - intros. - pose proof (weakened_legal_nested_efield_spec _ _ _ _ _ _ H H0 H4). - rewrite H2 in H5. - split. - + - erewrite classify_add_typeconv - by (apply typeconv_typeconv'_eq; eassumption). - destruct (typeof ei); inv H1. - try (exists Unsigned; reflexivity); (* Archi.ptr64 = true *) - destruct i; simpl; eexists; reflexivity. (* Archi.ptr64 = false *) - + eapply isBinOpResultType_add_ptr_ptrofs; [auto | apply typeconv_typeconv'_eq; eassumption | |]. - - destruct H3 as [_ [? [_ [_ ?]]]]. - eapply @nested_field_type_complete_legal_cosu_type with (gfs := gfs) in H3; auto. - rewrite H2 in H3. - exact H3. - - destruct (typeof (nested_efield e efs tts)); try solve [inv H5]; - apply eqb_type_false; try (unfold int_or_ptr_type; congruence). - Opaque eqb_type. simpl in H5. Transparent eqb_type. - destruct (eqb_type (Tpointer t1 a0) int_or_ptr_type) eqn:?H. - * apply eqb_type_true in H6. - unfold int_or_ptr_type in *; inv H5; inv H6. - * apply eqb_type_false in H6; auto. -Qed. - -Lemma array_op_facts: forall ei rho t_root e efs gfs tts t n a t0 p, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - is_int_type (typeof ei) = true -> - nested_field_type t_root gfs = Tarray t n a -> - field_compatible t_root gfs p -> - efield_denote efs gfs rho -> - (exists si, Cop.classify_add (typeof (nested_efield e efs tts)) (typeof ei) = Cop.add_case_pi t si) /\ - isBinOpResultType Cop.Oadd (nested_efield e efs tts) ei (tptr t0) = tc_isptr (nested_efield e efs tts). -Proof. - intros. - pose proof (weakened_legal_nested_efield_spec _ _ _ _ _ _ H H0 H4). - rewrite H2 in H5. - split. - + - erewrite classify_add_typeconv - by (apply typeconv_typeconv'_eq; eassumption). - destruct (typeof ei); inv H1. - destruct i; simpl; eexists; reflexivity. - + eapply isBinOpResultType_add_ptr; [auto | apply typeconv_typeconv'_eq; eassumption | |]. - - destruct H3 as [_ [? [_ [_ ?]]]]. - eapply @nested_field_type_complete_legal_cosu_type with (gfs := gfs) in H3; auto. - rewrite H2 in H3. - exact H3. - - destruct (typeof (nested_efield e efs tts)); try solve [inv H5]; - apply eqb_type_false; try (unfold int_or_ptr_type; congruence). - Opaque eqb_type. simpl in H5. Transparent eqb_type. - destruct (eqb_type (Tpointer t1 a0) int_or_ptr_type) eqn:?H. - * apply eqb_type_true in H6. - unfold int_or_ptr_type in *; inv H5; inv H6. - * apply eqb_type_false in H6; auto. -Qed. - - -Lemma Ptrofs_repr_Int_signed_special: - Archi.ptr64=false -> forall i, Ptrofs.repr (Int.signed (Int.repr i)) = Ptrofs.repr i. -Proof. -intros. -apply Ptrofs.eqm_samerepr. -unfold Ptrofs.eqm. -rewrite (Ptrofs.modulus_eq32 H). -change (Zbits.eqmod Int.modulus (Int.signed (Int.repr i)) i). -rewrite Int.signed_repr_eq. -if_tac. -apply Zbits.eqmod_sym. -apply Zbits.eqmod_mod. -computable. -apply Zbits.eqmod_sym. -eapply Zbits.eqmod_trans. -apply Zbits.eqmod_mod. -computable. -rewrite <- (Z.sub_0_r (i mod Int.modulus)) at 1. -apply Zbits.eqmod_sub. -apply Zbits.eqmod_refl. -hnf. exists (-1). lia. -Qed. - -Lemma Ptrofs_repr_Int_unsigned_special: - Archi.ptr64=false -> forall i, Ptrofs.repr (Int.unsigned (Int.repr i)) = Ptrofs.repr i. -Proof. -intros. -pose proof (Ptrofs.agree32_repr H i). -hnf in H0. -rewrite <- H0. -apply Ptrofs.repr_unsigned. -Qed. - -Lemma Ptrofs_repr_Int64_unsigned_special: - Archi.ptr64=true -> forall i, Ptrofs.repr (Int64.unsigned (Int64.repr i)) = Ptrofs.repr i. -Proof. -intros. -pose proof (Ptrofs.agree64_repr H i). -hnf in H0. -rewrite <- H0. -apply Ptrofs.repr_unsigned. -Qed. - -Definition sem_add_ptr_ptrofs t si := - if Archi.ptr64 then sem_add_ptr_long t else sem_add_ptr_int t si. - -Lemma sem_add_pptrofs_ptr_special: - forall t si p i, - complete_type cenv_cs t = true -> - isptr p -> - sem_add_ptr_ptrofs t si p (Vptrofs (Ptrofs.repr i)) = Some (offset_val (sizeof t * i) p). -Proof. - intros. - unfold sem_add_ptr_ptrofs, sem_add_ptr_int, sem_add_ptr_long. - destruct p; try contradiction. - unfold offset_val, Cop.sem_add_ptr_long, Cop.sem_add_ptr_int. - unfold Vptrofs, Cop.ptrofs_of_int, Ptrofs.of_ints, Ptrofs.of_intu, Ptrofs.of_int. - rewrite H. - destruct Archi.ptr64 eqn:Hp. - f_equal. f_equal. f_equal. rewrite Ptrofs.of_int64_to_int64 //. - rewrite <- ptrofs_mul_repr; f_equal. - f_equal. f_equal. f_equal. - destruct si; - rewrite <- ?ptrofs_mul_repr; - rewrite ptrofs_to_int_repr; - rewrite ?Ptrofs_repr_Int_signed_special ?Ptrofs_repr_Int_unsigned_special //. -Qed. - -Lemma sem_add_pl_ptr_special: - forall t p i, - complete_type cenv_cs t = true -> - isptr p -> - sem_add_ptr_long t p (Vlong (Int64.repr i)) = Some (offset_val (sizeof t * i) p). -Proof. - intros. - unfold sem_add_ptr_long. - rewrite H. - destruct p; try contradiction. - unfold offset_val, Cop.sem_add_ptr_long. - f_equal. f_equal. f_equal. - rewrite <- ptrofs_mul_repr; f_equal. - unfold Ptrofs.of_int64. - clear. - apply Ptrofs.eqm_samerepr. - unfold Ptrofs.eqm. - apply Zbits.eqmod_divides with Int64.modulus. - fold (Int64.eqm (Int64.unsigned (Int64.repr i)) i). - apply Int64.eqm_sym. - apply Int64.eqm_unsigned_repr. - destruct Archi.ptr64 eqn:Hp. - rewrite Ptrofs.modulus_eq64 //. - rewrite Ptrofs.modulus_eq32 //; apply power_nat_divide; computable. -Qed. - - -Lemma sem_add_pi_ptr_special: - forall t p i si, - complete_type cenv_cs t = true -> - isptr p -> - match si with - | Signed => Int.min_signed <= i <= Int.max_signed - | Unsigned => 0 <= i <= Int.max_unsigned - end -> - sem_add_ptr_int t si p (Vint (Int.repr i)) = Some (offset_val (sizeof t * i) p). -Proof. - intros. - unfold sem_add_ptr_int. - rewrite H. - destruct p; try contradiction. - unfold offset_val, Cop.sem_add_ptr_int. - unfold Cop.ptrofs_of_int, Ptrofs.of_ints, Ptrofs.of_intu, Ptrofs.of_int. - f_equal. f_equal. f_equal. - destruct si; rewrite <- ptrofs_mul_repr; f_equal. - rewrite Int.signed_repr; auto. - rewrite Int.unsigned_repr; auto. -Qed. - -Lemma sem_add_pi_ptr_special': - Archi.ptr64 = false -> - forall t p i si, - complete_type cenv_cs t = true -> - isptr p -> - sem_add_ptr_int t si p (Vint (Int.repr i)) = Some (offset_val (sizeof t * i) p). -Proof. - intros Hp. - intros. - unfold sem_add_ptr_int. - rewrite H. - destruct p; try contradiction. - unfold offset_val, Cop.sem_add_ptr_int. - unfold Cop.ptrofs_of_int, Ptrofs.of_ints, Ptrofs.of_intu, Ptrofs.of_int. - f_equal. f_equal. f_equal. - destruct si; rewrite <- ptrofs_mul_repr; f_equal. - apply (Ptrofs_repr_Int_signed_special Hp). - apply (Ptrofs_repr_Int_unsigned_special Hp). -Qed. - -Lemma sem_add_pl_ptr_special': - Archi.ptr64 = true -> - forall t p i, - complete_type cenv_cs t = true -> - isptr p -> - sem_add_ptr_long t p (Vlong (Int64.repr i)) = Some (offset_val (sizeof t * i) p). -Proof. - intros Hp. - intros. - unfold sem_add_ptr_long. - rewrite H. - destruct p; try contradiction. - unfold offset_val, Cop.sem_add_ptr_long. - f_equal. f_equal. f_equal. - rewrite (Ptrofs.agree64_of_int_eq (Ptrofs.repr i)); [| (apply Ptrofs.agree64_repr; auto)]. - rewrite ptrofs_mul_repr. auto. -Qed. - -Tactic Notation "simpl!" := simpl; unfold typecheck_lvalue; unfold typecheck_expr; fold typecheck_lvalue; fold typecheck_expr; simpl. - -Lemma array_ind_step_long: forall Delta ei i rho t_root e efs gfs tts t n a t0 p, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - is_long_type (typeof ei) = true -> - array_subsc_denote ei i rho -> - nested_field_type t_root gfs = Tarray t0 n a -> - tc_environ Delta rho -> - efield_denote efs gfs rho -> - field_compatible t_root gfs p -> - (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ tc_efield Delta efs rho - ⊢ ⌜field_address t_root gfs p = eval_LR (nested_efield e efs tts) (LR_of_type (Tarray t0 n a)) rho⌝ ∧ - tc_LR_strong Delta (nested_efield e efs tts) (LR_of_type (Tarray t0 n a)) rho) -> - (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ - tc_efield Delta (eArraySubsc ei :: efs) rho - ⊢ ⌜offset_val (gfield_offset (nested_field_type t_root gfs) (ArraySubsc i)) - (field_address t_root gfs p) = - eval_lvalue (nested_efield e (eArraySubsc ei :: efs) (t :: tts)) rho⌝ ∧ - tc_lvalue Delta (nested_efield e (eArraySubsc ei :: efs) (t :: tts)) rho). -Proof. - intros ? ? ? ? ? ? ? ? ? ? ? ? ? ? - LEGAL_NESTED_EFIELD_REC TYPE_MATCH ? ? NESTED_FIELD_TYPE TC_ENVIRON EFIELD_DENOTE FIELD_COMPATIBLE IH. - destruct (array_op_facts_long _ _ _ _ _ _ _ _ _ _ t _ LEGAL_NESTED_EFIELD_REC TYPE_MATCH H NESTED_FIELD_TYPE FIELD_COMPATIBLE EFIELD_DENOTE) as [CLASSIFY_ADD ISBINOP]. - pose proof field_address_isptr _ _ _ FIELD_COMPATIBLE as ISPTR. - rewrite tc_efield_ind. - Opaque assert_of. simpl!. Transparent assert_of. - rewrite bi.and_comm. rewrite -bi.and_assoc. - iApply bi.wand_trans; iSplitL. - iApply bi.and_mono; [apply entails_refl | rewrite bi.and_comm; exact IH]. - rewrite (add_andp _ _ (typecheck_expr_sound _ _ _ TC_ENVIRON)). - unfold_lift. - normalize. - iIntros "[[%H1 %H2] H]". - iApply (andp_right1 with "H"). - + apply bi.pure_intro. - assert (H3: Vlong (Int64.repr i) = eval_expr ei rho). { - clear - H1 H0 H. - destruct (typeof ei); inv H. - inv H0. rewrite <- H in H1; inv H1. - rewrite <- H. f_equal. - } - rewrite <- H3. - unfold force_val2, force_val. - unfold sem_add. - rewrite CLASSIFY_ADD. - rewrite sem_add_pl_ptr_special. - 2:{ - clear - NESTED_FIELD_TYPE FIELD_COMPATIBLE. - assert (H := field_compatible_nested_field _ _ _ FIELD_COMPATIBLE). - rewrite NESTED_FIELD_TYPE in H. - destruct H as [_ [? _]]. - simpl in H. - apply complete_legal_cosu_type_complete_type; auto. - } - 2: simpl in H2; rewrite <- H2; auto. - unfold gfield_offset; rewrite NESTED_FIELD_TYPE H2. - reflexivity. - + normalize. - unfold tc_lvalue. - Opaque isBinOpResultType. - Opaque assert_of. simpl!. Transparent assert_of. (* To protect denote_tc_assert *) - Transparent isBinOpResultType. - rewrite ISBINOP. - rewrite !denote_tc_assert_andp. - rewrite !monPred_at_and. - repeat apply andp_right1. - - apply bi.pure_intro. - simpl in H2; rewrite <- H2; auto. - - solve_andp. - - solve_andp. - - rewrite andb_false_r. simpl. apply bi.pure_intro; auto. - - apply bi.pure_intro. - simpl; unfold_lift. - rewrite <- H3. - normalize. -Qed. - - -Lemma array_ind_step_ptrofs: forall Delta ei i rho t_root e efs gfs tts t n a t0 p, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - is_ptrofs_type (typeof ei) = true -> - array_subsc_denote ei i rho -> - nested_field_type t_root gfs = Tarray t0 n a -> - tc_environ Delta rho -> - efield_denote efs gfs rho -> - field_compatible t_root gfs p -> - (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ tc_efield Delta efs rho - ⊢ ⌜field_address t_root gfs p = eval_LR (nested_efield e efs tts) (LR_of_type (Tarray t0 n a)) rho⌝ ∧ - tc_LR_strong Delta (nested_efield e efs tts) (LR_of_type (Tarray t0 n a)) rho) -> - tc_LR_strong Delta e (LR_of_type t_root) rho ∧ - tc_efield Delta (eArraySubsc ei :: efs) rho - ⊢ ⌜offset_val (gfield_offset (nested_field_type t_root gfs) (ArraySubsc i)) - (field_address t_root gfs p) = - eval_lvalue (nested_efield e (eArraySubsc ei :: efs) (t :: tts)) rho⌝ ∧ - tc_lvalue Delta (nested_efield e (eArraySubsc ei :: efs) (t :: tts)) rho. -Proof. - intros ? ? ? ? ? ? ? ? ? ? ? ? ? ? - LEGAL_NESTED_EFIELD_REC TYPE_MATCH ? ? NESTED_FIELD_TYPE TC_ENVIRON EFIELD_DENOTE FIELD_COMPATIBLE IH. - destruct (array_op_facts_ptrofs _ _ _ _ _ _ _ _ _ _ t _ LEGAL_NESTED_EFIELD_REC TYPE_MATCH H NESTED_FIELD_TYPE FIELD_COMPATIBLE EFIELD_DENOTE) as [CLASSIFY_ADD ISBINOP]. - pose proof field_address_isptr _ _ _ FIELD_COMPATIBLE as ISPTR. - rewrite tc_efield_ind. - Opaque assert_of. simpl!. Transparent assert_of. - rewrite bi.and_comm -bi.and_assoc. - iApply bi.wand_trans; iSplitL. - iApply bi.and_mono; [apply entails_refl | rewrite bi.and_comm; exact IH]. - rewrite (add_andp _ _ (typecheck_expr_sound _ _ _ TC_ENVIRON)). - iIntros; iStopProof. - normalize. - unfold_lift. - apply andp_right1; [apply bi.pure_intro | normalize]. - + - assert (H3: Vptrofs (Ptrofs.repr i) = eval_expr ei rho). { - clear - H1 H0 H. - unfold is_ptrofs_type, Vptrofs in *. - destruct Archi.ptr64 eqn:Hp. - destruct (typeof ei); inversion H; clear H. - inversion H0; subst. rewrite <- H in H1; inv H1. - rewrite <- H. f_equal. apply Ptrofs.agree64_to_int_eq. - apply Ptrofs.agree64_repr; auto. - destruct (typeof ei); inversion H; clear H. destruct i0; inversion H3. - inversion H0. 2: rewrite <- H in H1; inv H1. - rewrite <- H. f_equal. apply ptrofs_to_int_repr. - } - unfold_lift. - rewrite <- H3. - unfold force_val2, force_val. - unfold sem_add. - destruct CLASSIFY_ADD as [si CLASSIFY_ADD]. - rewrite CLASSIFY_ADD. - match goal with |- _ = match ?A _ _ with Some _ => _ | None => _ end => - change A with (sem_add_ptr_ptrofs t0 si) - end. - rewrite sem_add_pptrofs_ptr_special. - 2:{ - clear - NESTED_FIELD_TYPE FIELD_COMPATIBLE. - assert (H := field_compatible_nested_field _ _ _ FIELD_COMPATIBLE). - rewrite NESTED_FIELD_TYPE in H. - destruct H as [_ [? _]]. - simpl in H. - apply complete_legal_cosu_type_complete_type; auto. - } - 2: simpl in H2; rewrite <- H2; auto. - unfold gfield_offset; rewrite NESTED_FIELD_TYPE H2. - reflexivity. - + normalize. - unfold tc_lvalue. - Opaque isBinOpResultType. - Opaque assert_of. simpl!. Transparent assert_of. - Transparent isBinOpResultType. - rewrite ISBINOP. - rewrite !denote_tc_assert_andp. - rewrite !monPred_at_and. - repeat apply andp_right1. - - apply bi.pure_intro. - simpl in H2; rewrite <- H2; auto. - - solve_andp. - - solve_andp. - - rewrite andb_false_r. simpl. apply bi.pure_intro; auto. - - apply bi.pure_intro. - simpl; unfold_lift. - rewrite <- H3. - normalize. -Qed. - -Lemma array_ind_step: forall Delta ei i rho t_root e efs gfs tts t n a t0 p, - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - match typeconv (typeof ei) with - | Tint _ Signed _ => Int.min_signed <= i <= Int.max_signed - | Tint _ Unsigned _ => 0 <= i <= Int.max_unsigned - | _ => False - end -> - array_subsc_denote ei i rho -> - nested_field_type t_root gfs = Tarray t0 n a -> - tc_environ Delta rho -> - efield_denote efs gfs rho -> - field_compatible t_root gfs p -> - (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ tc_efield Delta efs rho - ⊢ ⌜field_address t_root gfs p = eval_LR (nested_efield e efs tts) (LR_of_type (Tarray t0 n a)) rho⌝ ∧ - tc_LR_strong Delta (nested_efield e efs tts) (LR_of_type (Tarray t0 n a)) rho) -> - tc_LR_strong Delta e (LR_of_type t_root) rho ∧ - tc_efield Delta (eArraySubsc ei :: efs) rho - ⊢ ⌜offset_val (gfield_offset (nested_field_type t_root gfs) (ArraySubsc i)) - (field_address t_root gfs p) = - eval_lvalue (nested_efield e (eArraySubsc ei :: efs) (t :: tts)) rho⌝ ∧ - tc_lvalue Delta (nested_efield e (eArraySubsc ei :: efs) (t :: tts)) rho. -Proof. - intros ? ? ? ? ? ? ? ? ? ? ? ? ? ? - LEGAL_NESTED_EFIELD_REC TYPE_MATCH H ? NESTED_FIELD_TYPE TC_ENVIRON EFIELD_DENOTE FIELD_COMPATIBLE IH. - rename H into H'. - assert (H: is_int_type (typeof ei) = true) - by (clear - H'; destruct (typeof ei) as [| | | [|] | | | | |]; try contradiction; auto). - destruct (array_op_facts _ _ _ _ _ _ _ _ _ _ t _ LEGAL_NESTED_EFIELD_REC TYPE_MATCH H NESTED_FIELD_TYPE FIELD_COMPATIBLE EFIELD_DENOTE) as [CLASSIFY_ADD ISBINOP]. - pose proof field_address_isptr _ _ _ FIELD_COMPATIBLE as ISPTR. - rewrite tc_efield_ind. - Opaque assert_of. simpl!. Transparent assert_of. - rewrite bi.and_comm -bi.and_assoc. - iApply bi.wand_trans; iSplitL. - iApply bi.and_mono; [apply entails_refl | rewrite bi.and_comm; exact IH]. - iIntros; iStopProof. - rewrite (add_andp _ _ (typecheck_expr_sound _ _ _ TC_ENVIRON)). - unfold_lift. - normalize. - apply andp_right1; [apply bi.pure_intro | normalize]. - + - assert (H3: Vint (Int.repr i) = eval_expr ei rho). { - clear - H1 H0 H. - inv H0; auto. - destruct (typeof ei); inv H. rewrite <- H2 in H1. - destruct i0,s; contradiction. - } - rewrite <- H3. - unfold force_val2, force_val. - unfold sem_add. - destruct CLASSIFY_ADD as [si CLASSIFY_ADD]. - rewrite CLASSIFY_ADD. - rewrite sem_add_pi_ptr_special. - 2:{ - clear - NESTED_FIELD_TYPE FIELD_COMPATIBLE. - assert (H := field_compatible_nested_field _ _ _ FIELD_COMPATIBLE). - rewrite NESTED_FIELD_TYPE in H. - destruct H as [_ [? _]]. - simpl in H. - apply complete_legal_cosu_type_complete_type; auto. - } - 2: simpl in H2; rewrite <- H2; auto. - unfold gfield_offset; rewrite NESTED_FIELD_TYPE H2. - reflexivity. - clear - H' CLASSIFY_ADD. - destruct (typeof (nested_efield e efs tts)) as [ | [ | | | ] [ | ]| [ | ] | [ | ] | | | | | ], - (typeof ei) as [ | [ | | | ] [ | ]| [ | ] | [ | ] | | | | | ]; inv CLASSIFY_ADD; try contradiction; auto. - + normalize. - unfold tc_lvalue. - Opaque isBinOpResultType. - Opaque assert_of. simpl!. Transparent assert_of. - Transparent isBinOpResultType. - rewrite ISBINOP. - rewrite !denote_tc_assert_andp. - rewrite !monPred_at_and. - repeat apply andp_right1. - - apply bi.pure_intro. - simpl in H2; rewrite <- H2; auto. - - solve_andp. - - solve_andp. - - rewrite andb_false_r. simpl. apply bi.pure_intro; auto. - - apply bi.pure_intro. - simpl; unfold_lift. - rewrite <- H3. - normalize. -Qed. - -Lemma struct_op_facts: forall Delta t_root e gfs efs tts i a i0 t rho - (PLAIN: plain_members (co_members (get_co i)) = true), - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - in_members i0 (co_members (get_co i)) -> - nested_field_type t_root gfs = Tstruct i a -> - efield_denote efs gfs rho -> - tc_lvalue Delta (nested_efield e efs tts) rho = - tc_lvalue Delta (nested_efield e (eStructField i0 :: efs) (t :: tts)) rho /\ - eval_field (typeof (nested_efield e efs tts)) i0 = - offset_val (field_offset cenv_cs i0 (co_members (get_co i))). -Proof. - intros. - pose proof (weakened_legal_nested_efield_spec _ _ _ _ _ _ H H0 H3). - rewrite H2 in H4; simpl in H4. - destruct (typeof (nested_efield e efs tts)) eqn:?H; inv H4. - 1: destruct i1; inv H7. - 1: match type of H7 with context [if ?A then _ else _] => destruct A end; inv H7. - unfold tc_lvalue, eval_field. - Opaque assert_of. simpl!. Transparent assert_of. - rewrite H5. - unfold field_offset, fieldlist.field_offset. - unfold get_co in *. - destruct (cenv_cs !! i1); [| inv H1]. - rewrite (plain_members_field_offset _ PLAIN _ _ H1). - split; auto. - rewrite tc_andp_TT2. - reflexivity. -Qed. - -Lemma struct_ind_step: forall Delta t_root e gfs efs tts i a i0 t rho p - (PLAIN: plain_members (co_members (get_co i0)) = true), - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - in_members i (co_members (get_co i0)) -> - nested_field_type t_root gfs = Tstruct i0 a -> - tc_environ Delta rho -> - efield_denote efs gfs rho -> - field_compatible t_root gfs p -> - (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ tc_efield Delta efs rho - ⊢ ⌜(field_address t_root gfs (eval_LR e (LR_of_type t_root) rho) = - eval_LR (nested_efield e efs tts) (LR_of_type (Tstruct i0 a)) rho)⌝ ∧ - tc_LR_strong Delta (nested_efield e efs tts) (LR_of_type (Tstruct i0 a)) rho) -> - tc_LR_strong Delta e (LR_of_type t_root) rho ∧ - tc_efield Delta (eStructField i :: efs) rho - ⊢ ⌜(offset_val (gfield_offset (nested_field_type t_root gfs) (StructField i)) - (field_address t_root gfs (eval_LR e (LR_of_type t_root) rho)) = - eval_lvalue (nested_efield e (eStructField i :: efs) (t :: tts)) rho)⌝ ∧ - tc_lvalue Delta (nested_efield e (eStructField i :: efs) (t :: tts)) rho. -Proof. - intros ? ? ? ? ? ? ? ? ? ? ? ? PLAIN - LEGAL_NESTED_EFIELD_REC TYPE_MATCH ? NESTED_FIELD_TYPE TC_ENVIRON EFIELD_DENOTE FIELD_COMPATIBLE IH. - destruct (struct_op_facts Delta _ _ _ _ _ _ _ _ t _ PLAIN LEGAL_NESTED_EFIELD_REC TYPE_MATCH H NESTED_FIELD_TYPE EFIELD_DENOTE) as [TC EVAL]. - rewrite tc_efield_ind; simpl. - iApply bi.wand_trans; iSplitL; [iApply IH | ]. iIntros; iStopProof. - unfold_lift. - normalize. - apply andp_right1; [apply bi.pure_intro | normalize]. - + rewrite EVAL H0 NESTED_FIELD_TYPE. - reflexivity. - + simpl in TC; rewrite <- TC. - apply derives_refl. -Qed. - -Lemma union_op_facts: forall Delta t_root e gfs efs tts i a i0 t rho - (PLAIN: plain_members (co_members (get_co i)) = true), - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - in_members i0 (co_members (get_co i)) -> - nested_field_type t_root gfs = Tunion i a -> - efield_denote efs gfs rho -> - tc_lvalue Delta (nested_efield e efs tts) rho = - tc_lvalue Delta (nested_efield e (eUnionField i0 :: efs) (t :: tts)) rho /\ - eval_field (typeof (nested_efield e efs tts)) i0 = offset_val 0. -Proof. - intros. - pose proof (weakened_legal_nested_efield_spec _ _ _ _ _ _ H H0 H3). - rewrite H2 in H4; simpl in H4. - destruct (typeof (nested_efield e efs tts)) eqn:?H; inv H4. - 1: destruct i1; inv H7. - 1: match type of H7 with context [if ?A then _ else _] => destruct A end; inv H7. - unfold tc_lvalue, eval_field. - Opaque assert_of. simpl!. Transparent assert_of. - rewrite H5. - unfold get_co in *. - destruct (cenv_cs !! i1); [| inv H1]. - rewrite (plain_members_union_field_offset _ PLAIN); auto. - split; [| normalize; auto]. - rewrite tc_andp_TT2. - reflexivity. -Qed. - -Lemma union_ind_step: forall Delta t_root e gfs efs tts i a i0 t rho p - (PLAIN: plain_members (co_members (get_co i0)) = true), - legal_nested_efield_rec t_root gfs tts = true -> - type_almost_match e t_root (LR_of_type t_root) = true -> - in_members i (co_members (get_co i0)) -> - nested_field_type t_root gfs = Tunion i0 a -> - tc_environ Delta rho -> - efield_denote efs gfs rho -> - field_compatible t_root gfs p -> - (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ tc_efield Delta efs rho - ⊢ ⌜field_address t_root gfs (eval_LR e (LR_of_type t_root) rho) = - eval_LR (nested_efield e efs tts) (LR_of_type (Tstruct i0 a)) rho⌝ ∧ - tc_LR_strong Delta (nested_efield e efs tts) (LR_of_type (Tstruct i0 a)) rho) -> - tc_LR_strong Delta e (LR_of_type t_root) rho ∧ - tc_efield Delta (eUnionField i :: efs) rho - ⊢ ⌜offset_val (gfield_offset (nested_field_type t_root gfs) (UnionField i)) - (field_address t_root gfs (eval_LR e (LR_of_type t_root) rho)) = - eval_lvalue (nested_efield e (eUnionField i :: efs) (t :: tts)) rho⌝ ∧ - tc_lvalue Delta (nested_efield e (eUnionField i :: efs) (t :: tts)) rho. -Proof. - intros ? ? ? ? ? ? ? ? ? ? ? ? PLAIN - LEGAL_NESTED_EFIELD_REC TYPE_MATCH ? NESTED_FIELD_TYPE TC_ENVIRON EFIELD_DENOTE FIELD_COMPATIBLE IH. - destruct (union_op_facts Delta _ _ _ _ _ _ _ _ t _ PLAIN LEGAL_NESTED_EFIELD_REC TYPE_MATCH H NESTED_FIELD_TYPE EFIELD_DENOTE) as [TC EVAL]. - rewrite tc_efield_ind; simpl. - iApply bi.wand_trans; iSplitL; [iApply IH | ]. iIntros; iStopProof. - unfold_lift. - normalize. - apply andp_right1; [apply bi.pure_intro | normalize]. - + rewrite EVAL H0 NESTED_FIELD_TYPE. - reflexivity. - + simpl in TC; rewrite <- TC. - apply derives_refl. -Qed. - -Definition lvalue_LR_of_type: forall Delta rho P p t e, - t = typeof e -> - tc_environ Delta rho -> - (P ⊢ ⌜p = eval_lvalue e rho⌝ ∧ tc_lvalue Delta e rho) -> - P ⊢ ⌜p = eval_LR e (LR_of_type t) rho⌝ ∧ tc_LR_strong Delta e (LR_of_type t) rho. -Proof. - intros. - destruct (LR_of_type t) eqn:?H. - + exact H1. - + rewrite (add_andp _ _ H1); clear H1. - normalize. - iIntros "[_ ?]". - unfold LR_of_type in H2. - subst. - destruct (typeof e) eqn:?H; inv H2. - iSplit. - - iPoseProof (By_reference_eval_expr with "[-]") as "%HH"; try done. - rewrite H; auto. - - iApply By_reference_tc_expr; auto. - rewrite H; auto. -Qed. - - Lemma eval_lvalue_nested_efield_aux: forall Delta t_root e efs gfs tts p, - field_compatible t_root gfs p -> - legal_nested_efield t_root e gfs tts (LR_of_type t_root) = true -> - local (`(eq p) (eval_LR e (LR_of_type t_root))) ∧ - tc_LR Delta e (LR_of_type t_root) ∧ - local (tc_environ Delta) ∧ - tc_efield Delta efs ∧ - local (efield_denote efs gfs) ⊢ - local (`(eq (field_address t_root gfs p)) - (eval_LR (nested_efield e efs tts) (LR_of_type (nested_field_type t_root gfs)))) ∧ - tc_LR_strong Delta (nested_efield e efs tts) (LR_of_type (nested_field_type t_root gfs)). -Proof. - (* Prepare *) - intros Delta t_root e efs gfs tts p FIELD_COMPATIBLE LEGAL_NESTED_EFIELD. - unfold local, lift1; split => rho; monPred.unseal. - unfold_lift. - normalize. - rename H0 into EFIELD_DENOTE, H into TC_ENVIRON. - trans (tc_LR_strong Delta e (LR_of_type t_root) rho ∧ tc_efield Delta efs rho). - { - repeat (apply bi.and_mono; auto). - rewrite -tc_LR_tc_LR_strong. - auto. - } - pose proof legal_nested_efield_weaken _ _ _ _ LEGAL_NESTED_EFIELD as [LEGAL_NESTED_EFIELD_REC TYPE_ALMOST_MATCH]. - rewrite -> field_compatible_field_address by auto. - clear LEGAL_NESTED_EFIELD. - - (* Induction *) - revert tts LEGAL_NESTED_EFIELD_REC; induction EFIELD_DENOTE; intros; - destruct tts; try solve [inversion LEGAL_NESTED_EFIELD_REC]; - [normalize; rewrite bi.and_elim_l // | ..]; - pose proof FIELD_COMPATIBLE as FIELD_COMPATIBLE_CONS; - apply field_compatible_cons in FIELD_COMPATIBLE; - destruct (nested_field_type t_root gfs) eqn:NESTED_FIELD_TYPE; try solve [inv FIELD_COMPATIBLE]; - rename LEGAL_NESTED_EFIELD_REC into LEGAL_NESTED_EFIELD_REC_CONS; - pose proof (proj1 (proj1 (andb_true_iff _ _) LEGAL_NESTED_EFIELD_REC_CONS) : legal_nested_efield_rec t_root gfs tts = true) as LEGAL_NESTED_EFIELD_REC; - (spec IHEFIELD_DENOTE; [tauto |]); - (spec IHEFIELD_DENOTE; [auto |]); - specialize (IHEFIELD_DENOTE tts LEGAL_NESTED_EFIELD_REC); - (apply lvalue_LR_of_type; [eapply typeof_nested_efield'; eauto; econstructor; eauto | eassumption |]); - destruct FIELD_COMPATIBLE as [? FIELD_COMPATIBLE]; - rewrite -> offset_val_nested_field_offset_ind by auto; - rewrite <- field_compatible_field_address in IHEFIELD_DENOTE |- * by auto. - + eapply array_ind_step; eauto. - + eapply array_ind_step_long; eauto. - + eapply array_ind_step_ptrofs; eauto. - + eapply struct_ind_step; eauto. - destruct FIELD_COMPATIBLE as [_ [H0 [_ [_ H1]]]]. - assert (H2 :=nested_field_type_complete_legal_cosu_type _ _ H0 H1). - rewrite NESTED_FIELD_TYPE in H2. simpl in H2. - unfold get_co. - destruct (cenv_cs !! i0); try discriminate. - destruct (co_su c); try discriminate; auto. - + eapply union_ind_step; eauto. - destruct FIELD_COMPATIBLE as [_ [H0 [_ [_ H1]]]]. - assert (H2 :=nested_field_type_complete_legal_cosu_type _ _ H0 H1). - rewrite NESTED_FIELD_TYPE in H2. simpl in H2. - unfold get_co. - destruct (cenv_cs !! i0); try discriminate. - destruct (co_su c); try discriminate; auto. -Qed. - -Lemma nested_efield_facts: forall Delta t_root e efs gfs tts lr p, - field_compatible t_root gfs p -> - LR_of_type t_root = lr -> - legal_nested_efield t_root e gfs tts lr = true -> - type_is_by_value (nested_field_type t_root gfs) = true -> - local (`(eq p) (eval_LR e (LR_of_type t_root))) ∧ - tc_LR Delta e (LR_of_type t_root) ∧ - local (tc_environ Delta) ∧ - tc_efield Delta efs ∧ - local (efield_denote efs gfs) ⊢ - local (`(eq (field_address t_root gfs p)) - (eval_lvalue (nested_efield e efs tts))) ∧ - tc_lvalue Delta (nested_efield e efs tts). -Proof. - intros. - subst lr. - rewrite eval_lvalue_nested_efield_aux //. - destruct (LR_of_type (nested_field_type t_root gfs)) eqn:?H; auto; try apply derives_refl. - unfold LR_of_type in H0. - destruct (nested_field_type t_root gfs) as [| [| | |] [|] | | [|] | | | | |]; inv H2; inv H0. -Qed. - -Lemma eval_lvalue_nested_efield: forall Delta t_root e efs gfs tts lr p, - field_compatible t_root gfs p -> - LR_of_type t_root = lr -> - legal_nested_efield t_root e gfs tts lr = true -> - type_is_by_value (nested_field_type t_root gfs) = true -> - local (`(eq p) (eval_LR e lr)) ∧ - tc_LR Delta e lr ∧ - local (tc_environ Delta) ∧ - tc_efield Delta efs ∧ - local (efield_denote efs gfs) ⊢ - local (`(eq (field_address t_root gfs p)) (eval_lvalue (nested_efield e efs tts))). -Proof. - intros. - subst lr. - rewrite eval_lvalue_nested_efield_aux //. - rewrite bi.and_elim_l. - destruct (LR_of_type (nested_field_type t_root gfs)) eqn:?H; auto; try apply derives_refl. - unfold LR_of_type in H0. - destruct (nested_field_type t_root gfs) as [| [| | |] [|] | | [|] | | | | |]; inv H2; inv H0. -Qed. - -Lemma tc_lvalue_nested_efield: forall Delta t_root e efs gfs tts lr p, - field_compatible t_root gfs p -> - LR_of_type t_root = lr -> - legal_nested_efield t_root e gfs tts lr = true -> - type_is_by_value (nested_field_type t_root gfs) = true -> - local (`(eq p) (eval_LR e lr)) ∧ - tc_LR Delta e lr ∧ - local (tc_environ Delta) ∧ - tc_efield Delta efs ∧ - local (efield_denote efs gfs) ⊢ - tc_lvalue Delta (nested_efield e efs tts). -Proof. - intros. - subst lr. - rewrite eval_lvalue_nested_efield_aux //. - rewrite bi.and_elim_r. - destruct (LR_of_type (nested_field_type t_root gfs)) eqn:?H; auto; try apply derives_refl. - unfold LR_of_type in H0. - destruct (nested_field_type t_root gfs) as [| [| | |] [|] | | [|] | | | | |]; inv H2; inv H0. -Qed. - -Fixpoint compute_nested_efield_rec {cs:compspecs} e lr_default := - match e with - | Efield e' id t => - match typeof e' with - | Tstruct id_str _ => - if eqb_type (field_type id (co_members (get_co id_str))) t - then match compute_nested_efield_rec e' LLLL with - | (e'', efs, lr) => (e'', eStructField id :: efs, lr) - end - else (e, nil, lr_default) - | Tunion id_uni _ => - if eqb_type (field_type id (co_members (get_co id_uni))) t - then match compute_nested_efield_rec e' LLLL with - | (e'', efs, lr) => (e'', eUnionField id :: efs, lr) - end - else (e, nil, lr_default) - | _ => (e, nil, lr_default) - end - | Ederef (Ebinop Cop.Oadd e' ei (Tpointer t a)) t' => - match typeof e' with - | Tarray t'' _ _ => - match eqb_type t t'', eqb_type t t', eqb_attr a noattr with - | true, true, true => - match compute_nested_efield_rec e' RRRR with - | (e'', efs, lr) => (e'', eArraySubsc ei :: efs, lr) - end - | _, _, _ => (e, nil, lr_default) - end - | Tpointer t'' _ => - match eqb_type t t'', eqb_type t t', eqb_attr a noattr, eqb_type (typeof e') int_or_ptr_type with - | true, true, true, false => (e', eArraySubsc ei :: nil, RRRR) - | _, _, _, _ => (e, nil, lr_default) - end - | _ => (e, nil, lr_default) - end - | _ => (e, nil, lr_default) - end. - -Definition compute_nested_efield {cs: compspecs} (e: expr): expr * list efield * LLRR := compute_nested_efield_rec e LLLL. - -Inductive compute_root_type: forall (t_from_e: type) (lr: LLRR) (t_root: type), Prop := - | compute_root_type_lvalue: forall t, compute_root_type t LLLL t - | compute_root_type_Tpointer_expr: forall t a1 n a2, compute_root_type (Tpointer t a1) RRRR (Tarray t n a2) - | compute_root_type_Tarray_expr: forall t n1 a1 n2 a2, compute_root_type (Tarray t n1 a1) RRRR (Tarray t n2 a2). - -(* which means (e, lr) is possible to be called by compute_nested_efield_rec *) -Definition LR_possible (e: expr) (lr: LLRR) : bool := - match lr with - | LLLL => match (typeof e) with - | Tarray _ _ _ => false - | _ => true - end - | RRRR => match (typeof e) with - | Tarray _ _ _ => true - | _ => false - end - end. - -Definition array_relexed_type_eq (t1 t2: type): Prop := - match t1, t2 with - | Tarray t1' _ _, Tarray t2' _ _ => t1' = t2' - | _, _ => t1 = t2 - end. - -Lemma compute_nested_efield_trivial: forall e rho lr_default, - forall e_root efs lr, - e_root = e -> efs = nil -> lr = lr_default -> - LR_possible e lr_default = true -> - forall t_root gfs, - exists tts, - compute_root_type (typeof e_root) lr t_root -> - efield_denote efs gfs rho -> - nested_efield e_root efs tts = e /\ - LR_of_type t_root = lr /\ - type_almost_match e_root t_root lr = true /\ - legal_nested_efield_rec t_root gfs tts = true /\ - match gfs with - | nil => array_relexed_type_eq t_root (typeof e) - | _ => nested_field_type t_root gfs = typeof e - end. -Proof. - intros. - exists nil. - intros. - subst. - unfold LR_possible in H2. - unfold type_almost_match. - Opaque eqb_type. - destruct (typeof e); inv H2; inv H3; inv H4; simpl; - try rewrite eqb_type_spec; auto. -Qed. - -Lemma compute_nested_efield_aux: forall e rho lr_default, - (LR_possible e lr_default = true -> - match compute_nested_efield_rec e lr_default with - | (e_root, efs, lr) => - forall t_root gfs, - exists tts, - compute_root_type (typeof e_root) lr t_root -> - efield_denote efs gfs rho -> - nested_efield e_root efs tts = e /\ - LR_of_type t_root = lr /\ - type_almost_match e_root t_root lr = true /\ - legal_nested_efield_rec t_root gfs tts = true /\ - match gfs with - | nil => array_relexed_type_eq t_root (typeof e) - | _ => nested_field_type t_root gfs = typeof e - end - end) /\ - forall t, - (LR_possible (Ederef e t) lr_default = true -> - match compute_nested_efield_rec (Ederef e t) lr_default with - | (e_root, efs, lr) => - forall t_root gfs, - exists tts, - compute_root_type (typeof e_root) lr t_root -> - efield_denote efs gfs rho -> - nested_efield e_root efs tts = Ederef e t /\ - LR_of_type t_root = lr /\ - type_almost_match e_root t_root lr = true /\ - legal_nested_efield_rec t_root gfs tts = true /\ - match gfs with - | nil => array_relexed_type_eq t_root (typeof (Ederef e t)) - | _ => nested_field_type t_root gfs = typeof (Ederef e t) - end - end). -Proof. - intros ? ?. - induction e; intros ?; (split; [ | intros ?]); - try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl). - + destruct (IHe lr_default). apply (H0 t). - + destruct b, t; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl). - simpl. - destruct (typeof e1) eqn:?H; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); - destruct (eqb_type t t1) eqn:?H; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); - apply eqb_type_spec in H0; - destruct (eqb_type t t0) eqn:?H; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); - apply eqb_type_spec in H1; - destruct (eqb_attr a noattr) eqn:?H; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); - apply eqb_attr_spec in H2; - [destruct (eqb_type ((Tpointer t1 a0)) int_or_ptr_type) eqn:HH; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); apply eqb_type_false in HH |]. - - subst. - intros. - exists (t0 :: nil). - intros. - inv H1; inv H2. - * inv H9. - unfold type_almost_match. - rewrite H in H4 |- *; inv H4. - simpl. - change (nested_field_type (Tarray t0 n a2) (SUB i)) with t0. - apply eqb_type_false in HH. - rewrite HH. - rewrite !eqb_type_spec. - auto. - * inv H9. - unfold type_almost_match. - rewrite H in H4 |- *; inv H4. - simpl. - change (nested_field_type (Tarray t0 n a2) (SUB i)) with t0. - apply eqb_type_false in HH. - rewrite HH. - rewrite !eqb_type_spec. - auto. - * inv H9. - unfold type_almost_match. - rewrite H in H4 |- *; inv H4. - simpl. - change (nested_field_type (Tarray t0 n a2) (SUB i)) with t0. - apply eqb_type_false in HH. - rewrite HH. - rewrite !eqb_type_spec. - auto. - * inv H9. - unfold type_almost_match. - rewrite H in H4 |- *; inv H4. - * inv H9. - unfold type_almost_match. - rewrite H in H4 |- *; inv H4. - * inv H9. - unfold type_almost_match. - rewrite H in H4 |- *; inv H4. - - subst. - destruct (IHe1 RRRR) as [IH _]; spec IH; [unfold LR_possible; rewrite H; auto |]. - clear IHe1 IHe2. - destruct (compute_nested_efield_rec e1 RRRR) as ((?, ?), ?). - intros. - destruct gfs; [exists nil; intros _ HHH; inv HHH |]. - specialize (IH t_root gfs). - destruct IH as [tts IH]. - exists (t0 :: tts). - intros. - inv H2. - { - specialize (IH H1 H10). - destruct IH as [IH1 [IH2 [IH3 [IH4 IH5]]]]. - simpl. - rewrite IH1 IH4. - simpl. - rewrite eqb_type_spec. - assert (nested_field_type t_root (gfs SUB i) = t0); auto. - rewrite nested_field_type_ind; destruct gfs. - * destruct t_root; inv IH5; auto. - * rewrite IH5. auto. - }{ - specialize (IH H1 H10). - destruct IH as [IH1 [IH2 [IH3 [IH4 IH5]]]]. - simpl. - rewrite IH1 IH4. - simpl. - rewrite eqb_type_spec. - assert (nested_field_type t_root (gfs SUB i) = t0); auto. - rewrite nested_field_type_ind; destruct gfs. - * destruct t_root; inv IH5; auto. - * rewrite IH5. auto. - }{ - specialize (IH H1 H10). - destruct IH as [IH1 [IH2 [IH3 [IH4 IH5]]]]. - simpl. - rewrite IH1 IH4. - simpl. - rewrite eqb_type_spec. - assert (nested_field_type t_root (gfs SUB i) = t0); auto. - rewrite nested_field_type_ind; destruct gfs. - * destruct t_root; inv IH5; auto. - * rewrite IH5. auto. - } - + Opaque field_type. simpl. Transparent field_type. - destruct (typeof e) eqn:?H; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); - destruct (eqb_type (field_type i (co_members (get_co i0))) t) eqn:?H; try exact (compute_nested_efield_trivial _ _ _ _ _ _ eq_refl eq_refl eq_refl); - apply eqb_type_spec in H0. - - intros. - destruct (IHe LLLL) as [IH _]; clear IHe. - spec IH; [unfold LR_possible; rewrite H; auto |]. - destruct (compute_nested_efield_rec e LLLL) as ((?, ?), ?). - intros. - destruct gfs; [exists nil; intros _ HHH; inv HHH |]. - specialize (IH t_root gfs). - destruct IH as [tts IH]. - exists (t :: tts); intros. - revert H0; inv H3; intros. - specialize (IH H2 H8). - destruct IH as [IH1 [IH2 [IH3 [IH4 IH5]]]]. - simpl. - rewrite IH1 IH4. - simpl. - rewrite eqb_type_spec. - assert (nested_field_type t_root (gfs DOT i) = t); auto. - rewrite nested_field_type_ind; destruct gfs. - * destruct t_root; inv IH5; auto. - * rewrite IH5. auto. - - intros. - destruct (IHe LLLL) as [IH _]; clear IHe. - spec IH; [unfold LR_possible; rewrite H; auto |]. - destruct (compute_nested_efield_rec e LLLL) as ((?, ?), ?). - intros. - destruct gfs; [exists nil; intros _ HHH; inv HHH |]. - specialize (IH t_root gfs). - destruct IH as [tts IH]. - exists (t :: tts); intros. - revert H0; inv H3; intros. - specialize (IH H2 H8). - destruct IH as [IH1 [IH2 [IH3 [IH4 IH5]]]]. - simpl. - rewrite IH1 IH4. - simpl. - rewrite eqb_type_spec. - assert (nested_field_type t_root (gfs UDOT i) = t); auto. - rewrite nested_field_type_ind; destruct gfs. - * destruct t_root; inv IH5; auto. - * rewrite IH5. auto. -Qed. - -Lemma compute_nested_efield_lemma: forall e rho, - type_is_by_value (typeof e) = true -> - match compute_nested_efield e with - | (e_root, efs, lr) => - forall t_root gfs, - exists tts, - compute_root_type (typeof e_root) lr t_root -> - efield_denote efs gfs rho -> - nested_efield e_root efs tts = e /\ - LR_of_type t_root = lr /\ - legal_nested_efield t_root e_root gfs tts lr = true /\ - nested_field_type t_root gfs = typeof e - end. -Proof. - intros. - destruct (compute_nested_efield_aux e rho LLLL) as [? _]. - unfold compute_nested_efield. - destruct (compute_nested_efield_rec e LLLL) as ((?, ?), ?). - - intros. - spec H0; [unfold LR_possible; destruct (typeof e); inv H; auto |]. - specialize (H0 t_root gfs). - destruct H0 as [tts ?]. - exists tts. - intros. - specialize (H0 H1 H2). - destruct H0 as [? [? [? [? ?]]]]. - assert (nested_field_type t_root gfs = typeof e); - [| split; [| split; [| split]]; auto]. - + destruct gfs; auto. - destruct t_root, (typeof e); inv H6; auto; inv H. - + unfold legal_nested_efield. - rewrite H5. - rewrite H4. - destruct gfs; auto. - unfold type_almost_match', type_almost_match in *. - destruct l0, t_root; try rewrite H4; auto. - destruct tts; [| inv H5]. - inv H2. - rewrite <- H7 in H. - inv H. -Qed. - -End CENV. diff --git a/floyd/field_at.v.crashcoqide b/floyd/field_at.v.crashcoqide deleted file mode 100644 index 429e52627..000000000 --- a/floyd/field_at.v.crashcoqide +++ /dev/null @@ -1,3121 +0,0 @@ -Require Import VST.floyd.base2. -Require Import VST.floyd.client_lemmas. -Require Import VST.floyd.type_induction. -Require Import VST.floyd.nested_pred_lemmas. -Require Import VST.floyd.nested_field_lemmas. -Require Import VST.floyd.mapsto_memory_block. -Require Import VST.floyd.reptype_lemmas. -Require VST.floyd.aggregate_pred. Import VST.floyd.aggregate_pred.aggregate_pred. -Require Import VST.floyd.data_at_rec_lemmas. -Require Import VST.floyd.jmeq_lemmas. -Require Import VST.zlist.sublist. -Require Import VST.floyd.local2ptree_typecheck. -Import LiftNotation. - -Local Unset SsrRewrite. - -(************************************************ - -Definition of nested_reptype_structlist, field_at, array_at, data_at, nested_sfieldlist_at - -************************************************) - -Section CENV. - -Context `{!VSTGS OK_ty Σ} {cs: compspecs}. - -Lemma struct_Prop_cons2: - forall it it' m (A: member -> Type) - (P: forall it, A it -> Prop) - (v: compact_prod (map A (it::it'::m))), - struct_Prop (it :: it' :: m) P v = - (P _ (fst v) /\ struct_Prop (it'::m) P (snd v)). -Proof. -intros. -destruct v. -reflexivity. -Qed. - -Lemma struct_Prop_ext_derives: forall m {A0 A1} (P0: forall it, A0 it -> Prop) (P1: forall it, A1 it -> Prop) v0 v1, - members_no_replicate m = true -> - (forall i d0 d1, in_members i m -> - P0 _ (proj_struct i m v0 d0) -> P1 _ (proj_struct i m v1 d1)) -> - struct_Prop m P0 v0 -> struct_Prop m P1 v1. -Proof. - intros. revert H1. - destruct m as [| a0 m]; [simpl; auto |]. - revert a0 v0 v1 H H0; induction m as [| a1 m]; intros. - + specialize (H0 (name_member a0)). - simpl in H0. - unfold field_type, Ctypes.field_type in H0. - simpl in H0. - rewrite if_true in H0 by auto. - specialize (H0 v0 v1). - spec H0; [left; reflexivity |]. - destruct (member_dec a0 a0); [ | congruence]. - unfold eq_rect_r in H0; rewrite <- !eq_rect_eq in H0. - simpl. auto. - + revert H1. - change (struct_Prop (a0 :: a1 :: m) P0 v0) with - (P0 a0 (fst v0) /\ struct_Prop (a1 :: m) P0 (snd v0)). - change (struct_Prop (a0 :: a1 :: m) P1 v1) with - (P1 a0 (fst v1) /\ struct_Prop (a1 :: m) P1 (snd v1)). - intro. - rewrite fieldlist.members_no_replicate_ind in H. - destruct H as [H H']. - specialize (IHm a1 (snd v0) (snd v1) H'). - split. - - destruct H1 as [H1 _]; revert H1. - specialize (H0 (name_member a0)). - unfold proj_struct in H0. - revert H0; unfold field_type; simpl. - rewrite if_true by auto. - destruct (member_dec a0 a0); [ | congruence]. - unfold eq_rect_r; rewrite <- !eq_rect_eq. - intros. apply (H0 (fst v0) (fst v1)); auto. - hnf. left; reflexivity. - - destruct H1 as [_ H1]; revert H1. - apply IHm; clear IHm. - assert (name_member a0 <> name_member a1) by (contradict H; left; auto). - intros. - specialize (H0 i). - assert (i<> name_member a0). contradict H1. subst i. contradiction. - clear H H'. - assert (get_member i (a0::a1::m) = get_member i (a1::m)) - by (simpl; rewrite if_false; auto). - unfold proj_struct in *. - rewrite H in H0. - specialize (H0 d0 d1). - spec H0; [unfold in_members; right; auto | ]. - assert (proj_compact_prod (get_member i (a1 :: m)) - (a0 :: a1 :: m) v0 d0 member_dec = - proj_compact_prod (get_member i (a1:: m)) (a1 :: m) - (snd v0) d0 member_dec). - clear - H1 H4. - unfold proj_compact_prod. unfold list_rect; cbv beta iota. - destruct (member_dec (get_member i (a1 :: m)) a0). - exfalso. subst a0. rewrite name_member_get in H1, H4. contradiction. - reflexivity. - rewrite H5 in H0; clear H5. - assert (proj_compact_prod (get_member i (a1 :: m)) - (a0 :: a1 :: m) v1 d1 member_dec = - proj_compact_prod (get_member i (a1 :: m)) (a1 :: m) - (snd v1) d1 member_dec). - clear - H1 H4. - unfold proj_compact_prod. unfold list_rect; cbv beta iota. - destruct (member_dec (get_member i (a1 :: m)) a0). - exfalso. subst a0. rewrite name_member_get in H1, H4. contradiction. - reflexivity. - rewrite H5 in H0; clear H5. - apply H0; auto. -Qed. - -Lemma struct_Prop_ext: forall m {A0 A1} (P0: forall it, A0 it -> Prop) (P1: forall it, A1 it -> Prop) v0 v1, - members_no_replicate m = true -> - (forall i d0 d1, in_members i m -> - P0 _ (proj_struct i m v0 d0) = P1 _ (proj_struct i m v1 d1)) -> - struct_Prop m P0 v0 = struct_Prop m P1 v1. -Proof. - intros. - apply prop_ext; split; eapply struct_Prop_ext_derives; eauto; intros; revert H2; - erewrite H0 by auto; eauto. -Qed. - -Definition field_at (sh: Share.t) (t: type) (gfs: list gfield) (v: reptype (nested_field_type t gfs)) (p: val): mpred := - ⌜field_compatible t gfs p⌝ ∧ - at_offset (data_at_rec sh (nested_field_type t gfs) v) (nested_field_offset t gfs) p. -Arguments field_at sh t gfs v p : simpl never. - -Definition field_at_ (sh: Share.t) (t: type) (gfs: list gfield) (p: val): mpred := - field_at sh t gfs (default_val (nested_field_type t gfs)) p. - -Arguments field_at_ sh t gfs p : simpl never. - -Definition data_at (sh: Share.t) (t: type) (v: reptype t) := field_at sh t nil v. -Global Typeclasses Opaque data_at. - -Definition data_at_ (sh: Share.t) (t: type) := field_at_ sh t nil. - -Definition nested_reptype_structlist t gfs (m: members) := - compact_prod (map (fun it => reptype (nested_field_type t (StructField (name_member it) :: gfs))) m). - -Definition nested_reptype_unionlist t gfs (m: members) := - compact_sum (map (fun it => reptype (nested_field_type t (UnionField (name_member it) :: gfs))) m). - -Lemma map_members_ext: forall A (f f':member -> A) (m: list member), - members_no_replicate m = true -> - (forall i, in_members i m -> f (get_member i m) = f' (get_member i m)) -> - map f m = map f' m. -Proof. - intros. - induction m as [| a0 m]. - + reflexivity. - + simpl. - rewrite members_no_replicate_ind in H. - f_equal. - - specialize (H0 (name_member a0)). - unfold field_type, in_members in H0. - simpl in H0; if_tac in H0; [| congruence]. - apply H0; auto. - - apply IHm. tauto. - intros. - specialize (H0 i). - unfold in_members in H0. - simpl in H0; if_tac in H0; [subst; tauto |]. - apply H0; auto. -Defined. - -Lemma nested_reptype_structlist_lemma: forall t gfs id a, - nested_field_type t gfs = Tstruct id a -> - reptype (nested_field_type t gfs) = nested_reptype_structlist t gfs (co_members (get_co id)). -Proof. - intros. - rewrite H, reptype_eq. - unfold reptype_structlist, nested_reptype_structlist. - f_equal. - apply map_members_ext; [apply get_co_members_no_replicate |]. - intros. - rewrite nested_field_type_ind, H. - simpl. - auto. -Defined. - -Lemma nested_reptype_unionlist_lemma: forall t gfs id a, - nested_field_type t gfs = Tunion id a -> - reptype (nested_field_type t gfs) = nested_reptype_unionlist t gfs (co_members (get_co id)). -Proof. - intros. - rewrite H, reptype_eq. - unfold reptype_unionlist, nested_reptype_unionlist. - f_equal. - apply map_members_ext; [apply get_co_members_no_replicate |]. - intros. - rewrite nested_field_type_ind, H. - simpl. - auto. -Defined. - -Definition nested_sfieldlist_at sh t gfs m (v: nested_reptype_structlist t gfs m) p: mpred := - match m with - | nil => ⌜field_compatible t gfs p⌝ ∧ emp - | _ => struct_pred m (fun it v p => - withspacer sh - (nested_field_offset t gfs + - (field_offset cenv_cs (name_member it) m + sizeof (field_type (name_member it) m))) - (nested_field_offset t gfs + - field_offset_next cenv_cs (name_member it) m (sizeof (nested_field_type t gfs))) - (field_at sh t (StructField (name_member it) :: gfs) v) p) v p - end. - -Definition nested_ufieldlist_at sh t gfs m (v: nested_reptype_unionlist t gfs m) (p: val): mpred := - match m with - | nil => ⌜field_compatible t gfs p⌝ ∧ emp - | _ => union_pred m (fun it v p => - withspacer sh - (nested_field_offset t gfs + sizeof (field_type (name_member it) m)) - (nested_field_offset t gfs + sizeof (nested_field_type t gfs)) - (field_at sh t (UnionField (name_member it) :: gfs) v) p) v p - end. - -Definition array_at (sh: Share.t) (t: type) (gfs: list gfield) (lo hi: Z) - (v: list (reptype (nested_field_type t (ArraySubsc 0 :: gfs)))) (p: val) : mpred := - ⌜field_compatible0 t (ArraySubsc lo :: gfs) p /\ - field_compatible0 t (ArraySubsc hi :: gfs) p⌝ ∧ - array_pred lo hi - (fun i v => at_offset (data_at_rec sh (nested_field_type t (ArraySubsc 0 :: gfs)) v) - (nested_field_offset t (ArraySubsc i :: gfs))) v p. - -Definition array_at_ (sh: Share.t) (t: type) (gfs: list gfield) (lo hi: Z) : val -> mpred := - array_at sh t gfs lo hi (Zrepeat (default_val _) (hi-lo)). - -(************************************************ - -field_compatible, local_facts, isptr and offset_zero properties - -************************************************) - -Lemma field_at_local_facts: - forall sh t path v c, - field_at sh t path v c ⊢ ⌜field_compatible t path c /\ value_fits (nested_field_type t path) v⌝. -Proof. - intros. - unfold field_at, at_offset. - rewrite data_at_rec_value_fits. - by iIntros "(% & %)"; iPureIntro. -Qed. - -Lemma field_at_compatible': - forall sh t path v c, - field_at sh t path v c ⊣⊢ - ⌜field_compatible t path c⌝ ∧ field_at sh t path v c. -Proof. -intros. -iSplit; last by iIntros "(_ & $)". -rewrite bi.and_comm; iApply add_and. -rewrite field_at_local_facts. -normalize. -Qed. - -Lemma field_at__local_facts: forall sh t gfs p, - field_at_ sh t gfs p ⊢ ⌜field_compatible t gfs p⌝. -Proof. - intros. - unfold field_at_, field_at. - normalize. -Qed. - -Lemma data_at_local_facts: - forall sh t v p, data_at sh t v p ⊢ ⌜field_compatible t nil p /\ value_fits t v⌝. -Proof. intros. apply field_at_local_facts. Qed. - -Lemma data_at__local_facts: forall sh t p, data_at_ sh t p ⊢ ⌜field_compatible t nil p⌝. -Proof. intros. - apply field_at__local_facts. -Qed. - -Lemma array_at_local_facts: forall sh t gfs lo hi v p, - array_at sh t gfs lo hi v p ⊢ - ⌜field_compatible0 t (ArraySubsc lo :: gfs) p - /\ field_compatible0 t (ArraySubsc hi :: gfs) p - /\ Zlength v = hi - lo - /\ Forall (value_fits (nested_field_type t (ArraySubsc 0 :: gfs))) v⌝. -Proof. - intros. - unfold array_at. - rewrite array_pred_local_facts. - 2: { intros. - unfold at_offset. - apply data_at_rec_value_fits. } - normalize. -Qed. - -Lemma array_at__local_facts: forall sh t gfs lo hi p, - array_at_ sh t gfs lo hi p ⊢ - ⌜field_compatible0 t (ArraySubsc lo :: gfs) p - /\ field_compatible0 t (ArraySubsc hi :: gfs) p⌝. -Proof. - intros. - unfold array_at_. - rewrite array_at_local_facts; eauto. - apply bi.pure_mono; intuition. -Qed. - -Lemma field_at_isptr: forall sh t gfs v p, - field_at sh t gfs v p ⊣⊢ ⌜isptr p⌝ ∧ field_at sh t gfs v p. -Proof. intros. eapply local_facts_isptr; [apply field_at_local_facts | intros [? ?]; auto]. Qed. - -Lemma field_at_offset_zero: forall sh t gfs v p, - field_at sh t gfs v p ⊣⊢ field_at sh t gfs v (offset_val 0 p). -Proof. intros. apply local_facts_offset_zero. - intros. rewrite field_at_isptr; normalize. -Qed. - -Lemma field_at__isptr: forall sh t gfs p, - field_at_ sh t gfs p ⊣⊢ ⌜isptr p⌝ ∧ field_at_ sh t gfs p. -Proof. intros. - intros. eapply local_facts_isptr; [apply field_at__local_facts | intros [? ?]; auto]. -Qed. - -Lemma field_at__offset_zero: forall sh t gfs p, - field_at_ sh t gfs p ⊣⊢ field_at_ sh t gfs (offset_val 0 p). -Proof. intros. apply local_facts_offset_zero. - intros. rewrite field_at__isptr; normalize. -Qed. - -Lemma data_at_isptr: forall sh t v p, data_at sh t v p ⊣⊢ ⌜isptr p⌝ ∧ data_at sh t v p. -Proof. intros. eapply local_facts_isptr; [apply data_at_local_facts | intros [? ?]; auto]. -Qed. - -Lemma data_at_offset_zero: forall sh t v p, data_at sh t v p ⊣⊢ data_at sh t v (offset_val 0 p). -Proof. intros. rewrite <- local_facts_offset_zero. reflexivity. - intros; rewrite data_at_isptr; normalize. -Qed. - -Lemma data_at__isptr: forall sh t p, data_at_ sh t p ⊣⊢ ⌜isptr p⌝ ∧ data_at_ sh t p. -Proof. intros. eapply local_facts_isptr; [apply data_at__local_facts | intros [? ?]; auto]. -Qed. - -Lemma data_at__offset_zero: forall sh t p, data_at_ sh t p ⊣⊢ data_at_ sh t (offset_val 0 p). -Proof. intros. apply field_at__offset_zero. Qed. - -(************************************************ - -Ext lemmas of array_at - -************************************************) - -Lemma array_at_ext_derives: forall sh t gfs lo hi v0 v1 p, - Zlength v0 = Zlength v1 -> - (forall i u0 u1, - lo <= i < hi -> - JMeq u0 (Znth (i-lo) v0) -> - JMeq u1 (Znth (i-lo) v1) -> - field_at sh t (ArraySubsc i :: gfs) u0 p ⊢ - field_at sh t (ArraySubsc i :: gfs) u1 p) -> - array_at sh t gfs lo hi v0 p ⊢ array_at sh t gfs lo hi v1 p. -Proof. - intros until p. intro ZL; intros. - unfold array_at, field_at. - normalize. - eapply array_pred_ext_derives. - 1: intro; lia. - intros. - specialize (H i). - clear ZL. - revert v0 v1 H. - unfold field_at. - rewrite @nested_field_type_ArraySubsc with (i := i). - intros. - specialize (H (Znth (i - lo) v0) (Znth (i - lo) v1)). - do 3 (spec H; [auto |]). - rewrite !prop_true_andp in H by (apply (field_compatible_range _ lo hi); auto). - auto. -Qed. - -Lemma array_at_ext: forall sh t gfs lo hi v0 v1 p, - Zlength v0 = Zlength v1 -> - (forall i u0 u1, - lo <= i < hi -> - JMeq u0 (Znth (i-lo) v0) -> - JMeq u1 (Znth (i-lo) v1) -> - field_at sh t (ArraySubsc i :: gfs) u0 p ⊣⊢ - field_at sh t (ArraySubsc i :: gfs) u1 p) -> - array_at sh t gfs lo hi v0 p ⊣⊢ array_at sh t gfs lo hi v1 p. -Proof. - intros. - iSplit; iApply array_at_ext_derives; try done; intros; [rewrite H0 | rewrite <- H0]; done. -Qed. - -(************************************************ - -Unfold and split lemmas - -************************************************) - -Lemma field_at_Tarray: forall sh t gfs t0 n a v1 v2 p, - legal_nested_field t gfs -> - nested_field_type t gfs = Tarray t0 n a -> - 0 <= n -> - JMeq v1 v2 -> - field_at sh t gfs v1 p ⊣⊢ array_at sh t gfs 0 n v2 p. -Proof. - intros. - unfold field_at, array_at. - revert v1 v2 H2; - rewrite (nested_field_type_ind t (ArraySubsc 0 :: gfs)). - rewrite H0; unfold gfield_type. - intros. - rewrite data_at_rec_eq. - rewrite at_offset_array_pred. - apply bi.and_proper. - + f_equiv. - rewrite !field_compatible0_cons, H0. - assert (0 <= 0 <= n) by lia. - assert (0 <= n <= n) by lia. - tauto. - + apply (JMeq_trans (unfold_reptype_JMeq _ v1)) in H2. - forget (unfold_reptype v1) as v1'. - clear v1. - cbv iota beta in v1'. - apply JMeq_eq in H2. - rewrite Z.max_r by lia. - apply array_pred_ext. - - subst; auto. - - intros. - rewrite at_offset_eq. - rewrite <- at_offset_eq2. - rewrite !at_offset_eq. - rewrite (nested_field_offset_ind t (ArraySubsc i :: gfs)) - by (apply legal_nested_field0_field; simpl; unfold legal_field; rewrite H0; auto). - rewrite H0. - subst; auto. -Qed. - -Lemma not_ptr_False {prop:bi}: forall (A : prop) p, (A ⊢ ⌜isptr p⌝) <-> (~ isptr p -> A ⊣⊢ False). -Proof. - intros. - split; intros. - + iSplit; last by iIntros "[]". - rewrite H; iIntros (?); done. - + destruct (isptr_dec p); first by iIntros "_". - rewrite H; last done. - iIntros "[]". -Qed. - -Ltac solve_ptr_derives := - repeat rewrite isptr_offset_val; - apply derives_refl. - -Lemma field_at_isptr': - forall sh t path v c, field_at sh t path v c ⊢ ⌜isptr c⌝. -Proof. -intros. -rewrite field_at_local_facts. -iIntros "(($ & _) & _)". -Qed. - -Ltac solve_nptr p A := - let H := fresh "H" in - match A with - | (?B ∗ ?C) => - try solve [assert (~ isptr p -> B ⊣⊢ False) as H by solve_nptr p B; - intro; rewrite H by auto; apply bi.False_sep]; - try solve [assert (~ isptr p -> C ⊣⊢ False) as H by solve_nptr p C; - intro; rewrite H by auto; apply bi.sep_False] - | (?B ∧ ?C) => - try solve [assert (~ isptr p -> B ⊣⊢ False) as H by solve_nptr p B; - intro; rewrite H by auto; apply bi.False_and]; - try solve [assert (~ isptr p -> C ⊣⊢ False) as H by solve_nptr p C; - intro; rewrite H by auto; apply bi.and_False] - | _ => apply (proj1 (not_ptr_False A p)); solve_ptr p A - end -with solve_ptr p A := - let p0 := fresh "p" in - match A with - | (_ ∗ _) => apply (proj2 (not_ptr_False A p)); solve_nptr p A - | (_ ∧ _) => apply (proj2 (not_ptr_False A p)); solve_nptr p A - | ⌜_ /\ _⌝ => destruct A as [_ A]; solve_ptr p A - | ⌜field_compatible _ _ ?q⌝ => etrans; first apply (bi.pure_mono _ _ (field_compatible_isptr _ _ _)); solve_ptr_derives - | ⌜field_compatible0 _ _ ?q⌝ => etrans; first apply (bi.pure_mono _ _ (field_compatible0_isptr _ _ _)); solve_ptr_derives - | (memory_block _ _ ?q) => etrans; first apply (memory_block_local_facts _ _ _); solve_ptr_derives - | (withspacer _ _ _ ?P p) => apply withspacer_preserve_local_facts; - intro p0; solve_ptr p0 (P p0) - | (at_offset ?P _ ?q) => trans ⌜isptr q⌝; - [apply at_offset_preserve_local_facts; intro p0; solve_ptr p0 (P p0) | - solve_ptr_derives] - | (field_at _ _ _ _ p) => apply field_at_isptr' - end. - -Ltac destruct_ptr p := - let b := fresh "b" in - let ofs := fresh "OFS" in - match goal with - | |- ?A ⊣⊢ ?B => - let H := fresh "H" in - let H0 := fresh "H" in - assert (~ isptr p -> A ⊣⊢ False) as H by solve_nptr p A; - assert (~ isptr p -> B ⊣⊢ False) as H0 by solve_nptr p B; - destruct p as [| | | | | b ofs]; try (rewrite H, H0 by (simpl; congruence); reflexivity); - clear H H0; - inv_int ofs - | |- (?A ⊢ _) => - let H := fresh "H" in - assert (~ isptr p -> A ⊣⊢ False) as H by solve_nptr p A; - destruct p as [| | | | | b ofs]; try (rewrite H by (simpl; congruence); apply bi.False_elim); - clear H; - inv_int ofs - end. - -Lemma field_at_Tstruct: forall sh t gfs id a v1 v2 p, - nested_field_type t gfs = Tstruct id a -> - JMeq v1 v2 -> - field_at sh t gfs v1 p ⊣⊢ nested_sfieldlist_at sh t gfs (co_members (get_co id)) v2 p. -Proof. - intros. - unfold field_at, nested_sfieldlist_at. - revert v1 H0; rewrite H; intros. - rewrite data_at_rec_eq. - rewrite at_offset_struct_pred. - rewrite andp_struct_pred; [| apply _..]. - generalize (co_members (get_co id)) at 1 10; intro m; destruct m; [auto |]. - apply struct_pred_ext; [apply get_co_members_no_replicate |]. - - intros. - destruct_ptr p. - unfold field_at, fst, snd. - autorewrite with at_offset_db. - unfold offset_val. - solve_mod_modulus. - normalize. - destruct (legal_nested_field_dec t (StructField i :: gfs)). - 2:{ - assert (~field_compatible t gfs (Vptr b (Ptrofs.repr ofs))) - by (clear - n H H1; unfold field_compatible; contradict n; simpl; rewrite H; simpl; tauto). - assert (~field_compatible t - (gfs DOT name_member (get_member i (co_members (get_co id)))) - (Vptr b (Ptrofs.repr ofs))) - by (clear - n H H1; unfold field_compatible; simpl in *; rewrite H in *; simpl in *; tauto). - rewrite !prop_false_andp by auto; auto. - } - f_equiv. - { - f_equiv. - unfold field_compatible. - do 4 f_equiv. - simpl. - split; intro; try tauto. split; auto. - rewrite H. simpl. rewrite name_member_get. auto. - } - replace (field_offset cenv_cs (name_member (get_member i (co_members (get_co id))))) - with (field_offset cenv_cs i) - by (rewrite name_member_get; auto). - replace (field_offset_next cenv_cs (name_member (get_member i (co_members (get_co id))))) - with (field_offset_next cenv_cs i) - by (rewrite name_member_get; auto). - apply bi.sep_proper. - f_equiv. - rewrite name_member_get. - change (sizeof ?A) with (expr.sizeof A) in *. - rewrite sizeof_Tstruct. hnf; lia. - hnf; f_equal. f_equal. - rewrite name_member_get. lia. - match goal with |- data_at_rec _ _ _ ?A ⊣⊢ data_at_rec _ _ _ ?B => replace B with A end. - 2:{ f_equal. f_equal. - rewrite name_member_get. - rewrite @nested_field_offset_ind with (gfs := StructField i :: gfs) by auto. - unfold gfield_offset; rewrite H. lia. - } - erewrite data_at_rec_type_changable; first done. - { rewrite nested_field_type_ind. - simpl; rewrite H. - auto. } - apply (proj_compact_prod_JMeq _ (get_member i _) (co_members (get_co id)) _ _ (unfold_reptype v1) v2); auto. - * intros. - rewrite nested_field_type_ind, H. - unfold gfield_type. - rewrite In_field_type; auto. - apply get_co_members_no_replicate. - * apply in_get_member; auto. - * clear - H0. - eapply JMeq_trans; [apply (unfold_reptype_JMeq _ v1) | auto]. -Qed. - -Lemma field_at_Tunion: forall sh t gfs id a v1 v2 p, - nested_field_type t gfs = Tunion id a -> - JMeq v1 v2 -> - field_at sh t gfs v1 p ⊣⊢ nested_ufieldlist_at sh t gfs (co_members (get_co id)) v2 p. -Proof. - intros. - unfold field_at, nested_ufieldlist_at. - revert v1 H0; rewrite H; intros. - rewrite data_at_rec_eq. - rewrite at_offset_union_pred. - rewrite andp_union_pred; [| apply _..]. - generalize (eq_refl (co_members (get_co id))). - generalize (co_members (get_co id)) at 2 3 9; intro m; destruct m; [auto |]. - intro HH; assert (co_members (get_co id) <> nil) by congruence; clear HH. - apply union_pred_ext; [apply get_co_members_no_replicate | |]. - { - apply compact_sum_inj_JMeq; auto. - + intros. - rewrite nested_field_type_ind, H. - reflexivity. - + eapply JMeq_trans; [apply (unfold_reptype_JMeq _ v1) | auto]. - } - intros. - destruct_ptr p. - unfold field_at, fst, snd. - autorewrite with at_offset_db. - unfold offset_val. - solve_mod_modulus. - normalize. - destruct (legal_nested_field_dec t (UnionField i :: gfs)). - 2:{ - rewrite (bi.pure_False (field_compatible t (UnionField _ :: _) _)) - by (rewrite name_member_get; unfold field_compatible; tauto). - simpl in n. - rewrite H in n. - simpl in n. - rewrite bi.pure_False by (unfold field_compatible; tauto). - iSplit; iIntros "([] & ?)". - } - f_equiv. - apply bi.pure_iff. - rewrite name_member_get, field_compatible_cons, H; tauto. - apply bi.sep_proper. - rewrite name_member_get. - f_equiv. rewrite sizeof_Tunion. hnf; lia. - hnf; f_equal. f_equal. lia. - match goal with |- data_at_rec _ _ _ ?A ⊣⊢ data_at_rec _ _ _ ?B => replace B with A end. - 2:{ f_equal. f_equal. - rewrite name_member_get. - rewrite @nested_field_offset_ind with (gfs := UnionField i :: gfs) by auto. - unfold gfield_offset; rewrite H. lia. - } - erewrite data_at_rec_type_changable; first done. - rewrite name_member_get. - rewrite nested_field_type_ind. - rewrite H; reflexivity. - unfold proj_union. - apply (proj_compact_sum_JMeq _ (get_member i _) (co_members (get_co id)) d0 d1 (unfold_reptype v1) v2); auto. - * intros a0 ?. - rewrite nested_field_type_ind, H. - simpl. - auto. - * eapply JMeq_trans; [apply (unfold_reptype_JMeq _ v1) | auto]. -Qed. - -Lemma array_at_len_0: forall sh t gfs i p, - array_at sh t gfs i i nil p ⊣⊢ ⌜field_compatible0 t (ArraySubsc i :: gfs) p⌝ ∧ emp. -Proof. - intros. - unfold array_at. - rewrite array_pred_len_0 by lia. - apply bi.equiv_entails_2; normalize. -Qed. - -Lemma array_at_len_1: forall sh t gfs i v v' p, - JMeq v v' -> - array_at sh t gfs i (i + 1) (v :: nil) p ⊣⊢ field_at sh t (ArraySubsc i :: gfs) v' p. -Proof. - intros. - unfold array_at, field_at. - rewrite array_pred_len_1 by lia. - revert v' H. - rewrite @nested_field_type_ArraySubsc with (i := i). - intros. - apply JMeq_eq in H; rewrite H. - apply bi.and_proper; last done. - apply bi.pure_iff. - rewrite field_compatible_field_compatible0'. - reflexivity. -Qed. - -Lemma split2_array_at: forall sh t gfs lo mid hi v p, - lo <= mid <= hi -> - Zlength v = hi - lo -> - array_at sh t gfs lo hi v p ⊣⊢ - array_at sh t gfs lo mid (sublist 0 (mid-lo) v) p ∗ - array_at sh t gfs mid hi (sublist (mid-lo) (Zlength v) v) p. -Proof. - intros. - unfold array_at. - normalize. - apply andp_prop_ext. - + split; [| tauto]. - intros [? ?]. - assert (field_compatible0 t (gfs SUB mid) p) by (apply (field_compatible0_range _ lo hi); auto). - tauto. - + intros [? ?]. - rewrite @split_array_pred with (mid := mid) by auto. - rewrite H0; auto. -Qed. - -Lemma split3seg_array_at: forall sh t gfs lo ml mr hi v p, - lo <= ml -> - ml <= mr -> - mr <= hi -> - Zlength v = hi-lo -> - array_at sh t gfs lo hi v p ⊣⊢ - array_at sh t gfs lo ml (sublist 0 (ml-lo) v) p ∗ - array_at sh t gfs ml mr (sublist (ml-lo) (mr-lo) v) p ∗ - array_at sh t gfs mr hi (sublist (mr-lo) (hi-lo) v) p. -Proof. - intros. - rewrite split2_array_at with (lo := lo) (mid := ml) (hi := hi) by lia. - apply bi.sep_proper; first done. - assert (Zlength (sublist (ml - lo) (hi - lo) v) = hi - ml). - { - replace (hi - ml) with (hi - lo - (ml - lo)) by lia. - apply Zlength_sublist; lia. - } - rewrite H2. - rewrite split2_array_at with (lo := ml) (mid := mr) (hi := hi) by lia. - apply bi.sep_proper. - rewrite sublist_sublist; try lia. f_equiv. f_equal; lia. - rewrite Zlength_sublist by lia. - rewrite sublist_sublist; try lia. f_equiv. f_equal; lia. -Qed. - -Lemma split3_array_at: forall sh t gfs lo mid hi v v0 p, - lo <= mid < hi -> - Zlength v = hi-lo -> - JMeq v0 (Znth (mid-lo) v) -> - array_at sh t gfs lo hi v p ⊣⊢ - array_at sh t gfs lo mid (sublist 0 (mid-lo) v) p ∗ - field_at sh t (ArraySubsc mid :: gfs) v0 p ∗ - array_at sh t gfs (mid + 1) hi (sublist (mid+1-lo) (hi-lo) v) p. -Proof. - intros. - rename H0 into e; rename H1 into H0. - rewrite split3seg_array_at with (ml := mid) (mr := mid + 1) by lia. - apply bi.sep_proper; first done. - apply bi.sep_proper; last done. - replace (mid + 1 - lo) with (mid - lo + 1) by lia. - rewrite sublist_len_1 by lia. - rewrite array_at_len_1 with (v' :=v0); [auto |]. - apply JMeq_sym; auto. -Qed. - -(************************************************ - -Reroot lemmas - -************************************************) - -Lemma field_at_data_at: forall sh t gfs v (p: val), - field_at sh t gfs v p ⊣⊢ - data_at sh (nested_field_type t gfs) v (field_address t gfs p). -Proof. - intros. - unfold data_at, field_at. - rewrite (nested_field_offset_ind (nested_field_type t gfs) nil) by (simpl; tauto). - unfold field_address. - if_tac. - + unfold at_offset; normalize. - rewrite prop_true_andp; [auto |]. - destruct p; try (destruct H; contradiction). - generalize (field_compatible_nested_field t gfs (Vptr b i)); - unfold at_offset; solve_mod_modulus; intros. auto. - + apply bi.equiv_entails_2; normalize. destruct H0; contradiction. -Qed. - -Lemma field_at_data_at' : forall sh t gfs v p, field_at sh t gfs v p ⊣⊢ - ⌜field_compatible t gfs p⌝ ∧ - data_at sh (nested_field_type t gfs) v (offset_val (nested_field_offset t gfs) p). -Proof. - intros. - rewrite field_at_data_at. - unfold field_address. - if_tac. - - rewrite prop_true_andp; auto. - - rewrite prop_false_andp by auto. - rewrite data_at_isptr, prop_false_andp; auto. -Qed. - -Lemma field_at__data_at_: forall sh t gfs p, - field_at_ sh t gfs p ⊣⊢ - data_at_ sh (nested_field_type t gfs) (field_address t gfs p). -Proof. - intros. - unfold data_at_, field_at_. apply field_at_data_at. -Qed. - -Lemma lifted_field_at_data_at: forall sh t gfs v p, - assert_of (`(field_at sh t gfs) v p) ⊣⊢ - assert_of (`(data_at sh (nested_field_type t gfs)) v (`(field_address t gfs) p)). -Proof. - intros. - split => rho; unfold_lift; simpl. - apply field_at_data_at. -Qed. - -Lemma lifted_field_at__data_at_: forall sh t gfs p, - assert_of (`(field_at_ sh t gfs) p) ⊣⊢ - assert_of (`(data_at_ sh (nested_field_type t gfs)) (`(field_address t gfs) p)). -Proof. - intros. - split => rho; unfold_lift; simpl. - apply field_at__data_at_. -Qed. - -Lemma value_fits_JMeq: - forall t t' v v', - t=t' -> JMeq v v' -> value_fits t v -> value_fits t' v'. -Proof. -intros. subst. apply JMeq_eq in H0. subst. -auto. -Qed. - -Lemma array_at_data_at: forall sh t gfs lo hi v p, - lo <= hi -> - array_at sh t gfs lo hi v p ⊣⊢ - ⌜field_compatible0 t (ArraySubsc lo :: gfs) p⌝ ∧ - ⌜field_compatible0 t (ArraySubsc hi :: gfs) p⌝ ∧ - at_offset (data_at sh (nested_field_array_type t gfs lo hi) v) - (nested_field_offset t (ArraySubsc lo :: gfs)) p. -Proof. - intros. - unfold array_at. - rewrite at_offset_eq. - unfold data_at, field_at. - change (nested_field_type (nested_field_array_type t gfs lo hi) nil) - with (Tarray (nested_field_type t (gfs SUB 0)) - (hi - lo) (no_alignas_attr (attr_of_type (nested_field_type t gfs)))). - rewrite data_at_rec_eq. - rewrite <- at_offset_eq. - normalize. - apply andp_prop_ext. - + pose proof field_compatible0_nested_field_array t gfs lo hi p. - tauto. - + intros [? ?]. - rewrite at_offset_eq, <- at_offset_eq2. - rewrite at_offset_array_pred. - rewrite Z.max_r by lia. - eapply array_pred_shift; [reflexivity | lia |]. - intros. - rewrite at_offset_eq at 1. - rewrite at_offset_eq, <- at_offset_eq2, at_offset_eq. - f_equiv. - f_equiv. - rewrite @nested_field_offset_ind with (gfs := nil) by (apply (field_compatible0_nested_field_array t gfs lo hi p); auto). - assert (field_compatible0 t (gfs SUB i') p) - by (apply (field_compatible0_range _ lo hi); auto; lia). - rewrite @nested_field_offset_ind with (gfs := ArraySubsc i' :: _) by auto. - rewrite @nested_field_offset_ind with (gfs := ArraySubsc lo :: _) by auto. - rewrite @nested_field_type_ind with (gfs := ArraySubsc 0 :: _). - rewrite field_compatible0_cons in H4. - destruct (nested_field_type t gfs); try tauto. - unfold gfield_offset, gfield_type. - assert (sizeof t0 * i' = sizeof t0 * lo + sizeof t0 * i)%Z by (rewrite Zred_factor4; f_equal; lia). - hnf; lia. -Qed. - -Lemma array_at_data_at': -forall sh t gfs lo hi v p, - lo <= hi -> - field_compatible0 t (ArraySubsc lo :: gfs) p -> - field_compatible0 t (ArraySubsc hi :: gfs) p -> - array_at sh t gfs lo hi v p ⊣⊢ - data_at sh (nested_field_array_type t gfs lo hi) v - (field_address0 t (ArraySubsc lo::gfs) p). -Proof. - intros. - rewrite array_at_data_at by auto. - rewrite !prop_true_andp by auto. - unfold at_offset. - f_equiv. - unfold field_address0. - rewrite if_true; auto. -Qed. - -Lemma array_at_data_at'': -forall sh t gfs lo hi v p, - lo <= hi -> - field_compatible0 t (ArraySubsc hi :: gfs) p -> - array_at sh t gfs lo hi v p ⊣⊢ - data_at sh (nested_field_array_type t gfs lo hi) v - (field_address0 t (ArraySubsc lo::gfs) p). -Proof. - intros. - rewrite array_at_data_at by auto. - unfold at_offset. - unfold field_address0. - if_tac. - + rewrite !prop_true_andp by auto. - auto. - + apply bi.equiv_entails_2. - - normalize. - - rewrite data_at_isptr. - normalize. -Qed. - -Lemma array_at_data_at''': - forall sh t gfs lo hi v p t0 n a, - nested_field_type t gfs = Tarray t0 n a -> - lo <= hi <= n -> - array_at sh t gfs lo hi v p ⊣⊢ - data_at sh (nested_field_array_type t gfs lo hi) v - (field_address0 t (ArraySubsc lo::gfs) p). -Proof. - intros. - destruct H0. - rewrite array_at_data_at by auto. - unfold at_offset. - unfold field_address0. - if_tac. - + assert (field_compatible0 t (gfs SUB hi) p). - - rewrite field_compatible0_cons in *. - rewrite H in *. - destruct H2 as [[? ?] ?]. - split; [split |]; auto. - lia. - - rewrite !prop_true_andp by auto. - auto. - + apply bi.equiv_entails_2. - - normalize. - - rewrite data_at_isptr. - normalize. -Qed. - -Lemma split3seg_array_at': forall sh t gfs lo ml mr hi v p, - lo <= ml -> - ml <= mr -> - mr <= hi -> - Zlength v = hi-lo -> - array_at sh t gfs lo hi v p ⊣⊢ - array_at sh t gfs lo ml (sublist 0 (ml-lo) v) p ∗ - data_at sh (nested_field_array_type t gfs ml mr) - (sublist (ml-lo) (mr-lo) v) - (field_address0 t (ArraySubsc ml::gfs) p) ∗ - array_at sh t gfs mr hi (sublist (mr-lo) (hi-lo) v) p. -Proof. - intros. - rewrite (split3seg_array_at sh t gfs lo ml mr hi); auto. - rewrite (add_andp _ _ (array_at_local_facts sh t gfs mr hi _ _)). - normalize. - apply andp_prop_ext; [tauto |]. - intros [? [? _]]. - rewrite (array_at_data_at'' sh t gfs ml mr); auto. -Qed. - -(************************************************ - -Lemmas about underscore and memory_block - -************************************************) - -Lemma field_at_field_at_: forall sh t gfs v p, - field_at sh t gfs v p ⊢ field_at_ sh t gfs p. -Proof. - intros. - destruct (field_compatible_dec t gfs p). - + destruct_ptr p. - unfold field_at_, field_at. - apply bi.and_mono; first done. - pose proof field_compatible_nested_field _ _ _ f. - unfold field_compatible in H, f. - unfold offset_val in H. - autorewrite with at_offset_db in *. - unfold align_compatible, size_compatible in *. - revert H f; solve_mod_modulus; intros. - pose proof nested_field_offset_in_range t gfs. - spec H1; [tauto |]. - spec H1; [tauto |]. - change (sizeof ?A) with (expr.sizeof A) in *. - rewrite (Z.mod_small ofs) in * by lia. - rewrite (Z.mod_small (ofs + nested_field_offset t gfs)) in H - by (pose proof base.sizeof_pos (nested_field_type t gfs); lia). - apply data_at_rec_data_at_rec_; try tauto. - unfold expr.sizeof in *. - lia. - + unfold field_at_, field_at. - normalize. -Qed. - -Lemma field_at_field_at_default : forall sh t gfs v v' p, - v' = default_val (nested_field_type t gfs) -> - field_at sh t gfs v p ⊢ field_at sh t gfs v' p. -Proof. - intros; subst. - apply field_at_field_at_. -Qed. - -Lemma field_at__memory_block: forall sh t gfs p, - field_at_ sh t gfs p ⊣⊢ - memory_block sh (sizeof (nested_field_type t gfs)) (field_address t gfs p). -Proof. - intros. - unfold field_address. - destruct (field_compatible_dec t gfs p). - + unfold field_at_, field_at. - rewrite prop_true_andp by auto. - assert (isptr p) by auto; destruct p; try contradiction; clear H. rename i into ofs. - inv_int ofs. rename ofs0 into ofs. - unfold at_offset, offset_val. - solve_mod_modulus. - pose proof field_compatible_nested_field _ _ _ f. - revert H f; - unfold field_compatible; - unfold size_compatible, align_compatible, offset_val; - solve_mod_modulus; - intros. - pose proof nested_field_offset_in_range t gfs. - spec H1; [tauto |]. - spec H1; [tauto |]. - change (sizeof ?A) with (expr.sizeof A) in *. - rewrite (Z.mod_small ofs) in * by lia. - rewrite (Z.mod_small (ofs + nested_field_offset t gfs)) in H by (pose proof base.sizeof_pos (nested_field_type t gfs); lia). - rewrite memory_block_data_at_rec_default_val; first done; try tauto; unfold expr.sizeof in *; try lia. - + unfold field_at_, field_at. - rewrite memory_block_isptr. - apply bi.equiv_entails_2; normalize. -Qed. - -Lemma mapsto_zero_data_at_zero: - forall t sh p, - readable_share sh -> - complete_legal_cosu_type t = true -> - fully_nonvolatile (rank_type cenv_cs t) t = true -> - field_compatible t nil p -> - mapsto_zeros (sizeof t) sh p ⊢ data_at sh t (zero_val t) p. -Proof. -intros. -unfold data_at, field_at. -rewrite prop_true_andp by auto. -destruct H2 as [? [? [? [? ?]]]]. -unfold nested_field_offset, nested_field_rec. -unfold at_offset. -normalize. -destruct p; try contradiction. -rewrite <- (Ptrofs.repr_unsigned i). -apply mapsto_zeros_data_at_rec_zero_val; auto. -red in H4. -rep_lia. -Qed. - -Lemma data_at_data_at_ : forall sh t v p, - data_at sh t v p ⊢ data_at_ sh t p. -Proof. - intros. - apply field_at_field_at_. -Qed. - -Lemma data_at_data_at_default : forall sh t v v' p, - v' = default_val (nested_field_type t nil) -> - data_at sh t v p ⊢ data_at sh t v' p. -Proof. - intros; subst. - apply data_at_data_at_. -Qed. - -Lemma data_at__memory_block: forall sh t p, - data_at_ sh t p ⊣⊢ - ⌜field_compatible t nil p⌝ ∧ memory_block sh (sizeof t) p. -Proof. - intros. - unfold data_at_, data_at. - rewrite field_at__memory_block. - unfold field_address. - if_tac. - + normalize. - + unfold field_at_, field_at. - rewrite memory_block_isptr. - rewrite bi.pure_False by auto. - rewrite (bi.pure_False _ H). - iSplit; iIntros "([] & _)". -Qed. - -Lemma memory_block_data_at_: forall sh t p, - field_compatible t nil p -> - memory_block sh (sizeof t) p ⊣⊢ data_at_ sh t p. -Proof. - intros. - rewrite data_at__memory_block. - normalize. -Qed. - -Lemma data_at__memory_block_cancel: - forall sh t p, - data_at_ sh t p ⊢ memory_block sh (sizeof t) p. -Proof. - intros. - rewrite data_at__memory_block. - normalize. -Qed. - -Lemma data_at_memory_block: - forall sh t v p, - data_at sh t v p ⊢ memory_block sh (sizeof t) p. -Proof. - intros. - rewrite data_at_data_at_. - rewrite data_at__memory_block by auto. - iIntros "(_ & $)". -Qed. - -Lemma array_at_array_at_: forall sh t gfs lo hi v p, - array_at sh t gfs lo hi v p ⊢ array_at_ sh t gfs lo hi p. -Proof. - intros. - iIntros "H". - iDestruct (array_at_local_facts with "H") as %H. - iApply (array_at_ext_derives with "H"). - { rewrite Zlength_Zrepeat by (rewrite Zlength_correct in H; lia); lia. } - intros. - destruct (field_compatible0_dec t (ArraySubsc i :: gfs) p). - + generalize dependent u1; erewrite <- @nested_field_type_ArraySubsc with (i := i). - intros ? ->%JMeq_eq. unfold Znth. rewrite if_false by lia. - unfold Zrepeat; rewrite nth_repeat. - apply field_at_field_at_; auto. - + unfold field_at. - normalize. - contradiction n; apply field_compatible_field_compatible0; done. -Qed. - -Lemma withspacer_field_at__Tunion: forall sh t gfs i id a p, - nested_field_type t gfs = Tunion id a -> - in_members i (co_members (get_co id)) -> - withspacer sh - (nested_field_offset t gfs + - sizeof (field_type i (co_members (get_co id)))) - (nested_field_offset t gfs + sizeof (nested_field_type t gfs)) - (field_at_ sh t (gfs UDOT i)) p ⊣⊢ - memory_block sh (sizeof (nested_field_type t gfs)) (field_address t gfs p). -Proof. - intros. - rewrite withspacer_spacer. - destruct (field_compatible_dec t gfs p). - 2:{ - unfold field_at_. - assert (~ field_compatible t (gfs UDOT i) p) by (rewrite field_compatible_cons, H; tauto). - rewrite field_at_compatible'. - rewrite memory_block_isptr. - unfold field_address. - rewrite if_false by auto. - rewrite H. - apply bi.equiv_entails_2; normalize. - } - rewrite field_at__memory_block. - assert (field_compatible t (gfs UDOT i) p) by (rewrite field_compatible_cons, H; split; auto). - rewrite !field_compatible_field_address by auto. - rewrite !(nested_field_offset_ind _ (gfs UDOT _)) by auto. - unfold gfield_offset; rewrite H, Z.add_0_r. - rewrite !(nested_field_type_ind _ (gfs UDOT _)), H. - unfold gfield_type. - assert (isptr p) by auto. - destruct p; try tauto. - inv_int i0. - pose proof nested_field_offset_in_range t gfs as HH. - spec HH; [auto |]. - spec HH; [unfold field_compatible in *; tauto |]. - rewrite spacer_sepcon_memory_block. - + reflexivity. - + pose proof sizeof_pos (field_type i (co_members (get_co id))); lia. - + lia. - + change (sizeof ?A) with (expr.sizeof A) in *. - split. - - rewrite sizeof_Tunion. - erewrite co_consistent_sizeof by apply get_co_consistent. - rewrite @complete_legal_cosu_type_Tunion with (a := a) - by (rewrite <- H; apply nested_field_type_complete_legal_cosu_type; - unfold field_compatible in *; tauto). - pose proof align_le (sizeof_composite cenv_cs Union (co_members (get_co id))) - (co_alignof (get_co id)) (co_alignof_pos _). - unfold sizeof_composite in *. - pose proof sizeof_union_in_members _ _ H0. - unfold expr.sizeof in *. - lia. - - rewrite <- H. - unfold field_compatible in *. - unfold size_compatible in *. - revert H1; solve_mod_modulus; intros. - rewrite Zmod_small in H1 by lia. - lia. - + rewrite <- H. - unfold field_compatible, size_compatible in *. - rewrite Ptrofs.unsigned_repr in * by (unfold Ptrofs.max_unsigned; lia). - unfold expr.sizeof in *. - lia. -Qed. - -Lemma array_at_ramif: forall sh t gfs t0 n a lo hi i v v0 p, - nested_field_type t gfs = Tarray t0 n a -> - lo <= i < hi -> - JMeq v0 (Znth (i - lo) v) -> - array_at sh t gfs lo hi v p ⊢ field_at sh t (ArraySubsc i :: gfs) v0 p ∗ - ∀ v0 v0', ⌜JMeq v0 v0'⌝ → - (field_at sh t (ArraySubsc i :: gfs) v0 p -∗ - array_at sh t gfs lo hi (upd_Znth (i - lo) v v0') p). -Proof. - intros. - iIntros "H". - iDestruct (array_at_local_facts with "H") as %(? & ? & ? & ?). - erewrite (split3_array_at sh t gfs lo i hi) by (eauto; lia). - iDestruct "H" as "(? & $ & ?)". - clear dependent v0. - iIntros (v0 v0' ?) "?". - erewrite (split3_array_at sh t gfs lo i hi). - 2: auto. - 2:{ rewrite upd_Znth_Zlength by lia. - auto. } - 2:{ rewrite upd_Znth_same by lia. - done. } - rewrite @sublist_upd_Znth_l with (lo := 0) by lia. - rewrite @sublist_upd_Znth_r with (lo := (i + 1 - lo)) by lia. - iFrame. -Qed. - -Lemma nested_sfieldlist_at_ramif: forall sh t gfs id a i v p, - let d := default_val _ in - nested_field_type t gfs = Tstruct id a -> - in_members i (co_members (get_co id)) -> - nested_sfieldlist_at sh t gfs (co_members (get_co id)) v p ⊢ - field_at sh t (StructField (name_member (get_member i (co_members (get_co id)))) :: gfs) - (proj_struct i (co_members (get_co id)) v d) p ∗ - (∀ v0, - field_at sh t (StructField (name_member (get_member i (co_members (get_co id)))) :: gfs) v0 p -∗ - nested_sfieldlist_at sh t gfs (co_members (get_co id)) - (upd_struct i (co_members (get_co id)) v v0) p). -Proof. - intros. - pose proof (get_co_members_no_replicate id). - forget (co_members (get_co id)) as m. - destruct m; [inv H0|]. - revert v d H0; intros. - unfold nested_sfieldlist_at. - etrans. - { apply (struct_pred_ramif (m::m0) - (fun it v p => - withspacer sh - (nested_field_offset t gfs + - (field_offset cenv_cs (name_member it) (m::m0) + - sizeof (field_type (name_member it) (m::m0)))) - (nested_field_offset t gfs + - field_offset_next cenv_cs (name_member it) (m::m0) - (sizeof (nested_field_type t gfs))) - (field_at sh t (gfs DOT name_member it) v) p)); eauto. } - iIntros "(H & H1)". - iDestruct (withspacer_ramif_Q with "H") as "($ & H2)". - iIntros (?) "?"; iApply "H1"; iApply "H2"; done. -Qed. - -Lemma nested_ufieldlist_at_ramif: forall sh t gfs id a i v p, - let d := default_val _ in - nested_field_type t gfs = Tunion id a -> - in_members i (co_members (get_co id)) -> - nested_ufieldlist_at sh t gfs (co_members (get_co id)) v p ⊢ - field_at sh t (UnionField (name_member (get_member i (co_members (get_co id)))) :: gfs) - (proj_union i (co_members (get_co id)) v d) p ∗ - (∀ v0, - field_at sh t (UnionField (name_member (get_member i (co_members (get_co id)))) :: gfs) v0 p -∗ - nested_ufieldlist_at sh t gfs (co_members (get_co id)) - (upd_union i (co_members (get_co id)) v v0) p). -Proof. - intros. - pose proof (get_co_members_no_replicate id). - destruct (co_members (get_co id)) eqn:?; [inv H0|]. - revert v d H0; intros. - unfold nested_ufieldlist_at. - etrans. - { apply (union_pred_ramif (m::m0) - (fun it v p => - withspacer sh - (nested_field_offset t gfs + - sizeof - (field_type (name_member it) (m::m0))) - (nested_field_offset t gfs + - sizeof (nested_field_type t gfs)) - (field_at sh t (gfs UDOT name_member it) v) p)); try done. - instantiate (1 := default_val _). - intros. - rewrite !withspacer_spacer. - unfold fst. - fold (field_at_ sh t (gfs UDOT i) p). - rewrite field_at_field_at_. - rewrite <- !withspacer_spacer. - rewrite name_member_get. - rewrite <- Heqm. - erewrite !withspacer_field_at__Tunion; try eassumption; auto. - rewrite name_member_get. rewrite Heqm. auto. - rewrite Heqm; auto. - } - iIntros "(H & H1)". - iDestruct (withspacer_ramif_Q with "H") as "($ & H2)". - iIntros (?) "?"; iApply "H1"; iApply "H2"; done. -Qed. - -Lemma memory_block_valid_ptr: - forall sh n p, - sh ≠ Share.bot -> - n > 0 -> - memory_block sh n p ⊢ valid_pointer p. -Proof. - intros. - rewrite memory_block_isptr. - normalize. - destruct p; try tauto. - inv_int i. - replace (Vptr b (Ptrofs.repr ofs)) with (offset_val 0 (Vptr b (Ptrofs.repr ofs))) at 2. - + apply memory_block_valid_pointer with (i := 0); auto; lia. - + simpl. - rewrite ptrofs_add_repr, Z.add_0_r. - auto. -Qed. - -Lemma data_at__valid_ptr: - forall sh t p, - sh ≠ Share.bot -> - sizeof t > 0 -> - data_at_ sh t p ⊢ valid_pointer p. -Proof. - intros. - rewrite data_at__memory_block. - normalize. - apply memory_block_valid_ptr; auto. -Qed. - -Lemma data_at_valid_ptr: - forall sh t v p, - sh ≠ Share.bot -> - sizeof t > 0 -> - data_at sh t v p ⊢ valid_pointer p. -Proof. - intros. - rewrite data_at_data_at_. - apply data_at__valid_ptr; auto. -Qed. - -Lemma field_at_valid_ptr: - forall sh t path v p, - sh ≠ Share.bot -> - sizeof (nested_field_type t path) > 0 -> - field_at sh t path v p ⊢ valid_pointer (field_address t path p). -Proof. -intros. -rewrite field_at_data_at. -apply data_at_valid_ptr; auto. -Qed. - -Lemma field_at_valid_ptr0: - forall sh t path v p, - sh ≠ Share.bot -> - sizeof (nested_field_type t path) > 0 -> - nested_field_offset t path = 0 -> - field_at sh t path v p ⊢ valid_pointer p. -Proof. -intros. -assert_PROP (field_compatible t path p). -unfold field_at. -normalize. -pattern p at 2; replace p with (field_address t path p). -rewrite field_at_data_at. -apply data_at_valid_ptr; auto. -unfold field_address. rewrite if_true by auto. -rewrite H1. -normalize. -Qed. - -(************************************************ - -Other lemmas - -************************************************) - -Lemma compute_legal_nested_field_spec {prop:bi}: forall (P: prop) t gfs, - Forall (fun Q => P ⊢ ⌜Q⌝) (compute_legal_nested_field t gfs) -> - P ⊢ ⌜legal_nested_field t gfs⌝. -Proof. - intros. - induction gfs as [| gf gfs]. - + simpl. - by iIntros "?". - + simpl in H |- *. - unfold legal_field. - destruct (nested_field_type t gfs), gf; inversion H; subst; - try - match goal with - | HH : P ⊢ ⌜False⌝ |- - P ⊢ ⌜_⌝ => rewrite HH; apply bi.pure_mono; tauto - end. - - apply IHgfs in H3. - rewrite (add_andp _ _ H2). - rewrite (add_andp _ _ H3). - normalize. - - destruct_in_members i0 (co_members (get_co i)). - * apply IHgfs in H as ->. - apply bi.pure_mono; tauto. - * inversion H1. - - destruct_in_members i0 (co_members (get_co i)). - * apply IHgfs in H as ->. - apply bi.pure_mono; tauto. - * inv H. - rewrite H6; iIntros "[]". - - destruct_in_members i0 (co_members (get_co i)). - * apply IHgfs in H as ->. - apply bi.pure_mono; tauto. - * inversion H1. - - destruct_in_members i0 (co_members (get_co i)). - * apply IHgfs in H as ->. - apply bi.pure_mono; tauto. - * inv H. - rewrite H6; iIntros "[]". -Qed. - - -Lemma compute_legal_nested_field_spec': - forall t gfs, - Forall Datatypes.id (compute_legal_nested_field t gfs) -> - legal_nested_field t gfs. -Proof. - intros. - induction gfs as [| gf gfs]. - + simpl; auto. - + simpl in H|-*. - unfold legal_field. unfold nested_field_type in *. - destruct (nested_field_rec t gfs) as [[? ?] | ]. - destruct t0; try now inv H; contradiction. - destruct gf; try now inv H; contradiction. - inv H. split; auto. - destruct gf; try now inv H; contradiction. - destruct (compute_in_members i0 (co_members (get_co i))) eqn:?; - try now inv H; contradiction. - split; auto. - rewrite <- compute_in_members_true_iff; auto. - destruct gf; try now inv H; contradiction. - destruct (compute_in_members i0 (co_members (get_co i))) eqn:?; - try now inv H; contradiction. - split; auto. - rewrite <- compute_in_members_true_iff; auto. - inv H. contradiction. -Qed. - -Definition compute_legal_nested_field0 (t: type) (gfs: list gfield) : list Prop := - match gfs with - | nil => nil - | gf :: gfs0 => - match (nested_field_type t gfs0), gf with - | Tarray _ n _, ArraySubsc i => - (0 <= i <= n) :: compute_legal_nested_field t gfs0 - | Tstruct id _, StructField i => - if compute_in_members i (co_members (get_co id)) then compute_legal_nested_field t gfs else False%type :: nil - | Tunion id _, UnionField i => - if compute_in_members i (co_members (get_co id)) then compute_legal_nested_field t gfs else False%type :: nil - | _, _ => False%type :: nil - end - end. - -Lemma compute_legal_nested_field0_spec': - forall t gfs, - Forall Datatypes.id (compute_legal_nested_field0 t gfs) -> - legal_nested_field0 t gfs. -Proof. -intros. -destruct gfs; simpl in *. -auto. - unfold nested_field_type in *. - destruct (nested_field_rec t gfs) as [[? ?] | ]. - destruct t0; try now inv H; contradiction. - destruct g; try now inv H; contradiction. - inv H. split. - apply compute_legal_nested_field_spec'; auto. - apply H2. - destruct g; try now inv H; contradiction. - destruct (compute_in_members i0 (co_members (get_co i))) eqn:?; - try now inv H; contradiction. - split. - apply compute_legal_nested_field_spec'; auto. - hnf. rewrite compute_in_members_true_iff in Heqb. apply Heqb. - destruct g; try now inv H; contradiction. - destruct (compute_in_members i0 (co_members (get_co i))) eqn:?; - try now inv H; contradiction. - split. - apply compute_legal_nested_field_spec'; auto. - hnf. rewrite compute_in_members_true_iff in Heqb. apply Heqb. - inv H. contradiction. -Qed. - -Lemma splice_top_top: Share.splice Tsh Tsh = Tsh. -Proof. -unfold Share.splice. -unfold Share.Lsh, Share.Rsh. -change Share.top with Tsh. -case_eq (Share.split Tsh); intros L R ?. -simpl. -do 2 rewrite Share.rel_top1. -erewrite Share.split_together; eauto. -Qed. - -Lemma field_at_conflict: forall sh t fld p v v', - sh ≠ Share.bot -> - 0 < sizeof (nested_field_type t fld) -> - field_at sh t fld v p ∗ field_at sh t fld v' p ⊢ False. -Proof. - intros. - rewrite field_at_compatible'. - iIntros "(((% & % & % & % & %) & ?) & ?)". - destruct (nested_field_offset_in_range t fld); [done..|]. - assert (0 < sizeof (nested_field_type t fld) < Ptrofs.modulus). - { - destruct p; try done. - simpl in *. - inv_int i. - unfold expr.sizeof in *. - lia. - } - rewrite !field_at_field_at_. - rewrite field_at__memory_block by auto. - iApply (memory_block_conflict with "[$]"); first done; unfold Ptrofs.max_unsigned; lia. -Qed. - -Lemma data_at_conflict: forall sh t v v' p, - sh ≠ Share.bot -> - 0 < sizeof t -> - data_at sh t v p ∗ data_at sh t v' p ⊢ False. -Proof. - intros. unfold data_at. apply field_at_conflict; auto. -Qed. - -Lemma field_at__conflict: - forall sh t fld p, - sh ≠ Share.bot -> - 0 < sizeof (nested_field_type t fld) -> - field_at_ sh t fld p - ∗ field_at_ sh t fld p ⊢ False. -Proof. -intros. -apply field_at_conflict; auto. -Qed. - -Lemma sepcon_False_derives' {prop:bi}: - forall (P Q: prop), (Q ⊢ False) -> P ∗ Q ⊢ False. -Proof. - intros ?? ->. - iIntros "(_ & [])". -Qed. - -Lemma field_compatible_offset_isptr: -forall t path n c, field_compatible t path (offset_val n c) -> - isptr c. -Proof. -intros. -destruct H as [? _]. destruct c; try contradiction; auto. -Qed. - -Lemma field_compatible0_offset_isptr: -forall t path n c, field_compatible t path (offset_val n c) -> - isptr c. -Proof. -intros. -destruct H as [? _]. destruct c; try contradiction; auto. -Qed. - -Lemma is_pointer_or_null_field_address_lemma: - forall t path p, - is_pointer_or_null (field_address t path p) <-> - field_compatible t path p. -Proof. -intros. -unfold field_address. -if_tac; intuition (auto; try solve [contradiction]). -Qed. - -Lemma isptr_field_address_lemma: - forall t path p, - isptr (field_address t path p) <-> - field_compatible t path p. -Proof. -intros. -unfold field_address. -if_tac; intuition (auto; try solve [contradiction]). -Qed. - -Lemma eval_lvar_spec: forall id t rho, - match eval_lvar id t rho with - | Vundef => True - | Vptr b ofs => ofs = Ptrofs.zero - | _ => False - end. -Proof. - intros. - unfold eval_lvar. - destruct (Map.get (ve_of rho) id); auto. - destruct p. - destruct (eqb_type _ _); auto. -Qed. - -Lemma var_block_data_at_: - forall sh id t, - complete_legal_cosu_type t = true -> - Z.ltb (sizeof t) Ptrofs.modulus = true -> - is_aligned cenv_cs ha_env_cs la_env_cs t 0 = true -> - readable_share sh -> - var_block sh (id, t) ⊣⊢ assert_of (`(data_at_ sh t) (eval_lvar id t)). -Proof. - intros; split => rho. - unfold var_block; monPred.unseal. - unfold_lift; simpl. - apply Zlt_is_lt_bool in H0. - rewrite data_at__memory_block; try auto. - rewrite memory_block_isptr. - unfold local, lift1; unfold_lift. - pose proof eval_lvar_spec id t rho. - destruct (eval_lvar id t rho); simpl in *; normalize. - { iSplit; iIntros "((_ & []) & _)". } - subst. - apply bi.and_proper; last done. - apply bi.pure_iff. - unfold field_compatible. - unfold isptr, legal_nested_field, size_compatible, align_compatible. - change (Ptrofs.unsigned Ptrofs.zero) with 0. - rewrite Z.add_0_l. - assert (sizeof t <= Ptrofs.modulus) by lia. - assert (sizeof t <= Ptrofs.max_unsigned) by (unfold Ptrofs.max_unsigned; lia). - apply la_env_cs_sound in H1; tauto. -Qed. - -Lemma valid_pointer_weak: - forall a, valid_pointer a ⊢ weak_valid_pointer a. -Proof. -intros. -unfold valid_pointer, weak_valid_pointer. -iIntros "$". -Qed. - -Lemma valid_pointer_weak': - forall P q, (P ⊢ valid_pointer q) -> - P ⊢ weak_valid_pointer q. -Proof. -intros. -rewrite <- valid_pointer_weak; done. -Qed. - -Lemma valid_pointer_offset_zero: forall P q, - (P ⊢ valid_pointer (offset_val 0 q)) -> - P ⊢ valid_pointer q. -Proof. -intros. -destruct q; auto. -- rewrite H. - simpl valid_pointer. - iIntros "[]". -- rewrite offset_val_zero_Vptr in H. - auto. -Qed. - -End CENV. - -#[export] Hint Extern 2 (memory_block _ _ _ ⊢ valid_pointer _) => - (apply memory_block_valid_ptr; [auto with valid_pointer | rep_lia]) : valid_pointer. - -#[export] Hint Resolve valid_pointer_weak' : valid_pointer. - -#[export] Hint Extern 1 (_ ⊢ valid_pointer ?Q) => - lazymatch Q with - | offset_val _ _ => fail - | _ => apply valid_pointer_offset_zero - end : core. - -#[export] Hint Extern 2 (memory_block _ _ _ ⊢ weak_valid_pointer _) => - (apply memory_block_weak_valid_pointer; - [rep_lia | rep_lia | auto with valid_pointer]) : valid_pointer. - -Local Set SsrRewrite. (* for rewrite bi._ to work *) -Ltac field_at_conflict z fld := - apply (derives_trans _ False); [ | apply bi.False_elim]; - repeat rewrite bi.sep_assoc; - unfold data_at_, data_at, field_at_; - let x := fresh "x" in set (x := field_at _ _ fld _ z); pull_right x; - let y := fresh "y" in set (y := field_at _ _ fld _ z); pull_right y; - try (rewrite <- bi.sep_assoc; eapply sepcon_False_derives'); - subst x y; - apply field_at_conflict; auto; - try solve [simpl; (* This simpl seems safe enough, it's just simplifying (sizeof (nested_field_type _ _)) - and in any case it's followed by (computable) *) - computable]. - -Ltac data_at_conflict z := field_at_conflict z (@nil gfield). - -Ltac data_at_conflict_neq_aux1 A sh fld E x y := - match A with - | context [data_at sh _ _ y] => unify fld (@nil gfield) - | context [data_at_ sh _ y] => unify fld (@nil gfield) - | context [field_at sh _ fld _ y] => idtac - | context [field_at_ sh _ fld y] => idtac - end; - trans (⌜~ E⌝ ∧ A); - [apply bi.and_intro; [ | apply derives_refl]; - let H := fresh in - apply not_prop_right; intro H; - (rewrite H || rewrite (ptr_eq_e _ _ H)); - field_at_conflict y fld - | apply bi.pure_elim_l; - (* for this tactic to succeed, it must introduce a new hyp H1, - but rewriting H1 can fail, as the goal might be _-∗⌜C[~E]⌝ - for some context C *) - let H1 := fresh in fancy_intro H1; - rewrite ->?(bi.pure_True (~E)) by assumption - ]. - -Ltac data_at_conflict_neq_aux2 A E x y := - match A with - | context [data_at ?sh _ _ x] => data_at_conflict_neq_aux1 A sh (@nil gfield) E x y - | context [data_at_ ?sh _ x] => data_at_conflict_neq_aux1 A sh (@nil gfield) E x y - | context [field_at ?sh _ ?fld _ x] => data_at_conflict_neq_aux1 A sh fld E x y - | context [field_at_ ?sh _ ?fld x] => data_at_conflict_neq_aux1 A sh fld E x y - end. - -Ltac data_at_conflict_neq := - match goal with |- ?A ⊢ ?B => - match B with - | context [?x <> ?y] => data_at_conflict_neq_aux2 A (x=y) x y - | context [~ ptr_eq ?x ?y] => data_at_conflict_neq_aux2 A (ptr_eq x y) x y - end - end. -Local Unset SsrRewrite. - -Definition natural_aligned {cs: compspecs} (na: Z) (t: type): bool := (na mod (hardware_alignof ha_env_cs t) =? 0) && is_aligned cenv_cs ha_env_cs la_env_cs t 0. - -Definition natural_aligned_soundness {cs: compspecs}: Prop := - forall na ofs t, - complete_legal_cosu_type t = true -> - natural_aligned na t = true -> - (na | ofs) -> - align_compatible_rec cenv_cs t ofs. - -Lemma natural_aligned_sound {cs: compspecs}: - natural_aligned_soundness. -Proof. - intros. - hnf. - intros. - unfold natural_aligned in H0. - autorewrite with align in H0. - 2: eapply hardware_alignof_two_p; [exact cenv_consistent | exact ha_env_cs_consistent | exact ha_env_cs_complete]. - destruct H0. - apply la_env_cs_sound in H2; auto. - replace ofs with (ofs - 0) in H1 by lia. - eapply align_compatible_rec_hardware_alignof_divide; auto. - + exact cenv_consistent. - + exact cenv_legal_su. - + exact ha_env_cs_consistent. - + exact ha_env_cs_complete. - + eapply Z.divide_trans; eassumption. - + exact H2. -Qed. - -Definition natural_alignment := 8. - -(* TODO: change this name to malloc_compatible_ptr and merge the definition of isptr, size_compatible, align_compatible into something like: size_align_compatible_ptr *) -Definition malloc_compatible (n: Z) (p: val) : Prop := - match p with - | Vptr b ofs => (natural_alignment | Ptrofs.unsigned ofs) /\ - Ptrofs.unsigned ofs + n < Ptrofs.modulus - | _ => False - end. - -(* TODO: move these definitions and lemmas into a new file. *) -Lemma malloc_compatible_field_compatible: - forall (cs: compspecs) t p, - malloc_compatible (sizeof t) p -> - complete_legal_cosu_type t = true -> - natural_aligned natural_alignment t = true -> - field_compatible t nil p. -Proof. -intros. -destruct p; simpl in *; try contradiction. -destruct H. -eapply natural_aligned_sound in H; eauto. -pose proof (Ptrofs.unsigned_range i). -repeat split; simpl; auto; try lia. -Qed. - -#[export] Hint Extern 2 (field_compatible _ nil _) => - (apply malloc_compatible_field_compatible; - [assumption | reflexivity | reflexivity]) : core. - -Section local_facts. - -Context `{!VSTGS OK_ty Σ}. - -Lemma data_array_at_local_facts {cs: compspecs}: - forall t' n a sh (v: list (reptype t')) p, - data_at sh (Tarray t' n a) v p ⊢ - ⌜field_compatible (Tarray t' n a) nil p - /\ Zlength v = Z.max 0 n - /\ Forall (value_fits t') v⌝. -Proof. -intros. -rewrite data_at_local_facts. -apply bi.pure_mono. -intros [? ?]; split; auto. -Qed. - -Lemma data_array_at_local_facts' {cs: compspecs}: - forall t' n a sh (v: list (reptype t')) p, - n >= 0 -> - data_at sh (Tarray t' n a) v p ⊢ - ⌜field_compatible (Tarray t' n a) nil p - /\ Zlength v = n - /\ Forall (value_fits t') v⌝. -Proof. -intros. -rewrite data_array_at_local_facts. -apply bi.pure_mono. -intros [? [? ?]]; split3; auto. -rewrite Z.max_r in H1 by lia. auto. -Qed. - -End local_facts. - -Lemma value_fits_by_value {cs: compspecs}: - forall t v, - type_is_volatile t = false -> - type_is_by_value t = true -> - value_fits t v = tc_val' t (repinject t v). -Proof. -intros. -rewrite value_fits_eq; destruct t; inv H; inv H0; -simpl; rewrite H2; auto. -Qed. - -Ltac field_at_saturate_local := -unfold data_at; -match goal with |- field_at ?sh ?t ?path ?v ?c ⊢ _ => -rewrite field_at_local_facts; - let p := fresh "p" in set (p := nested_field_type t path); - simpl in p; unfold field_type in p; simpl in p; subst p; (* these simpls are probably not dangerous *) - try rewrite value_fits_by_value by reflexivity; - try match goal with |- context [repinject ?t ?v] => - change (repinject t v) with v - end; - apply derives_refl -end. - -Ltac data_at_valid_aux := - first [computable | unfold sizeof; simpl Ctypes.sizeof; rewrite ?Z.max_r by rep_lia; rep_lia | rep_lia]. - -#[export] Hint Extern 1 (data_at _ _ _ _ ⊢ valid_pointer _) => - (simple apply data_at_valid_ptr; [now auto | data_at_valid_aux]) : valid_pointer. - -#[export] Hint Extern 1 (field_at _ _ _ _ _ ⊢ valid_pointer _) => - (simple apply field_at_valid_ptr; [now auto | data_at_valid_aux]) : valid_pointer. - -#[export] Hint Extern 1 (data_at_ _ _ _ ⊢ valid_pointer _) => - (simple apply data_at__valid_ptr; [now auto | data_at_valid_aux]) : valid_pointer. - -#[export] Hint Extern 1 (field_at_ _ _ _ _ ⊢ valid_pointer _) => - (apply field_at_valid_ptr; [now auto | data_at_valid_aux]) : valid_pointer. - -#[export] Hint Extern 1 (field_at _ _ _ _ _ ⊢ _) => - (field_at_saturate_local) : saturate_local. - -#[export] Hint Extern 1 (data_at _ _ _ _ ⊢ _) => - (field_at_saturate_local) : saturate_local. - -#[export] Hint Resolve array_at_local_facts array_at__local_facts : saturate_local. - -#[export] Hint Resolve field_at__local_facts : saturate_local. -#[export] Hint Resolve data_at__local_facts : saturate_local. -#[export] Hint Extern 0 (data_at _ (Tarray _ _ _) _ _ ⊢ _) => - (apply data_array_at_local_facts'; lia) : saturate_local. -#[export] Hint Extern 0 (data_at _ (tarray _ _) _ _ ⊢ _) => - (apply data_array_at_local_facts'; lia) : saturate_local. -#[export] Hint Extern 1 (data_at _ (Tarray _ _ _) _ _ ⊢ _) => - (apply data_array_at_local_facts) : saturate_local. -#[export] Hint Extern 1 (data_at _ (tarray _ _) _ _ ⊢ _) => - (apply data_array_at_local_facts) : saturate_local. -#[export] Hint Rewrite <- @field_at_offset_zero: norm1. -#[export] Hint Rewrite <- @field_at__offset_zero: norm1. -#[export] Hint Rewrite <- @field_at_offset_zero: cancel. -#[export] Hint Rewrite <- @field_at__offset_zero: cancel. -#[export] Hint Rewrite <- @data_at__offset_zero: norm1. -#[export] Hint Rewrite <- @data_at_offset_zero: norm1. -#[export] Hint Rewrite <- @data_at__offset_zero: cancel. -#[export] Hint Rewrite <- @data_at_offset_zero: cancel. - - -(* We do these as specific lemmas, rather than - as Hint Resolve derives_refl, to limit their application - and make them fail faster *) - -Section cancel. - -Context `{!VSTGS OK_ty Σ}. - -Lemma data_at_cancel: - forall {cs: compspecs} sh t v p, - data_at sh t v p ⊢ data_at sh t v p. -Proof. intros. apply derives_refl. Qed. -Lemma field_at_cancel: - forall {cs: compspecs} sh t gfs v p, - field_at sh t gfs v p ⊢ field_at sh t gfs v p. -Proof. intros. apply derives_refl. Qed. - -Lemma data_at_field_at_cancel: - forall {cs: compspecs} sh t v p, - data_at sh t v p ⊢ field_at sh t nil v p. -Proof. intros. apply derives_refl. Qed. -Lemma field_at_data_at_cancel: - forall {cs: compspecs} sh t v p, - field_at sh t nil v p ⊢ data_at sh t v p. -Proof. intros. apply derives_refl. Qed. - -Lemma field_at__data_at__cancel: - forall {cs: compspecs} sh t p, - field_at_ sh t nil p ⊢ data_at_ sh t p. -Proof. intros. apply derives_refl. Qed. - -Lemma data_at__field_at__cancel: - forall {cs: compspecs} sh t p, - data_at_ sh t p ⊢ field_at_ sh t nil p. -Proof. intros. apply derives_refl. Qed. - -End cancel. - -#[export] Hint Resolve data_at_cancel field_at_cancel - data_at_field_at_cancel field_at_data_at_cancel : cancel. - -#[export] Hint Resolve field_at__data_at__cancel data_at__field_at__cancel : cancel. - -(* We do these as Hint Extern, instead of Hint Resolve, - to limit their application and make them fail faster *) - -#[export] Hint Extern 2 (field_at _ _ _ _ _ ⊢ field_at_ _ _ _ _) => - (simple apply field_at_field_at_) : cancel. - -#[export] Hint Extern 2 (field_at _ _ _ _ _ ⊢ field_at _ _ _ _ _) => - (simple apply field_at_field_at_default; - match goal with |- _ = default_val _ => reflexivity end) : cancel. - -#[export] Hint Extern 1 (data_at _ _ _ _ ⊢ data_at_ _ _ _) => - (simple apply data_at_data_at_) : cancel. - -#[export] Hint Extern 1 (data_at _ _ _ _ ⊢ memory_block _ _ _) => - (simple apply data_at__memory_block_cancel) : cancel. - -#[export] Hint Extern 2 (data_at _ _ _ _ ⊢ data_at _ _ _ _) => - (simple apply data_at_data_at_default; - match goal with |- _ = default_val _ => reflexivity end) : cancel. - -(* too slow this way. -#[export] Hint Extern 2 (data_at _ _ _ _ ⊢ data_at _ _ _ _) => - (apply data_at_data_at_default; reflexivity) : cancel. -*) - -#[export] Hint Extern 2 (array_at _ _ _ _ _ _ _ ⊢ array_at_ _ _ _ _ _ _) => - (simple apply array_at_array_at_) : cancel. -#[export] Hint Extern 1 (isptr _) => (eapply field_compatible_offset_isptr; eassumption) : core. -#[export] Hint Extern 1 (isptr _) => (eapply field_compatible0_offset_isptr; eassumption) : core. -#[export] Hint Rewrite @is_pointer_or_null_field_address_lemma : entailer_rewrite. -#[export] Hint Rewrite @isptr_field_address_lemma : entailer_rewrite. - -Global Transparent alignof. (* MOVE ME *) - -Ltac simplify_project_default_val := -match goal with - | |- context [@fst ?A ?B (?x, ?y)] => - change (@fst A B (x,y)) with x - | |- context [@snd ?A ?B (?x, ?y)] => - change (@snd A B (x,y)) with y - | |- context [fst (@default_val ?cs ?t)] => - let E := fresh "E" in let D := fresh "D" in let H := fresh in - set (E := fst (@default_val cs t)); - set (D := @default_val cs t) in E; - unfold compact_prod_sigT_type in E; simpl in E; - assert (H := @default_val_eq cs t); - simpl in H; - match type of H with - @eq (@reptype cs t) _ (@fold_reptype _ _ (@pair ?A ?B ?x ?y)) => - change (@reptype cs t) with (@prod A B) in *; - change (@default_val cs t) with (x,y) in * - end; - clear H; subst D; simpl in E; subst E - | |- context [snd (@default_val ?cs ?t)] => - let E := fresh "E" in let D := fresh "D" in let H := fresh in - set (E := snd (@default_val cs t)); - set (D := @default_val cs t) in E; - unfold compact_prod_sigT_type in E; simpl in E; - assert (H := @default_val_eq cs t); - simpl in H; - match type of H with - @eq (@reptype cs t) _ (@fold_reptype _ _ (@pair ?A ?B ?x ?y)) => - change (@reptype cs t) with (@prod A B) in *; - change (@default_val cs t) with (x,y) in * - end; - clear H; subst D; simpl in E; subst E -end. - -Definition field_at_mark `{!VSTGS OK_ty Σ} cs := field_at(cs := cs). -Definition field_at_hide `{!VSTGS OK_ty Σ} cs := field_at(cs := cs). -Definition data_at_hide `{!VSTGS OK_ty Σ} cs := data_at(cs := cs). - -Ltac find_field_at N := - match N with - | S O => change (field_at(cs := ?cs)) with (field_at_mark cs) at 1; - change (field_at_hide ?cs) with (field_at(cs := cs)) - | S ?k => change (field_at(cs := ?cs)) with (field_at_hide cs) at 1; - find_field_at k - end. - -Ltac find_data_at N := - match N with - | S O => match goal with |- context[data_at ?sh ?t] => - change (data_at(cs := ?cs) sh t) with (field_at_mark cs sh t nil) at 1 - end; - change (data_at_hide ?cs) with (data_at(cs := cs)) - | S ?k => change (data_at(cs := ?cs)) with (data_at_hide cs) at 1; - find_data_at k - end. - -Definition protect (T: Type) (x: T) := x. -Global Opaque protect. - -Section lemmas. - -Context `{!VSTGS OK_ty Σ}. - -Lemma field_at_ptr_neq {cs: compspecs} : - forall sh t fld p1 p2 v1 v2, - sh ≠ Share.bot -> - 0 < sizeof (nested_field_type t (fld :: nil)) -> - field_at sh t (fld::nil) v1 p1 ∗ - field_at sh t (fld::nil) v2 p2 - ⊢ - ⌜~ ptr_eq p1 p2⌝. -Proof. - intros. - apply not_prop_right; intros. - rewrite -> (ptr_eq_e _ _ H1). - apply field_at_conflict; try assumption. -Qed. - -Lemma field_at_ptr_neq_andp_emp {cs: compspecs} : - forall sh t fld p1 p2 v1 v2, - sh ≠ Share.bot -> - 0 < sizeof (nested_field_type t (fld :: nil)) -> - field_at sh t (fld::nil) v1 p1 ∗ - field_at sh t (fld::nil) v2 p2 - ⊢ - field_at sh t (fld::nil) v1 p1 ∗ - field_at sh t (fld::nil) v2 p2 ∗ - (⌜~ ptr_eq p1 p2⌝ ∧ emp). -Proof. - intros. - iIntros "H". - iDestruct (field_at_ptr_neq with "H") as %?; [done..|]. - iDestruct "H" as "($ & $)"; done. -Qed. - -Lemma field_at_ptr_neq_null {cs: compspecs} : - forall sh t fld v p, - field_at sh t fld v p ⊢ ⌜~ ptr_eq p nullval⌝. -Proof. - intros. - rewrite -> field_at_isptr. - normalize. apply bi.pure_intro. - destruct p; unfold nullval; simpl in *; tauto. -Qed. - -Lemma spacer_share_join: - forall sh1 sh2 sh J K q, - sepalg.join sh1 sh2 sh -> - spacer sh1 J K q ∗ spacer sh2 J K q ⊣⊢ spacer sh J K q. -Proof. - intros. - unfold spacer. - if_tac. { apply bi.sep_emp. } - unfold at_offset. - apply memory_block_share_join; auto. -Qed. - -Lemma struct_pred_cons2: - forall it it' m (A: member -> Type) - (P: forall it, A it -> val -> mpred) - (v: compact_prod (map A (it::it'::m))) - (p: val), - struct_pred (it :: it' :: m) P v p = - (P _ (fst v) p ∗ struct_pred (it'::m) P (snd v) p). -Proof. -intros. -destruct v; reflexivity. -Qed. - -Lemma union_pred_cons2: - forall it it' m (A: member -> Type) - (P: forall it, A it -> val -> mpred) - (v: compact_sum (map A (it::it'::m))) - (p: val), - union_pred (it :: it' :: m) P v p = - match v with inl v => P _ v p | inr v => union_pred (it'::m) P v p end. -Proof. -intros. -destruct v; reflexivity. -Qed. - -Lemma struct_pred_timeless: forall m {A} (P : forall it : member, A it -> val -> mpred) v p - (HP : forall it a p, (P it a p ⊣⊢ emp) \/ Timeless (P it a p)), - (struct_pred m P v p ⊣⊢ emp) \/ Timeless (struct_pred m P v p). -Proof. - intros. - induction m as [| a1 m]; intros; auto. - destruct m; eauto. - rewrite struct_pred_cons2. - destruct (HP a1 v.1 p) as [Hemp | Htimeless], (IHm v.2) as [Hemp' | Htimeless']. - - left; rewrite Hemp, Hemp'; apply bi.sep_emp. - - right; rewrite Hemp. - eapply bi.Timeless_proper; first apply bi.emp_sep; done. - - right; rewrite Hemp'. - eapply bi.Timeless_proper; first apply bi.sep_emp; done. - - right; apply _. -Qed. - -Lemma spacer_timeless : forall sh a b p, b - a > 0 -> Timeless (spacer sh a b p). -Proof. - intros; unfold spacer. - rewrite if_false by lia. - by apply memory_block_timeless. -Qed. - -Lemma withspacer_timeless : forall sh a b P p, a <= b -> Timeless (P p) -> Timeless (withspacer sh a b P p). -Proof. - intros; unfold withspacer. - if_tac; last apply bi.sep_timeless; try apply _. - apply spacer_timeless; lia. -Qed. - -Lemma data_at_rec_timeless {cs:compspecs} (sh : share) t (v : reptype t) p : sizeof t > 0 -> Timeless (data_at_rec sh t v p). -Proof. - revert v p. - type_induction t; intros; rewrite data_at_rec_eq; try apply _; - try (simple_if_tac; [by apply memory_block_timeless | apply _]). - - simpl in *. - unfold array_pred, aggregate_pred.array_pred. - apply bi.and_timeless; first apply _. - rewrite Z.sub_0_r, Z.max_r by lia. - assert (Ctypes.sizeof t > 0) by lia. - set (lo := 0). - assert (lo >= 0) by lia. - assert (Z.to_nat z > 0) as Hz by lia; clear H. - forget (Z.to_nat z) as n; clearbody lo. - match goal with |-context[aggregate_pred.rangespec _ _ ?Q] => set (P := Q) end. - assert (forall i v, Timeless (P i v)). - { intros; apply IH; auto. } - clearbody P; clear IH; generalize dependent lo; induction n; first lia; simpl; intros. - destruct (eq_dec n O). - + subst; simpl. eapply bi.Timeless_proper; first apply bi.sep_emp. - apply _. - + apply bi.sep_timeless; try apply _. - apply IHn; lia. - - edestruct struct_pred_timeless; last done. - + intros. - destruct (Z.gt_dec (sizeof (field_type (name_member it) (co_members (get_co id)))) 0). - * right; apply withspacer_timeless. - { -Abort. - -(*Lemma data_at_timeless {cs:compspecs} sh t v p : sizeof t > 0 -> Timeless (data_at sh t v p). -Proof. - intros. - apply bi.and_timeless; first apply _. - by apply data_at_rec_timeless. -Qed.*) - -Lemma data_at_rec_void: - forall {cs: compspecs} - sh t v q, t = Tvoid -> data_at_rec sh t v q = False. -Proof. - intros; subst; reflexivity. -Qed. - -Lemma field_at_share_join {cs: compspecs}: - forall sh1 sh2 sh t gfs v p, - sepalg.join sh1 sh2 sh -> - field_at sh1 t gfs v p ∗ field_at sh2 t gfs v p ⊣⊢ field_at sh t gfs v p. -Proof. -intros. -unfold field_at. -normalize. -apply andp_prop_ext; [tauto |]. -intros. -unfold at_offset. -destruct H0 as [? _]. -assert (isptr p) by (destruct H0; tauto). -destruct p; try inversion H1. -apply data_at_rec_share_join; auto. -Qed. - -Lemma field_at__share_join {cs: compspecs}: - forall sh1 sh2 sh t gfs p, - sepalg.join sh1 sh2 sh -> - field_at_ sh1 t gfs p ∗ field_at_ sh2 t gfs p ⊣⊢ field_at_ sh t gfs p. -Proof. intros. apply field_at_share_join. auto. Qed. - -Lemma data_at_share_join {cs: compspecs}: - forall sh1 sh2 sh t v p, - sepalg.join sh1 sh2 sh -> - data_at sh1 t v p ∗ data_at sh2 t v p ⊣⊢ data_at sh t v p. -Proof. intros. apply field_at_share_join; auto. Qed. - -Lemma data_at__share_join {cs: compspecs}: - forall sh1 sh2 sh t p, - sepalg.join sh1 sh2 sh -> - data_at_ sh1 t p ∗ data_at_ sh2 t p ⊣⊢ data_at_ sh t p. -Proof. intros. apply data_at_share_join; auto. Qed. - -Lemma data_at_conflict_glb: forall {cs: compspecs} sh1 sh2 t v v' p, - sepalg.nonidentity (Share.glb sh1 sh2) -> - 0 < sizeof t -> - data_at sh1 t v p ∗ data_at sh2 t v' p ⊢ False. -Proof. - intros. - pose (sh := Share.glb sh1 sh2). - assert (sepalg.join sh (Share.glb sh1 (Share.comp sh)) sh1). { - hnf. rewrite (Share.glb_commute sh1), <- Share.glb_assoc, Share.comp2. - rewrite Share.glb_commute, Share.glb_bot. - split; auto. - rewrite Share.distrib2, Share.comp1. - rewrite Share.glb_commute, Share.glb_top. - unfold sh. rewrite Share.lub_commute, Share.lub_absorb. auto. - } - assert (sepalg.join sh (Share.glb sh2 (Share.comp sh)) sh2). { - hnf. rewrite (Share.glb_commute sh2), <- Share.glb_assoc, Share.comp2. - rewrite Share.glb_commute, Share.glb_bot. - split; auto. - rewrite Share.distrib2, Share.comp1. - rewrite Share.glb_commute, Share.glb_top. - unfold sh. rewrite Share.glb_commute. - rewrite Share.lub_commute, Share.lub_absorb. auto. - } - rewrite <- (data_at_share_join _ _ _ _ _ _ H1). - rewrite <- (data_at_share_join _ _ _ _ _ _ H2). - iIntros "((H11 & H12) & (H21 & H22))". - iDestruct (data_at_conflict with "[$H11 $H21]") as "[]"; auto. -Qed. - -Lemma nonreadable_memory_block_field_at: - forall {cs: compspecs} - sh t gfs v p, - ~ readable_share sh -> - value_fits _ v -> - memory_block sh (sizeof (nested_field_type t gfs)) (field_address t gfs p) ⊣⊢ field_at sh t gfs v p. -Proof. - intros until p. intros NONREAD VF. - unfold field_address. - destruct (field_compatible_dec t gfs p). - + unfold field_at_, field_at. - rewrite prop_true_andp by auto. - assert (isptr p) by auto; destruct p; try contradiction; clear H. - inv_int i. - unfold at_offset, offset_val. - solve_mod_modulus. - pose proof field_compatible_nested_field _ _ _ f. - revert H f; - unfold field_compatible; - unfold size_compatible, align_compatible, offset_val; - solve_mod_modulus; - intros. - pose proof nested_field_offset_in_range t gfs. - spec H1; [tauto |]. - spec H1; [tauto |]. - rewrite (Z.mod_small ofs) in * by lia. - pose proof Zmod_le (ofs + nested_field_offset t gfs) Ptrofs.modulus. - spec H2; [pose proof Ptrofs.modulus_pos; lia |]. - revert H; solve_mod_modulus; intros. - rewrite Zmod_small in H by (pose proof sizeof_pos (nested_field_type t gfs); lia). - apply nonreadable_memory_block_data_at_rec; try tauto; try lia. - + unfold field_at_, field_at. - rewrite memory_block_isptr. - apply bi.equiv_entails_2; normalize. -Qed. - -Lemma nonreadable_memory_block_data_at: forall {cs: compspecs} sh t v p, - ~ readable_share sh -> - field_compatible t nil p -> - value_fits t v -> - memory_block sh (sizeof t) p ⊣⊢ data_at sh t v p. -Proof. - intros. - replace p with (field_address t nil p) at 1. - change t with (nested_field_type t nil) at 1. - apply nonreadable_memory_block_field_at; auto. - rewrite field_compatible_field_address by auto. - simpl. - change (nested_field_offset t nil) with 0. - apply isptr_offset_val_zero. - auto with field_compatible. -Qed. - -Lemma nonreadable_field_at_eq {cs: compspecs} : - forall sh t gfs v v' p, - ~ readable_share sh -> - (value_fits (nested_field_type t gfs) v <-> value_fits (nested_field_type t gfs) v') -> - field_at sh t gfs v p ⊣⊢ field_at sh t gfs v' p. -Proof. -intros. -rewrite !field_at_data_at. -apply bi.equiv_entails_2; saturate_local. -rewrite <- !nonreadable_memory_block_data_at; auto. -apply H0; auto. -destruct (readable_share_dec sh); try contradiction. -rewrite <- !nonreadable_memory_block_data_at; auto. -apply H0; auto. -Qed. - -Lemma nonreadable_readable_memory_block_data_at_join - {cs: compspecs}: - forall ash bsh psh t v p, - sepalg.join ash bsh psh -> - ~ readable_share ash -> - memory_block ash (sizeof t) p ∗ data_at bsh t v p ⊣⊢ data_at psh t v p. -Proof. -intros. -apply bi.equiv_entails_2; saturate_local. -rewrite @nonreadable_memory_block_data_at with (v:=v); auto. -unfold data_at. -erewrite field_at_share_join; eauto. -rewrite @nonreadable_memory_block_data_at with (v:=v); auto. -unfold data_at. -erewrite field_at_share_join; eauto. -Qed. - -Lemma nonreadable_data_at_eq {cs: compspecs}: - forall sh t v v' p, ~readable_share sh -> - (value_fits t v <-> value_fits t v') -> - data_at sh t v p ⊣⊢ data_at sh t v' p. -Proof. -intros. -unfold data_at. -apply nonreadable_field_at_eq; auto. -Qed. - -Lemma field_at_share_join_W {cs: compspecs}: - forall sh1 sh2 sh t gfs v1 v2 p, - sepalg.join sh1 sh2 sh -> - writable_share sh1 -> - field_at sh1 t gfs v1 p ∗ field_at sh2 t gfs v2 p ⊢ field_at sh t gfs v1 p. -Proof. - intros. - pose proof join_writable_readable H H0. - rewrite (add_andp _ _ (field_at_local_facts sh1 _ _ _ _)). - rewrite (add_andp _ _ (field_at_local_facts sh2 _ _ _ _)). - normalize. - rewrite (nonreadable_field_at_eq sh2 _ _ v2 v1) by (auto; tauto). - erewrite field_at_share_join by eauto. - auto. -Qed. - -Lemma data_at_share_join_W {cs: compspecs}: - forall sh1 sh2 sh t v1 v2 p, - sepalg.join sh1 sh2 sh -> - writable_share sh1 -> - data_at sh1 t v1 p ∗ data_at sh2 t v2 p ⊢ data_at sh t v1 p. -Proof. - intros. - apply field_at_share_join_W; auto. -Qed. - -Lemma value_fits_Tint_trivial {cs: compspecs} : - forall s a i, value_fits (Tint I32 s a) (Vint i). -Proof. -intros. -rewrite value_fits_eq; simpl. -unfold type_is_volatile; simpl. -destruct (attr_volatile a); auto. -hnf. intro. apply Coq.Init.Logic.I. -Qed. - -(* TODO: move all change type lemmas into one file. Also those change compspecs lemmas. *) -Lemma data_at_tuint_tint {cs: compspecs}: forall sh v p, data_at sh tuint v p ⊣⊢ data_at sh tint v p. -Proof. - intros. - unfold data_at, field_at. - apply bi.and_proper; last done. - unfold field_compatible. - apply bi.pure_iff. - assert (align_compatible tuint p <-> align_compatible tint p); [| tauto]. - destruct p; simpl; try tauto. - split; intros. - + eapply align_compatible_rec_by_value_inv in H; [| reflexivity]. - eapply align_compatible_rec_by_value; [reflexivity |]. - auto. - + eapply align_compatible_rec_by_value_inv in H; [| reflexivity]. - eapply align_compatible_rec_by_value; [reflexivity |]. - auto. -Qed. - -Lemma mapsto_field_at {cs: compspecs} sh t gfs v v' p: - type_is_by_value (nested_field_type t gfs) = true -> - type_is_volatile (nested_field_type t gfs) = false -> - field_compatible t gfs p -> - JMeq v v' -> - mapsto sh (nested_field_type t gfs) (field_address t gfs p) v ⊣⊢ field_at sh t gfs v' p. -Proof. - intros. - unfold field_at, at_offset. - rewrite by_value_data_at_rec_nonvolatile by auto. - apply (fun HH => JMeq_trans HH (JMeq_sym (repinject_JMeq _ v' H))) in H2. - apply JMeq_eq in H2. - rewrite prop_true_andp by auto. - f_equiv; auto. - apply field_compatible_field_address; auto. -Qed. - -Lemma mapsto_field_at_ramify {cs: compspecs} sh t gfs v v' w w' p: - type_is_by_value (nested_field_type t gfs) = true -> - type_is_volatile (nested_field_type t gfs) = false -> - JMeq v v' -> - JMeq w w' -> - field_at sh t gfs v' p ⊢ - mapsto sh (nested_field_type t gfs) (field_address t gfs p) v ∗ - (mapsto sh (nested_field_type t gfs) (field_address t gfs p) w -∗ - field_at sh t gfs w' p). -Proof. - intros. - unfold field_at, at_offset. - rewrite !by_value_data_at_rec_nonvolatile by auto. - apply (fun HH => JMeq_trans HH (JMeq_sym (repinject_JMeq _ v' H))) in H1; apply JMeq_eq in H1. - apply (fun HH => JMeq_trans HH (JMeq_sym (repinject_JMeq _ w' H))) in H2; apply JMeq_eq in H2. - normalize. - rewrite field_compatible_field_address by auto. - subst. - iIntros "$ $". -Qed. - -Lemma mapsto_data_at {cs: compspecs} sh t v v' p : (* not needed here *) - type_is_by_value t = true -> - type_is_volatile t = false -> - isptr p -> - size_compatible t p -> - align_compatible t p -> - complete_legal_cosu_type t = true -> - JMeq v v' -> - mapsto sh t p v ⊣⊢ data_at sh t v' p. -Proof. - intros. - unfold data_at, field_at, at_offset, offset_val. - simpl. - destruct p; inv H1. - rewrite ptrofs_add_repr_0_r. - rewrite by_value_data_at_rec_nonvolatile by auto. - apply (fun HH => JMeq_trans HH (JMeq_sym (repinject_JMeq _ v' H))) in H5; apply JMeq_eq in H5. - rewrite prop_true_andp; auto. - f_equiv; auto. - repeat split; auto. -Qed. - -Lemma mapsto_data_at' {cs: compspecs} sh t v v' p: - type_is_by_value t = true -> - type_is_volatile t = false -> - field_compatible t nil p -> - JMeq v v' -> - mapsto sh t p v ⊣⊢ data_at sh t v' p. -Proof. - intros. - unfold data_at, field_at, at_offset, offset_val. - simpl. - rewrite prop_true_andp by auto. - rewrite by_value_data_at_rec_nonvolatile by auto. - apply (fun HH => JMeq_trans HH (JMeq_sym (repinject_JMeq _ v' H))) in H2; apply JMeq_eq in H2. - f_equiv; auto. - destruct H1. destruct p; try contradiction. - rewrite ptrofs_add_repr_0_r. auto. -Qed. - -Lemma headptr_field_compatible: forall {cs: compspecs} t path p, - headptr p -> - complete_legal_cosu_type t = true -> - legal_nested_field t path -> - sizeof t < Ptrofs.modulus -> - align_compatible_rec cenv_cs t 0 -> - field_compatible t path p. -Proof. - intros. - destruct H as [b ?]; subst p. - repeat split; auto. -Qed. - -Lemma mapsto_data_at'' {cs: compspecs}: forall sh t v v' p, - ((type_is_by_value t) && (complete_legal_cosu_type t) && (negb (type_is_volatile t)) && is_aligned cenv_cs ha_env_cs la_env_cs t 0 = true)%bool -> - headptr p -> - JMeq v v' -> - mapsto sh t p v ⊣⊢ data_at sh t v' p. -Proof. - intros. - rewrite !andb_true_iff in H. - destruct H as [[[? ?] ?] ?]. - rewrite negb_true_iff in H3. - apply mapsto_data_at'; auto. - apply headptr_field_compatible; auto. - + destruct t; inv H; simpl; auto. - + destruct t as [| [ | | | ] ? | | [ | ] | | | | |]; inv H; reflexivity. - + apply la_env_cs_sound in H4; auto. -Qed. - -Lemma data_at_type_changable {cs}: forall (sh: Share.t) (t1 t2: type) v1 v2, - t1 = t2 -> - JMeq v1 v2 -> - data_at (cs := cs) sh t1 v1 = data_at sh t2 v2. -Proof. intros. subst. apply JMeq_eq in H0. subst v2. reflexivity. Qed. - -Lemma field_at_change_composite {cs_from cs_to} {CCE: change_composite_env cs_from cs_to}: forall (sh: Share.t) (t: type) gfs v1 v2 p, - JMeq v1 v2 -> - cs_preserve_type cs_from cs_to (coeq _ _) t = true -> - field_at (cs := cs_from) sh t gfs v1 p ⊣⊢ field_at (cs := cs_to) sh t gfs v2 p. -Proof. - intros. - unfold field_at. - apply andp_prop_ext. - + apply field_compatible_change_composite; auto. - + intros. - pose proof H1. - rewrite field_compatible_change_composite in H2 by auto. - rewrite nested_field_offset_change_composite by auto. - revert v1 H; rewrite nested_field_type_change_composite by auto. - intros. - apply data_at_rec_change_composite; auto. - apply nested_field_type_preserves_change_composite; auto. -Qed. - -Lemma field_at__change_composite {cs_from cs_to} {CCE: change_composite_env cs_from cs_to}: forall (sh: Share.t) (t: type) gfs p, - cs_preserve_type cs_from cs_to (coeq _ _) t = true -> - field_at_ (cs := cs_from) sh t gfs p ⊣⊢ field_at_ (cs := cs_to) sh t gfs p. -Proof. - intros. - apply field_at_change_composite; auto. - rewrite nested_field_type_change_composite by auto. - apply default_val_change_composite. - apply nested_field_type_preserves_change_composite; auto. -Qed. - -Lemma data_at_change_composite {cs_from cs_to} {CCE: change_composite_env cs_from cs_to}: forall (sh: Share.t) (t: type) v1 v2 p, - JMeq v1 v2 -> - cs_preserve_type cs_from cs_to (coeq _ _) t = true -> - data_at (cs := cs_from) sh t v1 p ⊣⊢ data_at (cs := cs_to) sh t v2 p. -Proof. - intros. - apply field_at_change_composite; auto. -Qed. - -Lemma data_at__change_composite {cs_from cs_to} {CCE: change_composite_env cs_from cs_to}: forall (sh: Share.t) (t: type) p, - cs_preserve_type cs_from cs_to (coeq _ _) t = true -> - data_at_ (cs := cs_from) sh t p ⊣⊢ data_at_ (cs := cs_to) sh t p. -Proof. - intros. - apply field_at__change_composite; auto. -Qed. - -(* TODO: rename and clean up all array_at_data_at lemmas. *) -Lemma array_at_data_at1 {cs : compspecs} : forall sh t gfs lo hi v p, - lo <= hi -> - field_compatible0 t (gfs SUB lo) p -> - field_compatible0 t (gfs SUB hi) p -> - array_at sh t gfs lo hi v p ⊣⊢ - at_offset (data_at sh (nested_field_array_type t gfs lo hi) v) - (nested_field_offset t (ArraySubsc lo :: gfs)) p. -Proof. - intros. rewrite array_at_data_at by auto. unfold at_offset. apply bi.equiv_entails_2; normalize. -Qed. - -Lemma data_at_ext_derives {cs : compspecs} sh t v v' p q: v=v' -> p=q -> data_at sh t v p ⊢ data_at sh t v' q. -Proof. intros; subst. -apply derives_refl. -Qed. - -Lemma data_at_ext_eq {cs : compspecs} sh t v v' p q: v=v' -> p=q -> data_at sh t v p = data_at sh t v' q. -Proof. intros; subst. trivial. Qed. - -End lemmas. - -(* does not simplify array indices, because that might be too expensive *) -Ltac simpl_compute_legal_nested_field := - repeat match goal with - | |- context [ compute_legal_nested_field ?T ?L ] => - let r := eval hnf in (compute_legal_nested_field T L) in - change (compute_legal_nested_field T L) with r - end. - -Ltac solve_legal_nested_field_in_entailment := - match goal with - | |- _ ⊢ ⌜legal_nested_field ?t_root ?gfs⌝ => - try unfold t_root; - try unfold gfs; - try match gfs with - | (?gfs1 ++ ?gfs0) => try unfold gfs1; try unfold gfs0 - end - end; - first - [ apply bi.pure_intro; apply compute_legal_nested_field_spec'; - simpl_compute_legal_nested_field; - repeat apply Forall_cons; try apply Forall_nil; lia - | - apply compute_legal_nested_field_spec; - simpl_compute_legal_nested_field; - repeat apply Forall_cons; try apply Forall_nil; - try solve [apply bi.pure_intro; auto; lia]; - try solve [normalize; apply bi.pure_intro; auto; lia] - ]. - -Ltac headptr_field_compatible := - match goal with H: headptr ?P |- field_compatible _ _ ?P => - apply headptr_field_compatible; - [ apply H | reflexivity | | simpl; computable | apply la_env_cs_sound; reflexivity]; - apply compute_legal_nested_field_spec'; - simpl_compute_legal_nested_field; - repeat apply Forall_cons; try apply Forall_nil - end. - -#[export] Hint Extern 2 (field_compatible _ _ _) => headptr_field_compatible : field_compatible. - -(* BEGIN New experiments *) -Section new_lemmas. - -Context `{!VSTGS OK_ty Σ}. - -Lemma data_at_data_at_cancel {cs: compspecs}: forall sh t v v' p, - v = v' -> - data_at sh t v p ⊢ data_at sh t v' p. -Proof. intros. subst. apply derives_refl. Qed. - -Lemma field_at_field_at_cancel {cs: compspecs}: forall sh t gfs v v' p, - v = v' -> - field_at sh t gfs v p ⊢ field_at sh t gfs v' p. -Proof. intros. subst. apply derives_refl. Qed. - -Lemma data_at__data_at {cs: compspecs}: - forall sh t v p, v = default_val t -> data_at_ sh t p ⊢ data_at sh t v p. -Proof. -intros; subst; unfold data_at_; apply derives_refl. -Qed. - -Lemma data_at__eq : forall {cs : compspecs} sh t p, data_at_ sh t p = data_at sh t (default_val t) p. -Proof. - intros; unfold data_at_, data_at, field_at_; auto. -Qed. - -Lemma data_at_shares_join : forall {cs : compspecs} sh t v p shs sh1 (Hsplit : sepalg_list.list_join sh1 shs sh), - data_at sh1 t v p ∗ ([∗ list] sh ∈ shs, data_at sh t v p) ⊣⊢ - data_at sh t v p. -Proof. - induction shs; intros; simpl. - - inv Hsplit. - apply bi.sep_emp. - - inv Hsplit. - rewrite assoc, data_at_share_join; eauto; apply _. -Qed. - -Lemma data_at_shares_join_old : forall {cs : compspecs} sh t v p shs sh1 (Hsplit : sepalg_list.list_join sh1 shs sh), - data_at sh1 t v p ∗ fold_right bi_sep emp (map (fun sh => data_at sh t v p) shs) ⊣⊢ - data_at sh t v p. -Proof. - induction shs; intros; simpl. - - inv Hsplit. - apply bi.sep_emp. - - inv Hsplit. - rewrite assoc, data_at_share_join; eauto; apply _. -Qed. - -Lemma struct_pred_value_cohere : forall {cs : compspecs} m sh1 sh2 p t f off v1 v2 - (Hsh1 : readable_share sh1) (Hsh2 : readable_share sh2) - (IH : Forall (fun it : member => forall v1 v2 (p : val), - readable_share sh1 -> readable_share sh2 -> - data_at_rec sh1 (t it) v1 p ∗ data_at_rec sh2 (t it) v2 p ⊢ - data_at_rec sh1 (t it) v1 p ∗ data_at_rec sh2 (t it) v1 p) m), - struct_pred m (fun (it : member) v => - withspacer sh1 (f it + sizeof (t it)) (off it) (at_offset (data_at_rec sh1 (t it) v) (f it))) v1 p ∗ - struct_pred m (fun (it : member) v => - withspacer sh2 (f it + sizeof (t it)) (off it) (at_offset (data_at_rec sh2 (t it) v) (f it))) v2 p ⊢ - struct_pred m (fun (it : member) v => - withspacer sh1 (f it + sizeof (t it)) (off it) (at_offset (data_at_rec sh1 (t it) v) (f it))) v1 p ∗ - struct_pred m (fun (it : member) v => - withspacer sh2 (f it + sizeof (t it)) (off it) (at_offset (data_at_rec sh2 (t it) v) (f it))) v1 p. -Proof. - intros. - revert v1 v2; induction m; auto; intros. - inv IH. - destruct m. - - unfold withspacer, at_offset; simpl. - if_tac; auto. - match goal with |- (?P1 ∗ ?Q1) ∗ (?P2 ∗ ?Q2) ⊢ _ => trans ((P1 ∗ P2) ∗ (Q1 ∗ Q2)); - [cancel|] end. - rewrite H1; auto. - cancel. - - rewrite !struct_pred_cons2. - match goal with |- (?P1 ∗ ?Q1) ∗ (?P2 ∗ ?Q2) ⊢ _ => trans ((P1 ∗ P2) ∗ (Q1 ∗ Q2)); - [cancel|] end. - match goal with |- _ ⊢ (?P1 ∗ ?Q1) ∗ (?P2 ∗ ?Q2) => trans ((P1 ∗ P2) ∗ (Q1 ∗ Q2)); - [|cancel] end. - apply bi.sep_mono; auto. - unfold withspacer, at_offset; simpl. - if_tac; auto. - match goal with |- (?P1 ∗ ?Q1) ∗ (?P2 ∗ ?Q2) ⊢ _ => trans ((P1 ∗ P2) ∗ (Q1 ∗ Q2)); - [cancel|] end. - rewrite H1; auto. - cancel. -Qed. - -Lemma mapsto_value_eq: forall sh1 sh2 t p v1 v2, readable_share sh1 -> readable_share sh2 -> - v1 <> Vundef -> v2 <> Vundef -> mapsto sh1 t p v1 ∗ mapsto sh2 t p v2 ⊢ ⌜v1 = v2⌝. -Proof. - intros; unfold mapsto. - destruct (access_mode t); try solve [iIntros "([] & _)"]. - destruct (type_is_volatile t); try solve [iIntros "([] & _)"]. - destruct p; try solve [iIntros "([] & _)"]. - rewrite !if_true by done. - iIntros "([(_ & H1) | (-> & % & H1)] & [(_ & H2) | (-> & % & H2)])"; try solve [exfalso; pose proof (JMeq_refl Vundef); done]; - iApply res_predicates.address_mapsto_value_cohere; iFrame. -Qed. - -Lemma mapsto_value_cohere: forall sh1 sh2 t p v1 v2, readable_share sh1 -> - mapsto sh1 t p v1 ∗ mapsto sh2 t p v2 ⊢ mapsto sh1 t p v1 ∗ mapsto sh2 t p v1. -Proof. - intros; unfold mapsto. - destruct (access_mode t); try simple apply derives_refl. - destruct (type_is_volatile t); try simple apply derives_refl. - destruct p; try simple apply derives_refl. - rewrite if_true by done. - destruct (eq_dec v1 Vundef). - - subst; rewrite !prop_false_andp with (P := tc_val t Vundef), !bi.False_or, prop_true_andp; auto; - try apply tc_val_Vundef. - cancel. - if_tac. - + iIntros "[(% & ?) | (% & ?)]"; iRight; auto. - + Intros. iIntros "$"; iPureIntro; repeat split; auto. apply tc_val'_Vundef. - - rewrite !prop_false_andp with (P := v1 = Vundef), !bi.or_False; auto; Intros. - apply bi.and_intro; [apply bi.pure_intro; auto|]. - if_tac. - + iIntros "(H1 & H2)". - iAssert (∃ v2' : val, res_predicates.address_mapsto m v2' _ _) with "[H2]" as (v2') "H2". - { iDestruct "H2" as "[(% & ?) | (_ & $)]"; auto. } - iAssert ⌜v1 = v2'⌝ as %->. { iApply res_predicates.address_mapsto_value_cohere; iFrame. } - iFrame; eauto. - + apply bi.sep_mono; first done. - iIntros "((% & %) & $)"; iPureIntro; repeat split; auto. - apply tc_val_tc_val'; auto. -Qed. - -Lemma data_at_value_cohere : forall {cs : compspecs} sh1 sh2 t v1 v2 p, readable_share sh1 -> - type_is_by_value t = true -> type_is_volatile t = false -> - data_at sh1 t v1 p ∗ data_at sh2 t v2 p ⊢ - data_at sh1 t v1 p ∗ data_at sh2 t v1 p. -Proof. - intros; unfold data_at, field_at, at_offset. - iIntros "((% & ?) & (% & ?))". - rewrite !by_value_data_at_rec_nonvolatile by auto. - iDestruct (mapsto_value_cohere with "[-]") as "($ & $)"; auto; iFrame. -Qed. - -Lemma data_at_value_eq : forall {cs : compspecs} sh1 sh2 t v1 v2 p, - readable_share sh1 -> readable_share sh2 -> - is_pointer_or_null v1 -> is_pointer_or_null v2 -> - data_at sh1 (tptr t) v1 p ∗ data_at sh2 (tptr t) v2 p ⊢ ⌜v1 = v2⌝. -Proof. - intros; unfold data_at, field_at, at_offset; Intros. - rewrite !by_value_data_at_rec_nonvolatile by auto. - apply mapsto_value_eq; auto. - { intros X; subst; contradiction. } - { intros X; subst; contradiction. } -Qed. - -Lemma data_at_array_value_cohere : forall {cs : compspecs} sh1 sh2 t z a v1 v2 p, readable_share sh1 -> - type_is_by_value t = true -> type_is_volatile t = false -> - data_at sh1 (Tarray t z a) v1 p ∗ data_at sh2 (Tarray t z a) v2 p ⊢ - data_at sh1 (Tarray t z a) v1 p ∗ data_at sh2 (Tarray t z a) v1 p. -Proof. - intros; unfold data_at, field_at, at_offset. - iIntros "((% & H1) & (_ & H2))". - rewrite !bi.pure_True, !bi.True_and by done. - rewrite !data_at_rec_eq; simpl. - unfold array_pred, aggregate_pred.array_pred. - iDestruct "H1" as (?) "H1"; iDestruct "H2" as (?) "H2". - rewrite !bi.pure_True, !bi.True_and by done. - rewrite Z.sub_0_r in *. - rewrite Z2Nat_max0 in *. - clear H3 H4. - forget (offset_val 0 p) as p'; forget (Z.to_nat z) as n. - set (lo := 0) at 1 3 5 7; clearbody lo. - iInduction n as [|] "IH" forall (lo); auto; simpl; intros. - iDestruct "H1" as "(H1a & H1b)"; iDestruct "H2" as "(H2a & H2b)". - unfold at_offset. - rewrite !by_value_data_at_rec_nonvolatile by auto. - iDestruct (mapsto_value_cohere with "[$H1a $H2a]") as "($ & $)"; first done. - iApply ("IH" with "H1b H2b"). -Qed. - -Lemma field_at_array_inbounds : forall {cs : compspecs} sh t z a i v p, - field_at sh (Tarray t z a) (ArraySubsc i :: nil) v p ⊢ ⌜0 <= i < z⌝. -Proof. - intros. rewrite field_at_compatible'. - apply bi.pure_elim_l. intros. - apply bi.pure_intro. - destruct H as (_ & _ & _ & _ & _ & ?); auto. -Qed. - -Lemma field_at__field_at {cs: compspecs} : - forall sh t gfs v p, v = default_val (nested_field_type t gfs) -> field_at_ sh t gfs p ⊢ field_at sh t gfs v p. -Proof. -intros; subst; unfold field_at_; apply derives_refl. -Qed. - -Lemma data_at__field_at {cs: compspecs}: - forall sh t v p, v = default_val t -> data_at_ sh t p ⊢ field_at sh t nil v p. -Proof. -intros; subst; unfold data_at_; apply derives_refl. -Qed. - -Lemma field_at__data_at {cs: compspecs} : - forall sh t v p, v = default_val (nested_field_type t nil) -> field_at_ sh t nil p ⊢ data_at sh t v p. -Proof. -intros; subst; unfold field_at_; apply derives_refl. -Qed. - -Lemma field_at_data_at_cancel': forall {cs : compspecs} sh t v p, - field_at sh t nil v p ⊣⊢ data_at sh t v p. -Proof. - intros. apply bi.equiv_entails_2. - apply field_at_data_at_cancel. - apply data_at_field_at_cancel. -Qed. - -End new_lemmas. - -#[export] Hint Resolve data_at_data_at_cancel : cancel. -#[export] Hint Resolve data_at_data_at_cancel : cancel. -#[export] Hint Resolve field_at_field_at_cancel : cancel. -#[export] Hint Resolve data_at__data_at : cancel. -#[export] Hint Resolve field_at__field_at : cancel. -#[export] Hint Resolve data_at__field_at : cancel. -#[export] Hint Resolve field_at__data_at : cancel. - -#[export] Hint Extern 1 (_ = @default_val _ _) => - match goal with |- ?A = ?B => - let x := fresh "x" in set (x := B); hnf in x; subst x; - match goal with |- ?A = ?B => constr_eq A B; reflexivity - end end : core. - -#[export] Hint Extern 1 (_ = _) => - match goal with |- ?A = ?B => constr_eq A B; reflexivity end : cancel. - -(* enhance cancel to solve field_at and data_at *) - -#[export] Hint Rewrite - @field_at_data_at_cancel' - @field_at_data_at - @field_at__data_at_ : cancel. - -(* END new experiments *) - -Section more_lemmas. - -Context `{!VSTGS OK_ty Σ}. - -Lemma data_at__Tarray: - forall {CS: compspecs} sh t n a, - data_at_ sh (Tarray t n a) = - data_at sh (Tarray t n a) (Zrepeat (default_val t) n). -Proof. reflexivity. Qed. - -Lemma data_at__tarray: - forall {CS: compspecs} sh t n, - data_at_ sh (tarray t n) = - data_at sh (tarray t n) (Zrepeat (default_val t) n). -Proof. intros; apply data_at__Tarray; auto. Qed. - -Lemma data_at__Tarray': - forall {CS: compspecs} sh t n a v, - v = Zrepeat (default_val t) n -> - data_at_ sh (Tarray t n a) = data_at sh (Tarray t n a) v. -Proof. -intros. subst; reflexivity. -Qed. - -Lemma data_at__tarray': - forall {CS: compspecs} sh t n v, - v = Zrepeat (default_val t) n -> - data_at_ sh (tarray t n) = data_at sh (tarray t n) v. -Proof. intros; apply data_at__Tarray'; auto. Qed. - -Lemma change_compspecs_field_at_cancel: - forall {cs1 cs2: compspecs} {CCE : change_composite_env cs1 cs2} - (sh: share) (t1 t2: type) gfs - (v1: @reptype cs1 (@nested_field_type cs1 t1 gfs)) - (v2: @reptype cs2 (@nested_field_type cs2 t2 gfs)) - (p: val), - t1 = t2 -> - cs_preserve_type cs1 cs2 (@coeq cs1 cs2 CCE) t1 = true -> - JMeq v1 v2 -> - field_at (cs := cs1) sh t1 gfs v1 p ⊢ field_at (cs := cs2) sh t2 gfs v2 p. -Proof. -intros. -subst t2. -rewrite @field_at_change_composite with CCE; auto. -Qed. - -Lemma change_compspecs_data_at_cancel: - forall {cs1 cs2: compspecs} {CCE : change_composite_env cs1 cs2} - (sh: share) (t1 t2: type) - (v1: @reptype cs1 t1) (v2: @reptype cs2 t2) - (p: val), - t1 = t2 -> - cs_preserve_type cs1 cs2 (@coeq cs1 cs2 CCE) t1 = true -> - JMeq v1 v2 -> - data_at (cs := cs1) sh t1 v1 p ⊢ data_at (cs := cs2) sh t2 v2 p. -Proof. -intros. -apply change_compspecs_field_at_cancel; auto. -Qed. - -Lemma change_compspecs_field_at_cancel2: - forall {cs1 cs2: compspecs} {CCE : change_composite_env cs1 cs2} - (sh: share) (t1 t2: type) gfs - (p: val), - t1 = t2 -> - cs_preserve_type cs1 cs2 (@coeq cs1 cs2 CCE) t1 = true -> - field_at_ (cs := cs1) sh t1 gfs p ⊢ field_at_ (cs := cs2) sh t2 gfs p. -Proof. -intros. -subst t2. -apply @change_compspecs_field_at_cancel with CCE; auto. -pose proof (@nested_field_type_change_composite cs1 cs2 CCE t1 H0 gfs). -rewrite H. -apply @default_val_change_composite with CCE; auto. -apply nested_field_type_preserves_change_composite; auto. -Qed. - -Lemma change_compspecs_data_at_cancel2: - forall {cs1 cs2: compspecs} {CCE : change_composite_env cs1 cs2} - (sh: share) (t1 t2: type) - (p: val), - t1 = t2 -> - cs_preserve_type cs1 cs2 (@coeq cs1 cs2 CCE) t1 = true -> - data_at_ (cs := cs1) sh t1 p ⊢ data_at_ (cs := cs2) sh t2 p. -Proof. -intros. -apply change_compspecs_field_at_cancel2; auto. -Qed. - -Lemma change_compspecs_field_at_cancel3: - forall {cs1 cs2: compspecs} {CCE : change_composite_env cs1 cs2} - (sh: share) (t1 t2: type) gfs - (v1: @reptype cs1 (@nested_field_type cs1 t1 gfs)) - (p: val), - t1 = t2 -> - cs_preserve_type cs1 cs2 (@coeq cs1 cs2 CCE) t1 = true -> - field_at (cs := cs1) sh t1 gfs v1 p ⊢ field_at_ (cs := cs2) sh t2 gfs p. -Proof. -intros. -subst t2. -rewrite field_at_field_at_. -apply @change_compspecs_field_at_cancel2 with CCE; auto. -Qed. - -Lemma change_compspecs_data_at_cancel3: - forall {cs1 cs2: compspecs} {CCE : change_composite_env cs1 cs2} - (sh: share) (t1 t2: type) - (v1: @reptype cs1 t1) - (p: val), - t1 = t2 -> - cs_preserve_type cs1 cs2 (@coeq cs1 cs2 CCE) t1 = true -> - data_at (cs := cs1) sh t1 v1 p ⊢ data_at_ (cs := cs2) sh t2 p. -Proof. -intros. -apply @change_compspecs_field_at_cancel3 with CCE; auto. -Qed. - -Lemma data_at_nullptr: - forall {cs: compspecs} sh t p, - data_at sh size_t nullval p ⊣⊢ - data_at sh (tptr t) nullval p. -Proof. -intros. -unfold data_at, field_at. -apply bi.and_proper. -f_equiv. -unfold field_compatible; simpl. -intuition; destruct p; try auto; -(eapply align_compatible_rec_by_value_inv in H2; [ | reflexivity]; - eapply align_compatible_rec_by_value; [reflexivity | ]; - apply H2). -unfold at_offset. -rewrite !by_value_data_at_rec_nonvolatile by reflexivity. -simpl. -unfold nested_field_type; simpl. -rewrite <- mapsto_tuint_tptr_nullval with (t:=t). -done. -Qed. - -Lemma data_at_int_or_ptr_int: - forall {CS: compspecs} sh i p, - data_at sh int_or_ptr_type (Vptrofs i) p - = data_at sh size_t (Vptrofs i) p. -Proof. - intros. - unfold data_at, field_at. - simpl. f_equal. - f_equal. - unfold field_compatible. - f_equal. - f_equal. - f_equal. - f_equal. - unfold align_compatible. - destruct p; auto. - apply prop_ext; split; intro; - eapply align_compatible_rec_by_value_inv in H; - try reflexivity; - try (eapply align_compatible_rec_by_value; eauto). -Qed. - -Lemma data_at_int_or_ptr_ptr: - forall {CS: compspecs} sh t v p, - isptr v -> - data_at sh int_or_ptr_type v p - = data_at sh (tptr t) v p. -Proof. - intros. - destruct v; try contradiction. - clear H. - unfold data_at, field_at. - simpl. f_equal. - f_equal. - unfold field_compatible. - f_equal. - f_equal. - f_equal. - f_equal. - unfold align_compatible. - destruct p; auto. - apply prop_ext; split; intro; - eapply align_compatible_rec_by_value_inv in H; - try reflexivity; - try (eapply align_compatible_rec_by_value; eauto). - unfold at_offset. - unfold nested_field_type; simpl. - unfold data_at_rec; simpl. - unfold mapsto. - simpl. - destruct p; simpl; auto. - if_tac; auto. - f_equal. - simple_if_tac; auto. - f_equal. rewrite andb_false_r. reflexivity. - f_equal. rewrite andb_false_r. reflexivity. - f_equal. - f_equal. - f_equal. - unfold tc_val'. - unfold tc_val; simpl. - rewrite N.eqb_refl. - rewrite andb_false_r. reflexivity. -Qed. - -Lemma nonempty_writable0_glb (shw shr : share) : writable0_share shw -> readable_share shr -> - nonempty_share (Share.glb shw shr). - (* this lemma might be convenient for users *) -Proof. -intros Hshw Hshr. -apply leq_join_sub in Hshw. -apply Share.ord_spec2 in Hshw. -rewrite Share.glb_commute, <- Hshw, Share.distrib1, Share.glb_commute, Share.lub_commute. -apply readable_nonidentity. -apply readable_share_lub. -apply readable_glb. -assumption. -Qed. - -Lemma nonempty_writable_glb (shw shr : share) : writable_share shw -> readable_share shr -> - nonempty_share (Share.glb shw shr). - (* this lemma might be convenient for users *) -Proof. -intros Hshw Hshr. -apply nonempty_writable0_glb; try assumption. -apply writable_writable0; assumption. -Qed. - -End more_lemmas. - -Ltac unfold_data_at_ p := - match goal with |- context [data_at_ ?sh ?t p] => - let d := fresh "d" in set (d := data_at_ sh t p); - pattern d; - let g := fresh "goal" in - match goal with |- ?G d => set (g:=G) end; - revert d; - match t with - | Tarray ?t1 ?n _ => - erewrite data_at__Tarray' by apply eq_refl; - try change (default_val t1) with Vundef - | tarray ?t1 ?n => - erewrite data_at__tarray' by apply eq_refl; - try change (default_val t1) with Vundef - | _ => change (data_at_ sh t p) with (data_at sh t (default_val t) p); - try change (default_val t) with Vundef - end; - subst g; intro d; subst d; cbv beta - end. - -#[export] Hint Extern 2 (data_at_(cs := ?cs1) ?sh _ ?p ⊢ data_at_(cs := ?cs2) ?sh _ ?p) => - (tryif constr_eq cs1 cs2 then fail - else simple apply change_compspecs_data_at_cancel2; reflexivity) : cancel. - -#[export] Hint Extern 2 (data_at(cs := ?cs1) ?sh _ _ ?p ⊢ data_at_(cs := ?cs2) ?sh _ ?p) => - (tryif constr_eq cs1 cs2 then fail - else simple apply change_compspecs_data_at_cancel3; reflexivity) : cancel. - -#[export] Hint Extern 2 (data_at(cs := ?cs1) ?sh _ _ ?p ⊢ data_at(cs := ?cs2) ?sh _ _ ?p) => - (tryif constr_eq cs1 cs2 then fail - else simple apply change_compspecs_data_at_cancel; - [ reflexivity | reflexivity | apply JMeq_refl]) : cancel. - -#[export] Hint Extern 2 (field_at_(cs := ?cs1) ?sh _ ?gfs ?p ⊢ field_at_(cs := ?cs2) ?sh _ ?gfs ?p) => - (tryif constr_eq cs1 cs2 then fail - else simple apply change_compspecs_field_at_cancel2; reflexivity) : cancel. - -#[export] Hint Extern 2 (field_at(cs := ?cs1) ?sh _ ?gfs _ ?p ⊢ field_at_(cs := ?cs2) ?sh _ ?gfs ?p) => - (tryif constr_eq cs1 cs2 then fail - else simple apply change_compspecs_field_at_cancel3; reflexivity) : cancel. - -#[export] Hint Extern 2 (field_at(cs := ?cs1) ?sh _ ?gfs _ ?p ⊢ field_at(cs := ?cs2) ?sh _ ?gfs _ ?p) => - (tryif constr_eq cs1 cs2 then fail - else simple apply change_compspecs_field_at_cancel; - [ reflexivity | reflexivity | apply JMeq_refl]) : cancel. diff --git a/progs/list_dt.v.crashcoqide b/progs/list_dt.v.crashcoqide deleted file mode 100644 index 9e1c521b6..000000000 --- a/progs/list_dt.v.crashcoqide +++ /dev/null @@ -1,2634 +0,0 @@ -(* Require Import VST.floyd.proofauto. - TEMPORARILY replace "floyd.proofauto" - with all the imports in the list below. - This reduces makefile-based recompilation - when changing things in (e.g.) forward.v -*) -Require Import VST.floyd.base2. -Require Import VST.floyd.client_lemmas. -Require Import VST.floyd.closed_lemmas. -Require Import VST.floyd.nested_pred_lemmas. -Require Import VST.floyd.nested_field_lemmas. -Require Import VST.floyd.efield_lemmas. -Require Import VST.floyd.mapsto_memory_block. -Require Import VST.floyd.reptype_lemmas. -Require VST.floyd.aggregate_pred. Import VST.floyd.aggregate_pred.aggregate_pred. -Require Import VST.floyd.data_at_rec_lemmas. -Require Import VST.floyd.field_at. -Require Import VST.floyd.nested_loadstore. -(*Require Import VST.floyd.unfold_data_at.*) -Require Import VST.floyd.entailer. -Require Import VST.floyd.compat. -(* End TEMPORARILY *) - -Lemma int64_eq_e: forall i j, Int64.eq i j = true -> i=j. -Proof. intros. pose proof (Int64.eq_spec i j); rewrite H in H0; auto. Qed. - -Lemma ptrofs_eq_e: forall i j, Ptrofs.eq i j = true -> i=j. -Proof. intros. pose proof (Ptrofs.eq_spec i j); rewrite H in H0; auto. Qed. - -(*Lemma allp_andp1 {A}{ND: NatDed A}: forall B (any: B) (p: B -> A) q, andp (allp p) q = (allp (fun x => andp (p x) q)). -Proof. - intros. apply pred_ext. - apply allp_right; intro x. - apply andp_derives; auto. apply allp_left with x; auto. - apply andp_right. apply allp_right; intro x. apply allp_left with x. apply andp_left1; auto. - apply allp_left with any. apply andp_left2; auto. -Qed. - -Lemma allp_andp2 {A}{ND: NatDed A}: forall B (any: B) p (q: B -> A), - andp p (allp q) = (allp (fun x => andp p (q x))). -Proof. -intros. rewrite andp_comm. rewrite allp_andp1; auto. -f_equal. extensionality x. rewrite andp_comm; auto. -Qed.*) - -(*Lemma valid_pointer_offset_val_zero: - forall p, valid_pointer (offset_val 0 p) ⊣⊢ valid_pointer p. -Proof. - This isn't true, since nullval is valid but can't be offset. -Admitted.*) - -Class listspec {cs: compspecs} (list_structid: ident) (list_link: ident) (token: share -> val -> mpred):= - mk_listspec { - list_fields: members; - list_struct := Tstruct list_structid noattr; - list_members_eq: list_fields = co_members (get_co list_structid); - list_struct_complete_legal_cosu: complete_legal_cosu_type list_struct = true; (* TODO: maybe this line not useful? *) - list_link_type: nested_field_type list_struct (StructField list_link :: nil) = Tpointer list_struct noattr; - list_token := token; - list_plain: plain_members list_fields = true -}. - -Section LIST1. -Context {cs: compspecs}. -Context {list_structid: ident} {list_link: ident} {list_token: share -> val -> mpred}. - -Fixpoint all_but_link (f: members) : members := - match f with - | nil => nil - | cons it f' => if ident_eq (name_member it) list_link - then f' - else cons it (all_but_link f') - end. - -Lemma list_link_size_in_range (ls: listspec list_structid list_link list_token): - 0 < sizeof (nested_field_type list_struct (StructField list_link :: nil)) < Ptrofs.modulus. -Proof. - rewrite list_link_type. - unfold sizeof, Ctypes.sizeof. - destruct Archi.ptr64 eqn:Hp. - rewrite Ptrofs.modulus_eq64 by auto; computable. - rewrite Ptrofs.modulus_eq32 by auto; computable. -Qed. - -Definition elemtype (ls: listspec list_structid list_link list_token) := - compact_prod - (map (fun it => reptype (field_type (name_member it) list_fields)) (all_but_link list_fields)). - -Definition field_type' (F: members) (it: member) := - reptype (field_type (name_member it) F). - -Definition add_link_back' {F f: members} - (v: compact_prod (map (field_type' F) (all_but_link f))) : - compact_prod (map (field_type' F) f). - induction f as [| it0 f]. - + exact tt. - + destruct f as [| it1 f0]. - - exact (default_val _). - - change (all_but_link (it0 :: it1 :: f0)) - with (if ident_eq (name_member it0) list_link - then it1::f0 - else cons it0 (all_but_link (it1::f0))) - in v. - change (reptype (field_type (name_member it0) F) * compact_prod (map (field_type' F) (it1::f0)))%type. - destruct (ident_eq (name_member it0) list_link). - exact (default_val _, v). - destruct (all_but_link (it1 :: f0)) eqn:?. - simpl in Heqm. - destruct (ident_eq (name_member it1) list_link); [ | discriminate Heqm]. - subst f0. - exact (v, default_val _). - exact (fst v, IHf (snd v)). -Defined. - -Definition add_link_back - (F f : members) - (v : compact_prod - (map (fun it : member => reptype (field_type (name_member it) F)) - (all_but_link f))) - : compact_prod (map (fun it => reptype (field_type (name_member it) F)) f) - := -list_rect - (fun f0 : list (member) => - compact_prod (map (field_type' F) (all_but_link f0)) -> - compact_prod (map (field_type' F) f0)) - (fun _ : compact_prod (map (field_type' F) (all_but_link nil)) => tt) - (fun (it0 : member) (f0 : list member) - (IHf : compact_prod (map (field_type' F) (all_but_link f0)) -> - compact_prod (map (field_type' F) f0)) - (v0 : compact_prod (map (field_type' F) (all_but_link (it0 :: f0)))) => - match - f0 as l - return - (compact_prod (map (field_type' F) (all_but_link (it0 :: l))) -> - (compact_prod (map (field_type' F) (all_but_link l)) -> - compact_prod (map (field_type' F) l)) -> - compact_prod (map (field_type' F) (it0 :: l))) - with - | nil => - fun - (_ : compact_prod (map (field_type' F) (all_but_link (it0 :: nil)))) - (_ : compact_prod (map (field_type' F) (all_but_link nil)) -> - compact_prod (map (field_type' F) nil)) => - default_val (field_type (name_member it0) F) - | it1 :: f1 => - fun - (v1 : compact_prod - (map (field_type' F) (all_but_link (it0 :: it1 :: f1)))) - (IHf0 : compact_prod - (map (field_type' F) (all_but_link (it1 :: f1))) -> - compact_prod (map (field_type' F) (it1 :: f1))) => - (if ident_eq (name_member it0) list_link as s0 - return - (compact_prod - (map (field_type' F) - (if s0 then it1 :: f1 else it0 :: all_but_link (it1 :: f1))) -> - reptype (field_type (name_member it0) F) * - compact_prod (map (field_type' F) (it1 :: f1))) - then - fun v2 : compact_prod (map (field_type' F) (it1 :: f1)) => - (default_val (field_type (name_member it0) F), v2) - else - fun - v2 : compact_prod - (map (field_type' F) (it0 :: all_but_link (it1 :: f1))) => - match - all_but_link (it1 :: f1) as l - return - (all_but_link (it1 :: f1) = l -> - compact_prod (map (field_type' F) (it0 :: l)) -> - (compact_prod (map (field_type' F) l) -> - compact_prod (map (field_type' F) (it1 :: f1))) -> - reptype (field_type (name_member it0) F) * - compact_prod (map (field_type' F) (it1 :: f1))) - with - | nil => - fun (Heqm0 : all_but_link (it1 :: f1) = nil) - (v3 : compact_prod (map (field_type' F) (it0 :: nil))) - (IHf1 : compact_prod (map (field_type' F) nil) -> - compact_prod (map (field_type' F) (it1 :: f1))) => - let s0 := ident_eq (name_member it1) list_link in - (if s0 - return - ((if s0 then f1 else it1 :: all_but_link f1) = nil -> - reptype (field_type (name_member it0) F) * - compact_prod (map (field_type' F) (it1 :: f1))) - then - fun Heqm1 : f1 = nil => - eq_rect_r - (fun f2 : members => - (compact_prod (map (field_type' F) nil) -> - compact_prod (map (field_type' F) (it1 :: f2))) -> - reptype (field_type (name_member it0) F) * - compact_prod (map (field_type' F) (it1 :: f2))) - (fun - _ : compact_prod (map (field_type' F) nil) -> - compact_prod (map (field_type' F) (it1 :: nil)) => - (v3, default_val (field_type (name_member it1) F))) - Heqm1 IHf1 - else - fun Heqm1 : it1 :: all_but_link f1 = nil => - False_rect - (reptype (field_type (name_member it0) F) * - compact_prod (map (field_type' F) (it1 :: f1))) - (eq_rect (it1 :: all_but_link f1) - (fun e : members => - match e with - | nil => False%type - | _ :: _ => True%type - end) I nil Heqm1)) Heqm0 - | p :: m0 => - fun (_ : all_but_link (it1 :: f1) = p :: m0) - (v3 : compact_prod (map (field_type' F) (it0 :: p :: m0))) - (IHf1 : compact_prod (map (field_type' F) (p :: m0)) -> - compact_prod (map (field_type' F) (it1 :: f1))) => - (fst v3, IHf1 (snd v3)) - end eq_refl v2 IHf0) v1 - end v0 IHf) f v. - - -Definition list_data {ls: listspec list_structid list_link list_token} (v: elemtype ls): reptype list_struct. - unfold list_struct. - pose (add_link_back _ _ v: reptype_structlist _). - rewrite list_members_eq in r. - exact (@fold_reptype _ (Tstruct _ _) r). -Defined. - -Definition list_cell' (ls: listspec list_structid list_link list_token) sh v p := - (field_at_ sh list_struct (StructField list_link :: nil) p) -* (data_at sh list_struct (list_data v) p). - -Definition list_cell (ls: listspec list_structid list_link list_token) (sh: Share.t) - (v: elemtype ls) (p: val) : mpred := - !! field_compatible list_struct nil p && - struct_pred (all_but_link list_fields) - (fun it v => withspacer sh - (field_offset cenv_cs (name_member it) list_fields + sizeof (field_type (name_member it) list_fields)) - (field_offset_next cenv_cs (name_member it) list_fields (co_sizeof (get_co list_structid))) - (at_offset (data_at_rec sh (field_type (name_member it) list_fields) v) (field_offset cenv_cs (name_member it) list_fields))) - v p. - -Lemma struct_pred_type_changable: - forall m m' A F v v' p p', - m=m' -> - JMeq v v' -> - (forall it v, F it v p = F it v p') -> - struct_pred m (A := A) F v p = struct_pred m' (A := A) F v' p'. -Proof. -intros. -subst m'. apply JMeq_eq in H0. subst v'. -induction m. reflexivity. -destruct m. -destruct a; simpl; apply H1. -rewrite !struct_pred_cons2. -f_equal. -auto. -apply IHm. -Qed. - -Lemma list_cell_link_join: - forall (LS: listspec list_structid list_link list_token) sh v p, - list_cell LS sh v p - * spacer sh (field_offset cenv_cs list_link list_fields + - sizeof (field_type list_link list_fields)) - (field_offset_next cenv_cs list_link list_fields - (co_sizeof (get_co list_structid))) - (offset_val 0 p) - * field_at_ sh list_struct (StructField list_link :: nil) p - = data_at sh list_struct (list_data v) p. -Proof. -unfold list_cell, data_at_, data_at, field_at_, field_at; intros. -(*destruct (field_compatible_dec list_struct nil p); - [ | solve [apply pred_ext; normalize]].*) -Admitted. -(* -rewrite <- !gather_prop_left. -rewrite !(prop_true_andp _ _ f). -rewrite (prop_true_andp (field_compatible list_struct (StructField list_link :: nil) p)) - by admit. -normalize. -apply andp_prop_ext. -admit. -intro HV. -clear HV. -change (nested_field_type list_struct nil) with list_struct. -rewrite (data_at_rec_ind sh list_struct - (@fold_reptype cs (Tstruct list_structid noattr) - (@eq_rect members - (@list_fields cs list_structid list_link LS) - (fun m : members => @reptype_structlist cs m) - (@add_link_back - (@list_fields cs list_structid list_link LS) - (@list_fields cs list_structid list_link LS) v) - (co_members (@get_co cs list_structid)) - (@list_members_eq cs list_structid list_link LS)))). -simpl. -forget (co_sizeof (get_co list_structid)) as sz. -assert (FT: field_type list_link list_fields = tptr list_struct). - admit. -pose (P m := fun (it : ident * type) (v0 : reptype (field_type (fst it) m)) => - withspacer sh - (field_offset cenv_cs (fst it) m + - sizeof (field_type (fst it) m)) - (field_offset_next cenv_cs (fst it) m sz) - (at_offset (data_at_rec sh (field_type (fst it) m) v0) - (field_offset cenv_cs (fst it) m))). -fold (P list_fields). -fold (P (co_members (get_co list_structid))). -transitivity - (at_offset - (struct_pred (co_members (get_co list_structid)) - (P (co_members (get_co list_structid))) - ( - (eq_rect list_fields (fun m : members => reptype_structlist m) - (add_link_back _ _ v) (co_members (get_co list_structid)) - list_members_eq))) (nested_field_offset list_struct nil) p); - [ | f_equal; f_equal; rewrite unfold_fold_reptype; reflexivity ]. -unfold at_offset. -rewrite (data_at_rec_type_changable sh - (nested_field_type list_struct (StructField list_link :: nil)) - (tptr list_struct) - (default_val _) Vundef - list_link_type) - by (rewrite by_value_default_val; try reflexivity; - rewrite list_link_type; reflexivity). -set (ofs := Int.repr (nested_field_offset list_struct (StructField list_link :: nil))). -assert (Hofs: ofs = Int.repr (field_offset cenv_cs list_link list_fields)). { - unfold ofs. - clear. - f_equal. - unfold list_struct. - pose proof list_link_type. - unfold nested_field_offset. - simpl. rewrite list_members_eq. - unfold list_struct, nested_field_type in H; simpl in H. - destruct (compute_in_members list_link (co_members (get_co list_structid))); inv H. - reflexivity. - } -revert v; unfold elemtype. -fold (field_type' list_fields). -pose (m := list_fields). -pose (m' := co_members (get_co list_structid)). -set (H := list_members_eq). -clearbody H. -revert H. -change (forall (H: m=m') - (v : compact_prod (map (field_type' list_fields) (all_but_link m))), -struct_pred (all_but_link m) (P list_fields) v p * -spacer sh - (field_offset cenv_cs list_link list_fields + - sizeof (field_type list_link list_fields)) - (field_offset_next cenv_cs list_link list_fields sz) - p* -data_at_rec sh (tptr list_struct) Vundef - (offset_val ofs p) = -struct_pred m' - (P m') - (eq_rect m reptype_structlist - (add_link_back list_fields m v) m' H) - (offset_val (Int.repr 0) p)). -assert (MNR := get_co_members_no_replicate list_structid). -fold m' in MNR. -revert MNR. -clearbody m'. -intros. -subst m'. -rewrite <- eq_rect_eq. -assert (In list_link (map fst m)). { - unfold m. - rewrite list_members_eq. - pose proof list_link_type. - unfold nested_field_type in H. - unfold list_struct in H. unfold nested_field_rec in H. - destruct (compute_in_members list_link (co_members (get_co list_structid))) - eqn:?; inv H. - apply compute_in_members_true_iff; auto. -} -change (struct_pred (all_but_link m) (P list_fields) v p * -spacer sh - (field_offset cenv_cs list_link list_fields + - sizeof (field_type list_link list_fields)) - (field_offset_next cenv_cs list_link list_fields sz) - p* -data_at_rec sh (tptr list_struct) Vundef (offset_val ofs p) = -struct_pred m (P list_fields) - (add_link_back list_fields m v) - (offset_val (Int.repr 0) p)). -revert MNR H v; clearbody m. -induction m; intros; [inv H | ]. - simpl in H. - assert (H': In list_link (map fst m) -> fst a <> list_link). - clear - MNR. unfold members_no_replicate in MNR. - intros; simpl in *. destruct (id_in_list (fst a) (map fst m)) eqn:?. inv MNR. - apply id_in_list_false in Heqb. intro. congruence. - destruct H. -* (* list_link is the first field *) -clear H'. -destruct a. simpl in H. subst i. -destruct m. -Opaque field_offset. Opaque field_type. simpl. -Transparent field_offset. Transparent field_type. -assert ((if ident_eq list_link list_link then nil else (list_link, t) :: nil) = nil) - by (rewrite if_true; auto). -simpl in v. -assert (exists v' : compact_prod (map (field_type' list_fields) nil), JMeq v' v). { - revert H v. - clear. - pose (j := if ident_eq list_link list_link - then @nil (ident * type) else (list_link, t) :: @nil (ident * type)). - change (j = nil -> - forall - v : compact_prod (map (field_type' list_fields) j), - exists v' : compact_prod (map (field_type' list_fields) nil), JMeq v' v). - clearbody j. - intros; subst. exists v; reflexivity. -} -destruct H0 as [v' ?]. -replace (struct_pred - (if ident_eq list_link list_link then nil else (list_link, t) :: nil) - (P list_fields) v p) with - (struct_pred nil (P list_fields) v' p). -Focus 2. -if_tac; [ | congruence]. reflexivity. -Opaque field_offset. Opaque field_type. simpl. -Transparent field_offset. Transparent field_type. -rewrite emp_sepcon. -clear v' H0 H v IHm. -unfold P. -rewrite withspacer_spacer. -unfold at_offset. simpl @fst. -f_equal. -rewrite isptr_offset_val_zero by auto. -auto. -rewrite offset_offset_val, Int.add_zero_l. -rewrite Hofs. -apply equal_f. -apply data_at_rec_type_changable; auto. -rewrite FT. reflexivity. -assert (all_but_link ((list_link,t)::p0::m) = p0::m). -simpl. rewrite if_true by auto; reflexivity. -assert (all_but_link (p0::m) = p0::m). { - clear - MNR H. - admit. (* easy enough *) -} -rewrite struct_pred_cons2. -unfold P at 2. -rewrite withspacer_spacer. -rewrite Hofs. unfold at_offset. -rewrite offset_offset_val, Int.add_zero_l. -change (fst (list_link, t)) with list_link. -rewrite isptr_offset_val_zero by auto. -pull_right (spacer sh - (field_offset cenv_cs list_link list_fields + - sizeof (field_type list_link list_fields)) - (field_offset_next cenv_cs list_link list_fields sz) p). -f_equal. -rewrite sepcon_comm. -f_equal. -apply equal_f. -apply data_at_rec_type_changable; auto. -apply JMeq_trans with (B:= reptype (field_type list_link list_fields)) (y:= default_val (field_type list_link list_fields)). -rewrite FT. reflexivity. -match goal with |- JMeq ?A ?B => replace A with B end. -apply JMeq_refl. -clear. -revert v. -unfold all_but_link. -unfold add_link_back. -unfold list_rect at 1. -simpl @fst. -destruct (ident_eq list_link list_link); [ | exfalso; congruence]; intro. -simpl. reflexivity. - apply struct_pred_type_changable; auto. - clear. - revert v. - simpl. - destruct (ident_eq list_link list_link); [ | exfalso; congruence]; intro. - simpl. reflexivity. -* (* list link is not the first field *) - specialize (H' H). -destruct m; [inv H | ]. - rewrite struct_pred_cons2. - assert (all_but_link (a :: p0 :: m) = a :: all_but_link (p0::m)). { - clear - MNR H. forget (p0::m) as m'. clear p0 m. - induction m'. inv H. - unfold all_but_link; fold all_but_link. - unfold members_no_replicate in *. - rewrite map_cons in MNR. - unfold compute_list_norepet in MNR. - fold compute_list_norepet in MNR. - destruct (id_in_list (fst a) (map fst (a0 :: m'))) eqn:?; [discriminate | ]. - simpl in Heqb. rewrite orb_false_iff in Heqb. destruct Heqb. - apply Pos.eqb_neq in H0. - apply id_in_list_false in H1. - simpl in H. destruct H. - rewrite H in *. rewrite if_false by auto. auto. - rewrite if_false by congruence. auto. -} - unfold members_no_replicate in *. - simpl in MNR. - destruct ((fst a =? fst p0)%positive || id_in_list (fst a) (map fst m))%bool eqn:?; try discriminate. - rewrite orb_false_iff in Heqb. destruct Heqb. - apply Pos.eqb_neq in H1. - apply id_in_list_false in H2. - specialize (IHm MNR H). - destruct p0 as [i t]. -(* simpl in v'. *) - simpl in H1. clear MNR H. - destruct (ident_eq i list_link). - + subst i. - assert (exists v' : compact_prod (map (field_type' list_fields) (a :: m)), - JMeq v v'). { - revert v; clear - H0. - replace (all_but_link ((list_link, t) :: m)) with m in H0 - by (simpl; rewrite if_true by auto; auto). - rewrite H0. eexists; eauto. - } - destruct H as [v' H3]. - simpl in v'. - destruct m. - - - simpl in v'. - assert (exists v'': compact_prod - (map (field_type' list_fields) - (all_but_link ((list_link, t) :: nil))), JMeq v'' tt). { - clear. simpl. rewrite if_true by auto. exists tt; reflexivity. - } - destruct H as [v'' H4]. - specialize (IHm v''). - replace (struct_pred (all_but_link ((list_link, t) :: nil)) (P list_fields) v'' p) with - (struct_pred nil (P list_fields) tt p) in IHm - by (apply struct_pred_type_changable; auto; simpl; rewrite if_true; auto). - change (struct_pred nil (P list_fields) tt p) with emp in IHm. - rewrite emp_sepcon in IHm. - rewrite sepcon_assoc. rewrite IHm; clear IHm. - f_equal. - assert (exists v4: compact_prod - (map (field_type' list_fields) (a::nil)), JMeq v4 v). { - clear - H1. revert v. simpl. rewrite if_false by auto. rewrite if_true by auto. - eexists; eauto. - } - destruct H as [v4 H5]. - transitivity (struct_pred (a :: nil) (P list_fields) v4 (offset_val (Int.repr 0) p)). - apply struct_pred_type_changable; auto. - simpl. rewrite if_false by auto; rewrite if_true by auto. auto. - admit. (* see proof above *) - destruct a as [i' t']. - unfold struct_pred at 1. - unfold list_rect. - f_equal. - clear - H1 H5. simpl in H1. - admit. (* tedious *) - apply struct_pred_type_changable; auto. - clear - H1 H4 H3. - simpl in v'. - admit. (* tedious *) - - - simpl map at 1 in v'. cbv beta iota in v'. - destruct v' as [va vr]. - assert (exists vr' : compact_prod - (map (field_type' list_fields) - (all_but_link ((list_link, t) :: p0 :: m))), - JMeq vr vr'). { - clear - H1; simpl in H1. - simpl. rewrite if_true by auto. exists vr; eauto. - } destruct H as [vr' H4]. - specialize (IHm vr'). - replace (struct_pred (all_but_link (a :: (list_link, t) :: p0 :: m)) (P list_fields) v p) - with (P list_fields a - (fst (add_link_back list_fields (a :: (list_link, t) :: p0 :: m) v)) - (offset_val (Int.repr 0) p) * - struct_pred (all_but_link ((list_link, t) :: p0 :: m)) (P list_fields) vr' p). - rewrite !sepcon_assoc. f_equal. - rewrite <- sepcon_assoc. - rewrite IHm. - apply struct_pred_type_changable; auto. - clear - H3 H4 H1. - admit. (* tedious *) - clear - H3 H4 H1 H0. - transitivity (P list_fields a va p * - struct_pred (all_but_link ((list_link, t) :: p0 :: m)) (P list_fields) vr' p). - f_equal. - unfold P; rewrite !withspacer_spacer; f_equal. rewrite <- spacer_offset_zero. auto. - unfold at_offset. rewrite offset_offset_val. rewrite Int.add_zero_l. - f_equal. - admit. (* tedious *) - assert (exists v6: compact_prod (map (field_type' list_fields) (a :: p0 :: m)), - JMeq v v6). { - clear - H1 H0. - simpl all_but_link at 2 in H0. rewrite if_true in H0 by auto. - revert v; rewrite H0. intros. exists v; auto. - } destruct H as [v6 H]. - transitivity (struct_pred (a :: p0 :: m) (P list_fields) v6 p). - rewrite struct_pred_cons2. f_equal. - unfold P; rewrite !withspacer_spacer; f_equal. - unfold at_offset. - f_equal. rewrite H in H3. clear - H3. apply JMeq_eq in H3. subst; reflexivity. - apply struct_pred_type_changable; auto. - simpl. rewrite if_true by auto. auto. - rewrite H in H3. - clear - H3 H4. - eapply JMeq_trans. apply JMeq_sym. apply H4. - destruct v6. - clear - H3. simpl. - apply JMeq_eq in H3. inv H3; auto. - apply struct_pred_type_changable; auto. - simpl. rewrite if_false by auto. rewrite if_true by auto. auto. - + - assert (all_but_link ((i,t)::m) = (i,t)::all_but_link m). - simpl. rewrite if_false by auto; auto. - assert (exists v' : - (field_type' list_fields a * compact_prod (map (field_type' list_fields) (all_but_link ((i, t) :: m)))), JMeq v v'). { - clear - H H0 v. revert v; rewrite H0. rewrite H. - simpl. intros. exists v; reflexivity. - } destruct H3 as [v' Hv']. - destruct v' as [v1 vr]. - specialize (IHm vr). - replace (struct_pred (all_but_link (a :: (i, t) :: m)) (P list_fields) v p) - with (P list_fields a (fst (add_link_back list_fields (a :: (i, t) :: m) v)) - (offset_val (Int.repr 0) p) * - struct_pred (all_but_link ((i, t) :: m)) (P list_fields) vr p). - rewrite !sepcon_assoc. f_equal. - rewrite <- sepcon_assoc. - rewrite IHm. clear IHm. - apply struct_pred_type_changable; auto. - admit. (* tedious *) - assert (exists v'': compact_prod - (field_type' list_fields a :: field_type' list_fields (i,t) :: map (field_type' list_fields) (all_but_link m)), - JMeq v v''). { - clear - H H0. revert v; rewrite H0. rewrite H. intros; exists v. reflexivity. - } destruct H3 as [v'' Hv'']. - transitivity (struct_pred (a :: (i,t) :: all_but_link m) (P list_fields) v'' p). - rewrite struct_pred_cons2. - f_equal. - admit. (* tedious *) - apply struct_pred_type_changable; auto. - clear - Hv' Hv''. rewrite Hv'' in Hv'. simpl in v''. destruct v''. - clear - Hv'. - admit. (* tedious *) - apply struct_pred_type_changable; auto. - rewrite H0. rewrite H. auto. -Qed. -*) -Lemma list_cell_link_join_nospacer: - forall (LS: listspec list_structid list_link list_token) sh v p, - field_offset cenv_cs list_link list_fields + - sizeof (field_type list_link list_fields) = - field_offset_next cenv_cs list_link list_fields - (co_sizeof (get_co list_structid)) -> - list_cell LS sh v p * field_at_ sh list_struct (StructField list_link :: nil) p - ⊣⊢ data_at sh list_struct (list_data v) p. -Proof. -intros. -rewrite <- list_cell_link_join. -unfold spacer. rewrite if_true. rewrite bi.sep_emp. auto. -lia. -Qed. - -End LIST1. - -Module LsegGeneral. - -Section LIST2. -Context {cs: compspecs}. -Context {list_structid: ident} {list_link: ident} {list_token: share -> val -> mpred}. - -Fixpoint lseg (ls: listspec list_structid list_link list_token) (dsh psh: share) - (contents: list (val * elemtype ls)) (x z: val) : mpred := - match contents with - | (p,h)::hs => !! (p=x /\ ~ptr_eq x z) && - EX y:val, !! is_pointer_or_null y && - list_token dsh x * list_cell ls dsh h x - * field_at psh list_struct (StructField list_link ::nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x - * lseg ls dsh psh hs y z - | nil => !! (ptr_eq x z) && emp - end. - -Lemma lseg_unfold (ls: listspec list_structid list_link list_token): forall dsh psh contents v1 v2, - lseg ls dsh psh contents v1 v2 = - match contents with - | (p,h)::t => !! (p=v1 /\ ~ ptr_eq v1 v2) && EX tail: val, - !! is_pointer_or_null tail && - list_token dsh v1 * list_cell ls dsh h v1 - * field_at psh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) tail) v1 - * lseg ls dsh psh t tail v2 - | nil => !! (ptr_eq v1 v2) && emp - end. -Proof. - intros. - destruct contents as [ | [? ?] ?]; simpl; auto. -Qed. - -Lemma lseg_eq (ls: listspec list_structid list_link list_token): - forall dsh psh l v , - is_pointer_or_null v -> - lseg ls dsh psh l v v ⊣⊢ !!(l=nil) && emp. -Proof. -intros. -rewrite (lseg_unfold ls dsh psh l v v). -destruct l. -f_equiv. f_equiv. -split; intro; auto. -unfold ptr_eq. -unfold is_pointer_or_null in H. -destruct Archi.ptr64 eqn:Hp; -destruct v; inv H; auto; -unfold Ptrofs.cmpu; rewrite Ptrofs.eq_true; auto. -destruct p. -apply pred_ext; -apply bi.pure_elim_l; intro. -destruct H0. -contradiction H1. -destruct v; inv H; try split; auto; apply Ptrofs.eq_true. -inv H0. -Qed. - -Definition lseg_cons (ls: listspec list_structid list_link list_token) dsh psh (l: list (val * elemtype ls)) (x z: val) : mpred := - !! (~ ptr_eq x z) && - EX h:(elemtype ls), EX r:list (val * elemtype ls), EX y:val, - !!(l=(x,h)::r /\ is_pointer_or_null y) && - list_token dsh x * list_cell ls dsh h x * - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x * - lseg ls dsh psh r y z. - -Lemma lseg_unroll (ls: listspec list_structid list_link list_token): forall dsh psh l x z , - lseg ls dsh psh l x z ⊣⊢ - (!! (ptr_eq x z) && !! (l=nil) && emp) || lseg_cons ls dsh psh l x z. -Proof. -intros. -rewrite lseg_unfold at 1. -apply pred_ext; destruct l. -apply bi.pure_elim_l; intros. -rewrite prop_true_andp by auto. -rewrite prop_true_andp by auto. -apply bi.or_intro_l; auto. -destruct p. -rewrite <- bi.or_intro_r. -unfold lseg_cons. -apply bi.pure_elim_l; intros. -destruct H. -apply bi.exist_elim; intro tail. -normalize. -rewrite <- (bi.exist_intro e). rewrite TT_andp. -apply bi.exist_intro with l. -apply bi.exist_intro with tail. -repeat rewrite sepcon_andp_prop'. -apply andp_right. -apply prop_right; split; auto. -subst. -auto. -subst. auto. -apply orp_left. -rewrite andp_assoc; -do 2 (apply bi.pure_elim_l; intro). - rewrite prop_true_andp by auto. auto. -unfold lseg_cons. -apply bi.pure_elim_l; intros. -apply bi.exist_elim; intro h. -apply bi.exist_elim; intro r. -apply bi.exist_elim; intro y. -do 3 rewrite sepcon_andp_prop'. -apply bi.pure_elim_l; intros [? ?]. -inv H0. -destruct p. -apply orp_left. -rewrite andp_assoc; -do 2 (apply bi.pure_elim_l; intro). -inv H0. -unfold lseg_cons. -apply bi.pure_elim_l; intros. -apply bi.exist_elim; intro h. -apply bi.exist_elim; intro r. -apply bi.exist_elim; intro y. -do 3 rewrite sepcon_andp_prop'. -apply bi.pure_elim_l; intros [? ?]. -symmetry in H0; inv H0. - rewrite prop_true_andp by auto. -apply bi.exist_intro with y. -normalize. -Qed. - -Lemma lseg_unroll_nonempty1 (ls: listspec list_structid list_link list_token): - forall p P dsh psh h tail v1 v2, - ~ ptr_eq v1 v2 -> - is_pointer_or_null p -> - (P |-- list_token dsh v1 * list_cell ls dsh h v1 * - (field_at psh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) p) v1 * - lseg ls dsh psh tail p v2)) -> - P |-- lseg ls dsh psh ((v1,h)::tail) v1 v2. -Proof. intros. rewrite lseg_unroll. apply bi.or_intro_r. unfold lseg_cons. - rewrite prop_true_andp by auto. - apply bi.exist_intro with h. apply bi.exist_intro with tail. apply bi.exist_intro with p. - rewrite prop_true_andp by auto. - rewrite sepcon_assoc. - eapply derives_trans; [ apply H1 | ]. - apply sepcon_derives; auto. -Qed. - -Lemma lseg_neq (ls: listspec list_structid list_link list_token): - forall dsh psh s v v2, - ptr_neq v v2 -> - lseg ls dsh psh s v v2 = lseg_cons ls dsh psh s v v2. -intros. rewrite lseg_unroll. -apply pred_ext. apply orp_left; auto. -rewrite andp_assoc. -do 2 (apply bi.pure_elim_l; intro). -congruence. -apply bi.or_intro_r. auto. -Qed. - -Lemma lseg_nonnull (ls: listspec list_structid list_link list_token): - forall dsh psh s v, - typed_true (tptr list_struct) v -> - lseg ls dsh psh s v nullval = lseg_cons ls dsh psh s v nullval. -Proof. -intros. unfold nullval. -apply lseg_neq. -destruct v; inv H; intuition; try congruence. -intro. apply ptr_eq_e in H. -destruct Archi.ptr64 eqn:Hp; inv H. -inv H1. -intro. simpl in H. -destruct Archi.ptr64; congruence. -Qed. - -Lemma unfold_lseg_neq (ls: listspec list_structid list_link list_token): - forall P Q1 Q R (v v2: val) dsh psh (s: list (val * elemtype ls)), - (PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s v v2 :: R))) |-- - !! (ptr_neq v v2)) -> - PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s v v2 :: R))) |-- - EX hryp: elemtype ls * list (val * elemtype ls) * val * val, - match hryp with (h,r,y,p) => - !! (s=(p,h)::r /\ is_pointer_or_null y) && - !! (p=v) && - PROPx P (LOCALx Q - (SEPx (list_token dsh v :: list_cell ls dsh h v:: - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y v2 :: - R))) - end. -Proof. -intros. -apply derives_trans with -(PROPx P (LOCALx (Q1::Q) (SEPx (lseg_cons ls dsh psh s v v2 :: R)))). -apply derives_trans with -(!! ptr_neq v v2 && PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s v v2 :: R)))). -apply andp_right; auto. -intro rho; unfold PROPx,LOCALx,SEPx,local,tc_expr,tc_lvalue; unfold_lift; simpl. -unfold lift1; simpl. - repeat (apply bi.pure_elim_l; intro). - rewrite prop_true_andp by auto. - rewrite prop_true_andp by auto. -apply sepcon_derives; auto. -rewrite lseg_neq; auto. -intro rho; unfold PROPx,LOCALx,SEPx,local,tc_expr,tc_lvalue,lift2,lift1,lift0; simpl. - unfold_lift. - unfold lseg_cons. simpl. - apply bi.pure_elim_l; intro. - apply bi.pure_elim_l; intros [? ?]. - rewrite sepcon_andp_prop'. - apply bi.pure_elim_l; intro. - rewrite exp_sepcon1; apply bi.exist_elim; intro h. - rewrite exp_sepcon1; apply bi.exist_elim; intro r. - rewrite exp_sepcon1; apply bi.exist_elim; intro y. - repeat rewrite sepcon_andp_prop'. - apply bi.pure_elim_l; intros [? ?]. - subst. - apply bi.exist_intro with (h,r,y, v). - repeat rewrite prop_true_andp by auto. - repeat rewrite sepcon_assoc. - auto. -Qed. - -Lemma unfold_lseg_cons (ls: listspec list_structid list_link list_token): - forall P Q1 Q R e dsh psh s, - (PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s e nullval :: R))) |-- - !! (typed_true (tptr list_struct) e)) -> - PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s e nullval :: R))) |-- - EX hryp: elemtype ls * list (val * elemtype ls) * val * val, - match hryp with (h,r,y,p) => - !! (s=(p,h)::r /\ is_pointer_or_null y) && - !! (p=e)&& - PROPx P (LOCALx Q - (SEPx (list_token dsh e :: list_cell ls dsh h e :: - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) e :: - lseg ls dsh psh r y nullval :: - R))) - end. -Proof. -intros. apply unfold_lseg_neq. -eapply derives_trans. -apply H. normalize. -unfold local. super_unfold_lift. -unfold nullval. -intro. -apply ptr_eq_e in H1. subst. -normalize. -Qed. - -Lemma semax_lseg_neq (ls: listspec list_structid list_link list_token): - forall (Espec: OracleKind) - Delta P Q dsh psh s v v2 R c Post, - ~ (ptr_eq v v2) -> - (forall (h: elemtype ls) (r: list (val * elemtype ls)) (y: val), - s=(v,h)::r -> is_pointer_or_null y -> - semax Delta - (PROPx P (LOCALx Q - (SEPx (list_token dsh v :: list_cell ls dsh h v :: - field_at psh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y v2 :: - R)))) c Post) -> - semax Delta - (PROPx P (LOCALx Q (SEPx (lseg ls dsh psh s v v2 :: R)))) - c Post. -Proof. -intros. -rewrite lseg_neq by auto. -unfold lseg_cons. -apply semax_pre0 with - (EX h: elemtype ls, EX r: list (val * elemtype ls), EX y: val, - !!(s = (v, h) :: r /\ is_pointer_or_null y) && - PROPx P (LOCALx Q (SEPx (list_token dsh v :: list_cell ls dsh h v :: - field_at psh list_struct (StructField list_link :: nil) - (valinject - (nested_field_type list_struct - (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y v2 :: R)))). -go_lowerx; entailer. -Exists h r y. -rewrite <- ?sepcon_assoc. -normalize. - autorewrite with subst norm1 norm2; normalize. -Intros h r y. -apply semax_extract_prop; intros [? ?]. -eapply H0; eauto. -Qed. - - -Lemma semax_lseg_nonnull (ls: listspec list_structid list_link list_token): - forall (Espec: OracleKind) - Delta P Q dsh psh s v R c Post, - ENTAIL Delta, PROPx P (LOCALx Q - (SEPx (lseg ls dsh psh s v nullval :: R))) |-- - !!(typed_true (tptr list_struct) v) -> - (forall (h: elemtype ls) (r: list (val * elemtype ls)) (y: val), - s=(v,h)::r -> is_pointer_or_null y -> - semax Delta - (PROPx P (LOCALx Q - (SEPx (list_token dsh v :: list_cell ls dsh h v :: - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y nullval :: - R)))) c Post) -> - semax Delta - (PROPx P (LOCALx Q (SEPx (lseg ls dsh psh s v nullval :: R)))) - c Post. -Proof. -intros. -assert_PROP (~ ptr_eq v nullval). -eapply derives_trans; [apply H |]. -normalize. -apply semax_lseg_neq; auto. -Qed. - -Lemma lseg_nil_eq (ls: listspec list_structid list_link list_token): - forall dsh psh p q, lseg ls dsh psh nil p q = !! (ptr_eq p q) && emp. -Proof. intros. - rewrite lseg_unroll. - apply pred_ext. - apply orp_left. - rewrite andp_assoc. - apply andp_derives; auto. -rewrite prop_true_andp by auto. auto. - unfold lseg_cons. normalize. inv H0. - apply bi.or_intro_l. rewrite andp_assoc. - rewrite (prop_true_andp (_ = _)) by auto. auto. -Qed. - -Lemma lseg_cons_eq (ls: listspec list_structid list_link list_token): - forall dsh psh h r x z , - lseg ls dsh psh (h::r) x z = - !!(x = fst h /\ ~ ptr_eq x z) && - (EX y : val, - !!(is_pointer_or_null y) && - list_token dsh x * list_cell ls dsh (snd h) x * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x * - lseg ls dsh psh r y z). -Proof. - intros. rewrite lseg_unroll. - apply pred_ext. - apply orp_left. - rewrite andp_assoc. - apply bi.pure_elim_l; intro. - apply bi.pure_elim_l; intro. - inv H0. - unfold lseg_cons. - normalize. - symmetry in H0; inv H0. - apply bi.exist_intro with y. normalize. - autorewrite with subst norm1 norm2; normalize. - normalize. destruct h as [p h]. simpl in *. - apply bi.or_intro_r. - unfold lseg_cons. - rewrite prop_true_andp by auto. - apply bi.exist_intro with h. apply bi.exist_intro with r. apply bi.exist_intro with y. - normalize. - autorewrite with subst norm1 norm2; normalize. -Qed. - -Definition lseg_cons_right (ls: listspec list_structid list_link list_token) - dsh psh (l: list (val * elemtype ls)) (x z: val) : mpred := - !! (~ ptr_eq x z) && - EX h:(elemtype ls), EX r:list (val * elemtype ls), EX y:val, - !!(l=r++(y,h)::nil /\ is_pointer_or_null y) && - list_token dsh y * list_cell ls dsh h y * - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls dsh psh r x y. - -Lemma lseg_cons_right_neq (ls: listspec list_structid list_link list_token): forall dsh psh l x h y w z, - sepalg.nonidentity psh -> - list_token dsh y * list_cell ls dsh h y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls dsh psh l x y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) w) z - |-- lseg ls dsh psh (l++(y,h)::nil) x z * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) w) z. -Proof. -intros. rename H into SH. -assert (SZ: 0 < sizeof (nested_field_type list_struct (DOT list_link))). - unfold sizeof; rewrite list_link_type; simpl; destruct Archi.ptr64; computable. -rewrite (field_at_isptr _ _ _ _ z). -normalize. -revert x; induction l; simpl; intros. -* -normalize. - autorewrite with subst norm1 norm2; normalize. - apply bi.exist_intro with z. - entailer!. -* -destruct a as [v el]. -normalize. -apply bi.exist_intro with x0. -normalize. -rewrite <- ?sepcon_assoc. - autorewrite with subst norm1 norm2; normalize. -specialize (IHl x0). -entailer. -pull_right (list_cell ls dsh el x). -apply sepcon_derives; auto. -pull_right (field_at psh list_struct (StructField list_link :: nil) - (valinject - (nested_field_type list_struct (StructField list_link :: nil)) x0) - x). -pull_right (list_token dsh x). -apply sepcon_derives; auto. -apply sepcon_derives; auto. -Qed. - -Lemma lseg_cons_right_null (ls: listspec list_structid list_link list_token): forall dsh psh l x h y, - list_token dsh y * list_cell ls dsh h y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) nullval) y * - lseg ls dsh psh l x y - |-- lseg ls dsh psh (l++(y,h)::nil) x nullval. -Proof. -intros. -revert x; induction l; simpl; intros. -* -normalize. - autorewrite with subst norm1 norm2; normalize. -apply bi.exist_intro with nullval. -apply andp_right. -apply not_prop_right; intro. -apply ptr_eq_e in H. subst y. -entailer!. -destruct H. contradiction H. -rewrite prop_true_andp by reflexivity. -rewrite prop_true_andp - by (unfold nullval; destruct Archi.ptr64 eqn:Hp; simpl; auto). -normalize. -* -destruct a as [v el]. -normalize. -apply bi.exist_intro with x0. -normalize. - autorewrite with subst norm1 norm2; normalize. -specialize (IHl x0). -apply andp_right. -rewrite prop_and. -apply andp_right; [ | apply prop_right; auto]. -apply not_prop_right; intro. -apply ptr_eq_e in H0. subst x. -entailer. -destruct H2; contradiction H2. -eapply derives_trans. -2: apply sepcon_derives; [ | eassumption]; apply derives_refl. -clear IHl. -cancel. -Qed. - - -Lemma lseg_cons_right_list (ls: listspec list_structid list_link list_token): forall dsh psh l l' x h y z, - sepalg.nonidentity psh -> - list_token dsh y * list_cell ls dsh h y - * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y - * lseg ls dsh psh l x y - * lseg ls dsh psh l' z nullval - |-- lseg ls dsh psh (l++(y,h)::nil) x z * lseg ls dsh psh l' z nullval. -Proof. -intros. -destruct l'. -rewrite lseg_nil_eq. -normalize. -rewrite prop_true_andp by apply ptr_eq_nullval. -apply lseg_cons_right_null. -rewrite lseg_cons_eq. -Intros u. Exists u. subst z. -rewrite <- ?sepcon_assoc. -rewrite !prop_true_andp by auto. -normalize. -apply sepcon_derives; auto. -pull_right (list_cell ls dsh (snd p) (fst p)). -pull_right (list_token dsh (fst p)). -apply sepcon_derives; auto. -apply sepcon_derives; auto. -apply lseg_cons_right_neq; auto. -Qed. - -Lemma lseg_unroll_right (ls: listspec list_structid list_link list_token): forall sh sh' l x z , - lseg ls sh sh' l x z = (!! (ptr_eq x z) && !! (l=nil) && emp) || lseg_cons_right ls sh sh' l x z. -Abort. (* not likely true *) - - -Lemma lseg_local_facts: - forall ls dsh psh contents p q, - lseg ls dsh psh contents p q |-- - !! (is_pointer_or_null p /\ (p=q <-> contents=nil)). -Proof. -intros. -rewrite lseg_unfold. -destruct contents. -apply bi.pure_elim_l; intro. -unfold ptr_eq in H. -apply prop_right. -destruct p; try contradiction; simpl; auto. -destruct q; try contradiction; auto. -unfold Int.cmpu in H. -destruct H as [? [? ?]]. -apply int_eq_e in H0. -apply int_eq_e in H1. subst. rewrite H. -split; auto; split; auto. -destruct q; try contradiction; auto. -unfold Int64.cmpu in H. -destruct H as [? [? ?]]. -apply int64_eq_e in H0. -apply int64_eq_e in H1. subst. rewrite H. -split3; auto. -destruct q; try contradiction. -destruct H; subst. -unfold Ptrofs.cmpu in H0. -apply ptrofs_eq_e in H0. -subst. tauto. -destruct p0. -normalize. -rewrite field_at_isptr. -normalize. - autorewrite with subst norm1 norm2; normalize. -apply prop_right. -split. intro; subst q. -contradiction H. normalize. -intros. discriminate. -Qed. - -Definition lseg_cell (ls: listspec list_structid list_link list_token) - (dsh psh : share) - (v: elemtype ls) (x y: val) := - !!is_pointer_or_null y && list_token dsh x * list_cell ls dsh v x * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x. - -Lemma lseg_cons_eq2: forall - (ls : listspec list_structid list_link list_token) (dsh psh : share) (h : elemtype ls) - (r : list (val * elemtype ls)) - (x x' z : val), lseg ls dsh psh ((x',h) :: r) x z = - !!(x=x' /\ ~ ptr_eq x z) && (EX y : val, lseg_cell ls dsh psh h x y * lseg ls dsh psh r y z). -Proof. - intros. - rewrite -> lseg_cons_eq. - unfold lseg_cell. - normalize. -Qed. - -Lemma list_append: forall {dsh psh: share} - {ls : listspec list_structid list_link list_token} (hd mid tl:val) ct1 ct2 P, - (forall x tl', lseg_cell ls dsh psh x tl tl' * P tl |-- FF) -> - (lseg ls dsh psh ct1 hd mid) * lseg ls dsh psh ct2 mid tl * P tl|-- - (lseg ls dsh psh (ct1 ++ ct2) hd tl) * P tl. -Proof. - intros. - revert hd; induction ct1; simpl; intros; auto. - * - normalize. - * - destruct a as [v a]. - normalize. - autorewrite with subst norm1 norm2; normalize. - apply bi.exist_intro with y. - apply andp_right. - apply not_prop_right; intro. apply ptr_eq_e in H1; subst hd. - clear IHct1. - unfold lseg_cell in H. - specialize (H a y). - rewrite prop_true_andp in H by auto. - apply derives_trans with - (lseg ls dsh psh ct1 y mid * lseg ls dsh psh ct2 mid tl * FF). - cancel. auto. - rewrite sepcon_FF; auto. - normalize. - specialize (IHct1 y). clear H. - do 2 rewrite sepcon_assoc. - eapply derives_trans. - apply sepcon_derives. - apply derives_refl. - rewrite <- !sepcon_assoc; eassumption. - cancel. -Qed. - -Lemma list_append_null: - forall - (ls: listspec list_structid list_link list_token) - (dsh psh: share) - (hd mid: val) ct1 ct2, - lseg ls dsh psh ct1 hd mid * lseg ls dsh psh ct2 mid nullval |-- - lseg ls dsh psh (ct1++ct2) hd nullval. -Proof. -intros. - rewrite <- sepcon_emp. - eapply derives_trans; [ | apply (list_append hd mid nullval ct1 ct2 (fun _ => emp))]. - normalize. - intros. - unfold lseg_cell. simpl. saturate_local. destruct H. contradiction H. -Qed. - -Lemma sizeof_list_struct_pos (LS: listspec list_structid list_link list_token) : - sizeof list_struct > 0. -Admitted. - -End LIST2. - -#[export] Hint Rewrite @lseg_nil_eq : norm. - -#[export] Hint Rewrite @lseg_eq using reflexivity: norm. - -#[export] Hint Resolve lseg_local_facts : saturate_local. -End LsegGeneral. - -Module LsegSpecial. -Import LsegGeneral. - -Section LIST. -Context {cs: compspecs}. -Context {list_structid: ident} {list_link: ident} {list_token: share -> val -> mpred}. - -Definition lseg (ls: listspec list_structid list_link list_token) (sh: share) - (contents: list (elemtype ls)) (x y: val) : mpred := - EX al:list (val*elemtype ls), - !! (contents = map snd al) && - LsegGeneral.lseg ls sh sh al x y. - -Lemma lseg_unfold (ls: listspec list_structid list_link list_token): forall sh contents v1 v2, - lseg ls sh contents v1 v2 = - match contents with - | h::t => !! (~ ptr_eq v1 v2) && EX tail: val, - !! is_pointer_or_null tail && - list_token sh v1 * list_cell ls sh h v1 - * field_at sh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) tail) v1 - * lseg ls sh t tail v2 - | nil => !! (ptr_eq v1 v2) && emp - end. -Proof. - intros. - unfold lseg. - revert v1; induction contents; intros. - apply pred_ext. - normalize. destruct al; inv H. - rewrite LsegGeneral.lseg_nil_eq; auto. - apply bi.exist_intro with nil. - apply bi.pure_elim_l; intro. - normalize. - apply pred_ext. - apply bi.exist_elim; intros [ | [v1' a'] al]. - normalize. inv H. - apply bi.pure_elim_l; intro. - symmetry in H; inv H. - rewrite LsegGeneral.lseg_cons_eq; auto. - apply bi.pure_elim_l; intros [? ?]. - simpl in H; subst v1'. - apply bi.exist_elim; intro y. - normalize. apply bi.exist_intro with y. normalize. - repeat apply sepcon_derives; auto. - apply bi.exist_intro with al; normalize. - normalize. - apply bi.exist_intro with ((v1,a)::al); normalize. - simpl. - normalize. apply bi.exist_intro with tail. normalize. - autorewrite with subst norm1 norm2; normalize. -Qed. - -Lemma lseg_eq (ls: listspec list_structid list_link list_token): - forall sh l v , - is_pointer_or_null v -> - lseg ls sh l v v = !!(l=nil) && emp. -Proof. -intros. -unfold lseg. -apply pred_ext. -normalize. -rewrite LsegGeneral.lseg_eq by auto. normalize. -apply bi.exist_intro with nil. -normalize. -Qed. - -Definition lseg_cons (ls: listspec list_structid list_link list_token) sh (l: list (elemtype ls)) (x z: val) : mpred := - !! (~ ptr_eq x z) && - EX h:(elemtype ls), EX r:list (elemtype ls), EX y:val, - !!(l=h::r /\ is_pointer_or_null y) && - list_token sh x * list_cell ls sh h x * - field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x * - lseg ls sh r y z. - -Lemma lseg_unroll (ls: listspec list_structid list_link list_token): forall sh l x z , - lseg ls sh l x z = - (!! (ptr_eq x z) && !! (l=nil) && emp) || lseg_cons ls sh l x z. -Proof. -intros. -unfold lseg, lseg_cons. -apply pred_ext. -* -apply bi.exist_elim; intros. -apply bi.pure_elim_l; intro. -rewrite LsegGeneral.lseg_unroll. -apply orp_left; [apply bi.or_intro_l | apply bi.or_intro_r]. -rewrite andp_assoc; repeat (apply bi.pure_elim_l; intro). -subst. simpl. -normalize. -unfold LsegGeneral.lseg_cons. -apply bi.pure_elim_l; intro. -rewrite prop_true_andp by auto. -apply exp_derives; intro h. -apply bi.exist_elim; intro r; apply bi.exist_intro with (map snd r). -apply exp_derives; intro y. -normalize. -subst l. -unfold lseg. -cancel. -apply bi.exist_intro with r; normalize. -* -apply orp_left. -rewrite andp_assoc; repeat (apply bi.pure_elim_l; intro). -subst. -apply bi.exist_intro with nil. -simpl. normalize. - autorewrite with subst norm1 norm2; normalize. -apply bi.pure_elim_l; intro. -apply bi.exist_elim; intro h. -apply bi.exist_elim; intro r. -apply bi.exist_elim; intro y. -normalize. -unfold lseg. -normalize. -apply bi.exist_intro with ((x,h)::al). -normalize. -simpl. -normalize. -apply bi.exist_intro with y. -normalize. - autorewrite with subst norm1 norm2; normalize. -Qed. - -Lemma lseg_unroll_nonempty1 (ls: listspec list_structid list_link list_token): - forall p P sh h (tail: list (elemtype ls)) v1 v2, - ~ ptr_eq v1 v2 -> - is_pointer_or_null p -> - (P |-- list_token sh v1 * list_cell ls sh h v1 * - (field_at sh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) p) v1 * - lseg ls sh tail p v2)) -> - P |-- lseg ls sh (h::tail) v1 v2. -Proof. intros. rewrite lseg_unroll. apply bi.or_intro_r. unfold lseg_cons. - rewrite prop_true_andp by auto. - apply bi.exist_intro with h. apply bi.exist_intro with tail. apply bi.exist_intro with p. - rewrite prop_true_andp by auto. - rewrite sepcon_assoc. - eapply derives_trans; [ apply H1 | ]. - apply sepcon_derives; auto. -Qed. - -Lemma lseg_neq (ls: listspec list_structid list_link list_token): - forall sh s v v2, - ptr_neq v v2 -> - lseg ls sh s v v2 = lseg_cons ls sh s v v2. -intros. rewrite lseg_unroll. -apply pred_ext. apply orp_left; auto. -rewrite andp_assoc. -do 2 (apply bi.pure_elim_l; intro). -congruence. -apply bi.or_intro_r. auto. -Qed. - -Lemma lseg_nonnull (ls: listspec list_structid list_link list_token): - forall sh s v, - typed_true (tptr list_struct) v -> - lseg ls sh s v nullval = lseg_cons ls sh s v nullval. -Proof. -intros. unfold nullval. -apply lseg_neq. -unfold typed_true, strict_bool_val in H. -simpl in H. -destruct Archi.ptr64 eqn:Hp. -* -destruct v; inv H. -destruct (Int64.eq i Int64.zero); inv H1. -intro; apply ptr_eq_e in H; inv H. -* -destruct v; inv H. -destruct (Int.eq i Int.zero); inv H1. -intro; apply ptr_eq_e in H; inv H. -Qed. - -Lemma unfold_lseg_neq (ls: listspec list_structid list_link list_token): - forall P Q1 Q R (v v2: val) sh (s: list (elemtype ls)), - (PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls sh s v v2 :: R))) |-- - !! (ptr_neq v v2)) -> - PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls sh s v v2 :: R))) |-- - EX hryp: elemtype ls * list (elemtype ls) * val, - match hryp with (h,r,y) => - !! (s=h::r /\ is_pointer_or_null y) && - PROPx P (LOCALx Q - (SEPx (list_token sh v :: list_cell ls sh h v:: - field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls sh r y v2 :: - R))) - end. -Proof. -intros. -apply derives_trans with -(PROPx P (LOCALx (Q1::Q) (SEPx (lseg_cons ls sh s v v2 :: R)))). -apply derives_trans with -(!! (ptr_neq v v2) && PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls sh s v v2 :: R)))). -apply andp_right; auto. -intro rho; unfold PROPx,LOCALx,SEPx,local,tc_expr,tc_lvalue; unfold_lift; simpl. -unfold lift1; simpl. - repeat (apply bi.pure_elim_l; intro). - rewrite prop_true_andp by auto. - rewrite prop_true_andp by auto. -apply sepcon_derives; auto. -rewrite lseg_neq; auto. -intro rho; unfold PROPx,LOCALx,SEPx,local,tc_expr,tc_lvalue,lift2,lift1,lift0; simpl. - unfold_lift. - unfold lseg_cons. simpl. - apply bi.pure_elim_l; intro. - apply bi.pure_elim_l; intros [? ?]. - rewrite sepcon_andp_prop'. - apply bi.pure_elim_l; intro. - rewrite exp_sepcon1; apply bi.exist_elim; intro h. - rewrite exp_sepcon1; apply bi.exist_elim; intro r. - rewrite exp_sepcon1; apply bi.exist_elim; intro y. - repeat rewrite sepcon_andp_prop'. - apply bi.pure_elim_l; intros [? ?]. - subst. - apply bi.exist_intro with (h,r,y). - repeat rewrite prop_true_andp by auto. - repeat rewrite sepcon_assoc. - auto. -Qed. - -Lemma unfold_lseg_cons (ls: listspec list_structid list_link list_token): - forall P Q1 Q R e sh s, - (PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls sh s e nullval :: R))) |-- - !!(typed_true (tptr list_struct) e)) -> - PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls sh s e nullval :: R))) |-- - EX hryp: elemtype ls * list (elemtype ls) * val, - match hryp with (h,r,y) => - !! (s=h::r /\ is_pointer_or_null y) && - PROPx P (LOCALx Q - (SEPx (list_token sh e :: list_cell ls sh h e :: - field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) e :: - lseg ls sh r y nullval :: - R))) - end. -Proof. -intros. apply unfold_lseg_neq. -eapply derives_trans. -apply H. normalize. -unfold local. super_unfold_lift. -unfold nullval. -destruct e; inv H0; try congruence; auto. -intro. apply ptr_eq_e in H0. -destruct Archi.ptr64; inv H0. -Qed. - -Lemma semax_lseg_neq (ls: listspec list_structid list_link list_token): - forall (Espec: OracleKind) - Delta P Q sh s v v2 R c Post, - ~ (ptr_eq v v2) -> - (forall (h: elemtype ls) (r: list (elemtype ls)) (y: val), - s=h::r -> is_pointer_or_null y -> - semax Delta - (PROPx P (LOCALx Q - (SEPx (list_token sh v :: list_cell ls sh h v :: - field_at sh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls sh r y v2 :: - R)))) c Post) -> - semax Delta - (PROPx P (LOCALx Q (SEPx (lseg ls sh s v v2 :: R)))) - c Post. -Proof. -intros. -rewrite lseg_neq by auto. -unfold lseg_cons. -apply semax_pre0 with - (EX h: elemtype ls, EX r: list (elemtype ls), EX y: val, - !!(s = h :: r /\ is_pointer_or_null y) && - PROPx P (LOCALx Q (SEPx (list_token sh v :: list_cell ls sh h v :: - field_at sh list_struct (StructField list_link :: nil) - (valinject - (nested_field_type list_struct - (StructField list_link :: nil)) y) v :: - lseg ls sh r y v2 :: R)))). -go_lowerx; entailer. (* Intros h r y should work here, but doesn't. *) -Exists h r y. -rewrite <- ?sepcon_assoc. -normalize. - autorewrite with subst norm1 norm2; normalize. -Intros h r y. -apply semax_extract_prop; intros [? ?]. -eapply H0; eauto. -Qed. - - -Lemma semax_lseg_nonnull (ls: listspec list_structid list_link list_token): - forall (Espec: OracleKind) - Delta P Q sh s v R c Post, - ENTAIL Delta, PROPx P (LOCALx Q - (SEPx (lseg ls sh s v nullval :: R))) |-- - !!(typed_true (tptr list_struct) v) -> - (forall (h: elemtype ls) (r: list (elemtype ls)) (y: val), - s=h::r -> is_pointer_or_null y -> - semax Delta - (PROPx P (LOCALx Q - (SEPx (list_token sh v :: list_cell ls sh h v :: - field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls sh r y nullval :: - R)))) c Post) -> - semax Delta - (PROPx P (LOCALx Q (SEPx (lseg ls sh s v nullval :: R)))) - c Post. -Proof. -intros. -assert_PROP (~ ptr_eq v nullval). -eapply derives_trans; [apply H |]. -normalize. -apply semax_lseg_neq; auto. -Qed. - -Lemma lseg_nil_eq (ls: listspec list_structid list_link list_token): - forall sh p q, lseg ls sh nil p q = !! (ptr_eq p q) && emp. -Proof. intros. - rewrite lseg_unroll. - apply pred_ext. - apply orp_left. - rewrite andp_assoc. - apply andp_derives; auto. -rewrite prop_true_andp by auto. auto. - unfold lseg_cons. normalize. inv H0. - apply bi.or_intro_l. rewrite andp_assoc. - rewrite (prop_true_andp (_ = _)) by auto. auto. -Qed. - -Lemma lseg_cons_eq (ls: listspec list_structid list_link list_token): - forall sh h r x z , - lseg ls sh (h::r) x z = - !!(~ ptr_eq x z) && - (EX y : val, - !!(is_pointer_or_null y) && - list_token sh x * list_cell ls sh h x * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x * - lseg ls sh r y z). -Proof. - intros. rewrite lseg_unroll. - apply pred_ext. - apply orp_left. - rewrite andp_assoc. - apply bi.pure_elim_l; intro. - apply bi.pure_elim_l; intro. - inv H0. - unfold lseg_cons. - normalize. - symmetry in H0; inv H0. - apply bi.exist_intro with y. normalize. - apply bi.or_intro_r. - unfold lseg_cons. - apply andp_derives; auto. - apply bi.exist_intro with h. apply bi.exist_intro with r. apply exp_derives; intro y. - normalize. - autorewrite with subst norm1 norm2; normalize. -Qed. - -Definition lseg_cons_right (ls: listspec list_structid list_link list_token) - sh (l: list (elemtype ls)) (x z: val) : mpred := - !! (~ ptr_eq x z) && - EX h:(elemtype ls), EX r:list (elemtype ls), EX y:val, - !!(l=r++(h::nil) /\ is_pointer_or_null y) && - list_token sh y * list_cell ls sh h y * - field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls sh r x y. - -Lemma lseg_cons_right_neq (ls: listspec list_structid list_link list_token): forall sh l x h y w z, - sepalg.nonidentity sh -> - list_token sh y * list_cell ls sh h y * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls sh l x y * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) w) z - |-- lseg ls sh (l++h::nil) x z * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) w) z. -Proof. -intros. -unfold lseg. -normalize. -apply bi.exist_intro with (al ++ (y,h)::nil). -rewrite prop_true_andp by (rewrite map_app; reflexivity). -eapply derives_trans; [ | apply LsegGeneral.lseg_cons_right_neq; auto]. -cancel. -Qed. - -Lemma lseg_cons_right_null (ls: listspec list_structid list_link list_token): forall sh l x h y, - list_token sh y * list_cell ls sh h y * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) nullval) y * - lseg ls sh l x y - |-- lseg ls sh (l++h::nil) x nullval. -Proof. -intros. -unfold lseg. -normalize. -apply bi.exist_intro with (al ++ (y,h)::nil). -rewrite prop_true_andp by (rewrite map_app; reflexivity). -eapply derives_trans; [ | apply LsegGeneral.lseg_cons_right_null]. -cancel. -Qed. - - -Lemma lseg_cons_right_list (ls: listspec list_structid list_link list_token): forall sh l l' x h y z, - sepalg.nonidentity sh -> - list_token sh y * list_cell ls sh h y * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls sh l x y * lseg ls sh l' z nullval - |-- lseg ls sh (l++h::nil) x z * lseg ls sh l' z nullval. -Proof. -intros. -destruct l'. -rewrite lseg_nil_eq. -normalize. -rewrite prop_true_andp by apply ptr_eq_nullval. -apply lseg_cons_right_null. -rewrite lseg_cons_eq. -Intros u. -Exists u. -rewrite !prop_true_andp by auto. -rewrite <- !sepcon_assoc. -apply sepcon_derives; auto. -pull_right (list_cell ls sh e z). -pull_right (list_token sh z). -apply sepcon_derives; auto. -apply sepcon_derives; auto. -apply lseg_cons_right_neq. -auto. -Qed. - -Lemma lseg_unroll_right (ls: listspec list_structid list_link list_token): forall sh l x z , - lseg ls sh l x z = (!! (ptr_eq x z) && !! (l=nil) && emp) || lseg_cons_right ls sh l x z. -Abort. (* not likely true *) - -Lemma lseg_local_facts: - forall ls sh contents p q, - lseg ls sh contents p q |-- - !! (is_pointer_or_null p /\ (p=q <-> contents=nil)). -Proof. -intros. -unfold lseg. -normalize. -eapply derives_trans; [apply LsegGeneral.lseg_local_facts |]. -normalize. -split; auto. -rewrite H. -clear. -destruct al; simpl; intuition; try congruence. -Qed. - -Definition lseg_cell (ls: listspec list_structid list_link list_token) - (sh : share) - (v: elemtype ls) (x y: val) := - !!is_pointer_or_null y && list_token sh x * list_cell ls sh v x * field_at sh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x. - -Lemma lseg_cons_eq2: forall - (ls : listspec list_structid list_link list_token) (sh : share) (h : elemtype ls) - (r : list (elemtype ls)) - (x z : val), lseg ls sh (h :: r) x z = - !!(~ ptr_eq x z) && (EX y : val, lseg_cell ls sh h x y * lseg ls sh r y z). -Proof. - intros. - rewrite -> lseg_cons_eq. - unfold lseg_cell. - normalize. -Qed. - -Lemma list_append: forall {sh: share} - {ls : listspec list_structid list_link list_token} (hd mid tl:val) ct1 ct2 P, - (forall x tl', lseg_cell ls sh x tl tl' * P tl |-- FF) -> - (lseg ls sh ct1 hd mid) * lseg ls sh ct2 mid tl * P tl|-- - (lseg ls sh (ct1 ++ ct2) hd tl) * P tl. -Proof. - intros. - unfold lseg. - normalize. - eapply derives_trans. - apply LsegGeneral.list_append. - intros. - eapply derives_trans; [ | apply (H x0 tl')]. - unfold lseg_cell, LsegGeneral.lseg_cell. - entailer. - apply bi.exist_intro with (x++al). - rewrite prop_true_andp; auto. - rewrite map_app; reflexivity. -Qed. - -Lemma list_append_null: - forall - (ls: listspec list_structid list_link list_token) - (sh: share) - (hd mid: val) ct1 ct2, - lseg ls sh ct1 hd mid * lseg ls sh ct2 mid nullval |-- - lseg ls sh (ct1++ct2) hd nullval. -Proof. -intros. - rewrite <- sepcon_emp. - eapply derives_trans; [ | apply (list_append hd mid nullval ct1 ct2 (fun _ => emp))]. - normalize. - intros. - unfold lseg_cell. simpl. saturate_local. destruct H. contradiction H. -Qed. - -Lemma list_cell_valid_pointer: - forall (LS: listspec list_structid list_link list_token) (sh: Share.t) v p, - sepalg.nonidentity sh -> - field_offset cenv_cs list_link list_fields + sizeof (field_type list_link list_fields) - = field_offset_next cenv_cs list_link list_fields (co_sizeof (get_co list_structid)) -> - list_cell LS sh v p * field_at_ sh list_struct (StructField list_link::nil) p - |-- valid_pointer p. -Proof. - intros ? ? ? ? NON_ID ?. - rewrite list_cell_link_join_nospacer; auto. - unfold data_at_, field_at_, data_at. - eapply derives_trans; [ apply field_at_valid_ptr; auto | ]. - change (nested_field_type list_struct nil) with list_struct. - apply LsegGeneral.sizeof_list_struct_pos. - unfold field_address. - if_tac; auto. - change (Int.repr (nested_field_offset list_struct nil)) with Int.zero. - rewrite valid_pointer_offset_val_zero; auto. - simpl. - change predicates_hered.FF with FF. apply FF_left. -Qed. - -Lemma lseg_valid_pointer: - forall (ls : listspec list_structid list_link list_token) sh contents p q R, - sepalg.nonidentity sh -> - field_offset cenv_cs list_link list_fields + sizeof (field_type list_link list_fields) - = field_offset_next cenv_cs list_link list_fields (co_sizeof (get_co list_structid)) -> - (R |-- valid_pointer q) -> - R * lseg ls sh contents p q |-- valid_pointer p. -Proof. -intros ? ? ? ? ? ? NON_ID ? ?. -destruct contents. -rewrite lseg_nil_eq. normalize. -unfold lseg; simpl. -normalize. -destruct al; inv H1. -rewrite LsegGeneral.lseg_cons_eq. -normalize. -destruct p0 as [p z]; simpl in *. -apply sepcon_valid_pointer2. -apply sepcon_valid_pointer1. -rewrite sepcon_assoc. -apply sepcon_valid_pointer2. -eapply derives_trans; [ | eapply list_cell_valid_pointer; eauto]. -apply sepcon_derives ; [ apply derives_refl | ]. -cancel. -Qed. - -End LIST. - -#[export] Hint Rewrite @lseg_nil_eq : norm. -#[export] Hint Rewrite @lseg_eq using reflexivity: norm. -#[export] Hint Resolve lseg_local_facts : saturate_local. - -Ltac resolve_lseg_valid_pointer := -match goal with - | |- ?Q |-- valid_pointer ?p => - match Q with context [lseg ?A ?B ?C p ?q] => - repeat rewrite <- sepcon_assoc; - pull_right (lseg A B C p q); - apply lseg_valid_pointer; [auto | reflexivity | ]; - auto 50 with valid_pointer - end - end. - -#[export] Hint Extern 10 (_ |-- valid_pointer _) => - resolve_lseg_valid_pointer : valid_pointer. - -Lemma list_cell_local_facts: - forall {cs: compspecs} {list_structid list_link: ident}{list_token} - (ls: listspec list_structid list_link list_token) sh v p, - list_cell ls sh v p |-- !! field_compatible list_struct nil p. -Proof. -intros. -unfold list_cell. -normalize. -Qed. -#[export] Hint Resolve list_cell_local_facts : saturate_local. - -End LsegSpecial. - -Module Links. - -Section LIST2. -Context {cs: compspecs}. -Context {list_structid: ident} {list_link: ident}{list_token: share -> val -> mpred}. - -Definition vund (ls: listspec list_structid list_link list_token) : elemtype ls := - compact_prod_gen - (fun it => default_val (field_type (name_member it) list_fields)) (@all_but_link list_link list_fields). - -Definition lseg (ls: listspec list_structid list_link list_token) (dsh psh: share) - (contents: list val) (x z: val) : mpred := - LsegGeneral.lseg ls dsh psh (map (fun v => (v, vund ls)) contents) x z. - -Lemma nonreadable_list_cell_eq: - forall (ls: listspec list_structid list_link list_token) sh v v' p, - ~ readable_share sh -> - list_cell ls sh v p = list_cell ls sh v' p. -Proof. -unfold list_cell; intros. - destruct (field_compatible_dec list_struct nil p); - [ | solve [ apply pred_ext; normalize ]]. - f_equal. - revert v v'; unfold elemtype. - set (m := all_but_link list_fields). - assert (PLAIN: plain_members m = true). { - generalize list_plain. subst m. set (al := list_fields). - induction al; simpl; intros; auto. - destruct a; [ | discriminate]. - if_tac; auto. - } - clearbody m. - induction m; intros. - reflexivity. - destruct a as [i t|]; [ |discriminate]. - assert (field_compatible (field_type i list_fields) nil - (offset_val (field_offset cenv_cs i list_fields) p)) - by admit. (* need to adjust the induction hypothesis to prove this *) - destruct m as [ | [i' t'|]]; [ | | discriminate]. - + Opaque field_type field_offset. - clear IHm; simpl. - Transparent field_type field_offset. - rewrite !withspacer_spacer. - f_equal. - admit. (* apply nonreadable_data_at_rec_eq; auto. *) (* list_cell should be defined by field_at instead of data_at_rec. *) - + - rewrite !struct_pred_cons2. - rewrite !withspacer_spacer. - f_equal. f_equal. - * admit. (* unfold at_offset. apply nonreadable_data_at_rec_eq; auto.*) - * apply IHm. - simpl; auto. -Admitted. - -Lemma cell_share_join: - forall (ls: listspec list_structid list_link list_token) ash bsh psh p v, - sepalg.join ash bsh psh -> - list_cell ls ash v p * list_cell ls bsh v p = list_cell ls psh v p. -Proof. - intros. - unfold list_cell. - destruct (field_compatible_dec list_struct nil p); - [ | solve [ apply pred_ext; normalize ]]. - normalize. - f_equal. - revert v; unfold elemtype. - set (m := all_but_link list_fields). - assert (PLAIN: plain_members m = true). { - generalize list_plain. subst m. set (al := list_fields). - induction al; simpl; intros; auto. - destruct a; [ | discriminate]. - if_tac; auto. - } - clearbody m. - induction m; intros. - simpl. rewrite emp_sepcon; auto. - destruct a as [i t|]; [ | discriminate]. - assert (field_compatible (field_type i list_fields) nil - (offset_val (field_offset cenv_cs i list_fields) p)) - by admit. (* need to adjust the induction hypothesis to prove this *) - destruct m as [ | [i' t'|]]; [ | | discriminate]. - + - clear IHm; simpl. rewrite !withspacer_spacer. - rewrite <- sepcon_assoc. - match goal with |- ?A * ?B * ?C * ?D = _ => - pull_left C; pull_left A - end. - rewrite sepcon_assoc. f_equal. - unfold spacer. if_tac. rewrite emp_sepcon; auto. - unfold at_offset. - apply memory_block_share_join; auto. - unfold at_offset. - assert (isptr p) by (auto with field_compatible). - destruct p; try inversion H1. - apply data_at_rec_share_join; auto. - + - rewrite !struct_pred_cons2. - rewrite !withspacer_spacer. - match goal with |- (?A * ?B * ?C) * (?A' * ?B' * ?C') = _ => - transitivity ((A * A') * (B * B') * (C * C')) - end. - rewrite <- ! sepcon_assoc. - repeat match goal with |- _ * ?A = _ => pull_right A; f_equal end. - f_equal. f_equal. - unfold spacer. if_tac. apply sepcon_emp. - unfold at_offset. - apply memory_block_share_join; auto. - unfold at_offset. - assert (isptr p) by (auto with field_compatible). - destruct p; try inversion H1. - apply data_at_rec_share_join; auto. - apply IHm. auto. -Admitted. - -Lemma join_cell_link (ls: listspec list_structid list_link list_token): - forall v' ash bsh psh p v, - sepalg.join ash bsh psh -> - ~ (readable_share ash) -> - readable_share bsh -> - list_cell ls ash v' p * list_cell ls bsh v p = list_cell ls psh v p. - Proof. - intros. - rewrite (nonreadable_list_cell_eq _ _ v' v _ H0). - apply cell_share_join; auto. -Qed. - -Lemma lseg_unfold (ls: listspec list_structid list_link list_token): forall dsh psh contents v1 v2, - lseg ls dsh psh contents v1 v2 = - match contents with - | p::t => !! (p=v1 /\ ~ ptr_eq v1 v2) && EX tail: val, - !! is_pointer_or_null tail && - list_token dsh v1 * list_cell ls dsh (vund ls) v1 - * field_at psh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) tail) v1 - * lseg ls dsh psh t tail v2 - | nil => !! (ptr_eq v1 v2) && emp - end. -Proof. - intros. - unfold lseg. - rewrite LsegGeneral.lseg_unfold. - revert v1; induction contents; simpl; intros; auto. -Qed. - -Lemma lseg_eq (ls: listspec list_structid list_link list_token): - forall dsh psh l v , - is_pointer_or_null v -> - lseg ls dsh psh l v v = !!(l=nil) && emp. -Proof. -intros. -rewrite (lseg_unfold ls dsh psh l v v). -destruct l. -f_equal. f_equal. -apply prop_ext; split; intro; auto. -normalize. -apply pred_ext; -apply bi.pure_elim_l; intro. -destruct H0. -contradiction H1. -destruct v; inv H; try split; auto. -unfold Ptrofs.cmpu. apply Ptrofs.eq_true. -inv H0. -Qed. - -Definition lseg_cons (ls: listspec list_structid list_link list_token) dsh psh - (l: list val) (x z: val) : mpred := - !! (~ ptr_eq x z) && - EX h:(elemtype ls), EX r:list val, EX y:val, - !!(l=x::r /\ is_pointer_or_null y) && - list_token dsh x * list_cell ls dsh h x * - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x * - lseg ls dsh psh r y z. - -Lemma lseg_unroll (ls: listspec list_structid list_link list_token): forall dsh psh l x z , - ~ (readable_share dsh) -> - lseg ls dsh psh l x z = - (!! (ptr_eq x z) && !! (l=nil) && emp) || lseg_cons ls dsh psh l x z. -Proof. -intros. -rename H into NR. -rewrite lseg_unfold at 1. -apply pred_ext; destruct l. -apply bi.pure_elim_l; intros. -rewrite prop_true_andp by auto. -rewrite prop_true_andp by auto. -apply bi.or_intro_l; auto. -apply bi.or_intro_r. -unfold lseg_cons. -apply bi.pure_elim_l; intros. -destruct H. -subst x. -apply bi.exist_elim; intro tail. -rewrite (prop_true_andp (~ptr_eq v z)) by auto. -apply bi.exist_intro with (vund ls). -apply bi.exist_intro with l. -apply bi.exist_intro with tail. -normalize. - autorewrite with subst norm1 norm2; normalize. -apply orp_left. -rewrite andp_assoc; -do 2 (apply bi.pure_elim_l; intro). - rewrite prop_true_andp by auto. auto. -unfold lseg_cons. -apply bi.pure_elim_l; intros. -apply bi.exist_elim; intro h. -apply bi.exist_elim; intro r. -apply bi.exist_elim; intro y. -do 3 rewrite sepcon_andp_prop'. -apply bi.pure_elim_l; intros [? ?]. -inv H0. -apply orp_left. -rewrite andp_assoc; -do 2 (apply bi.pure_elim_l; intro). -inv H0. -unfold lseg_cons. -apply bi.pure_elim_l; intros. -apply bi.exist_elim; intro h. -apply bi.exist_elim; intro r. -apply bi.exist_elim; intro y. -do 3 rewrite sepcon_andp_prop'. -apply bi.pure_elim_l; intros [? ?]. -symmetry in H0; inv H0. - rewrite prop_true_andp by auto. -apply bi.exist_intro with y. -normalize. -repeat (apply sepcon_derives; auto). -clear - NR. -apply derives_refl'; apply nonreadable_list_cell_eq; auto. -Qed. - -Lemma lseg_unroll_nonempty1 (ls: listspec list_structid list_link list_token): - forall p P dsh psh h tail v1 v2, - ~ (readable_share dsh) -> - ~ ptr_eq v1 v2 -> - is_pointer_or_null p -> - (P |-- list_token dsh v1 * list_cell ls dsh h v1 * - (field_at psh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) p) v1 * - lseg ls dsh psh tail p v2)) -> - P |-- lseg ls dsh psh (v1::tail) v1 v2. -Proof. intros. rewrite lseg_unroll by auto. apply bi.or_intro_r. unfold lseg_cons. - rewrite prop_true_andp by auto. - apply bi.exist_intro with h. apply bi.exist_intro with tail. apply bi.exist_intro with p. - rewrite prop_true_andp by auto. - rewrite sepcon_assoc. - eapply derives_trans; [ eassumption | ]. - apply sepcon_derives; auto. -Qed. - -Lemma lseg_neq (ls: listspec list_structid list_link list_token): - forall dsh psh s v v2, - ~ (readable_share dsh) -> - ptr_neq v v2 -> - lseg ls dsh psh s v v2 = lseg_cons ls dsh psh s v v2. -intros. rewrite lseg_unroll by auto. -apply pred_ext. apply orp_left; auto. -rewrite andp_assoc. -do 2 (apply bi.pure_elim_l; intro). -congruence. -apply bi.or_intro_r. auto. -Qed. - -Lemma lseg_nonnull (ls: listspec list_structid list_link list_token): - forall dsh psh s v, - ~ (readable_share dsh) -> - typed_true (tptr list_struct) v -> - lseg ls dsh psh s v nullval = lseg_cons ls dsh psh s v nullval. -Proof. -intros. unfold nullval. -apply lseg_neq; auto. -unfold typed_true, strict_bool_val in H0; simpl in H0. -destruct Archi.ptr64 eqn:?; - destruct v; inv H0; - first [ revert H2; simple_if_tac; discriminate | intro Hx; inv Hx]. -Qed. - -Lemma unfold_lseg_neq (ls: listspec list_structid list_link list_token): - forall P Q1 Q R (v v2: val) dsh psh (s: list val), - ~ (readable_share dsh) -> - (PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s v v2 :: R))) |-- - !! (ptr_neq v v2)) -> - PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s v v2 :: R))) |-- - EX hryp: elemtype ls * list val * val * val, - match hryp with (h,r,y,p) => - !! (s=p::r /\ is_pointer_or_null y) && - !! (p=v) && - PROPx P (LOCALx Q - (SEPx (list_token dsh v :: list_cell ls dsh h v:: - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y v2 :: - R))) - end. -Proof. -intros. -apply derives_trans with -(PROPx P (LOCALx (Q1::Q) (SEPx (lseg_cons ls dsh psh s v v2 :: R)))). -apply derives_trans with -(!! (ptr_neq v v2) && PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s v v2 :: R)))). -apply andp_right; auto. -intro rho; unfold PROPx,LOCALx,SEPx,local,tc_expr,tc_lvalue; unfold_lift; simpl. -unfold lift1; simpl. - repeat (apply bi.pure_elim_l; intro). - rewrite prop_true_andp by auto. - rewrite prop_true_andp by auto. -apply sepcon_derives; auto. -rewrite lseg_neq; auto. -intro rho; unfold PROPx,LOCALx,SEPx,local,tc_expr,tc_lvalue,lift2,lift1,lift0; simpl. - unfold_lift. - unfold lseg_cons. simpl. - apply bi.pure_elim_l; intro. - apply bi.pure_elim_l; intros [? ?]. - rewrite sepcon_andp_prop'. - apply bi.pure_elim_l; intro. - rewrite exp_sepcon1; apply bi.exist_elim; intro h. - rewrite exp_sepcon1; apply bi.exist_elim; intro r. - rewrite exp_sepcon1; apply bi.exist_elim; intro y. - repeat rewrite sepcon_andp_prop'. - apply bi.pure_elim_l; intros [? ?]. - subst. - apply bi.exist_intro with (h,r,y, v). - repeat rewrite prop_true_andp by auto. - repeat rewrite sepcon_assoc. - auto. -Qed. - -Lemma unfold_lseg_cons (ls: listspec list_structid list_link list_token): - forall P Q1 Q R e dsh psh s, - ~ (readable_share dsh) -> - (PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s e nullval :: R))) |-- - !! (typed_true (tptr list_struct) e)) -> - PROPx P (LOCALx (Q1::Q) (SEPx (lseg ls dsh psh s e nullval :: R))) |-- - EX hryp: elemtype ls * list val * val * val, - match hryp with (h,r,y,p) => - !! (s=p::r /\ is_pointer_or_null y) && - !! (p = e) && - PROPx P (LOCALx Q - (SEPx (list_token dsh e :: list_cell ls dsh h e :: - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) e :: - lseg ls dsh psh r y nullval :: - R))) - end. -Proof. -intros. apply unfold_lseg_neq; auto. -eapply derives_trans. -apply H0. normalize. -unfold local. super_unfold_lift. -unfold nullval. destruct e; inv H1; try congruence; auto. -intro. apply ptr_eq_e in H1. -destruct Archi.ptr64; inv H1. -Qed. - -Lemma semax_lseg_neq (ls: listspec list_structid list_link list_token): - forall (Espec: OracleKind) - Delta P Q dsh psh s v v2 R c Post, - ~ (readable_share dsh) -> - ~ (ptr_eq v v2) -> - (forall (h: elemtype ls) (r: list val) (y: val), - s=v::r -> is_pointer_or_null y -> - semax Delta - (PROPx P (LOCALx Q - (SEPx (list_token dsh v :: list_cell ls dsh h v :: - field_at psh list_struct (StructField list_link :: nil) - (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y v2 :: - R)))) c Post) -> - semax Delta - (PROPx P (LOCALx Q (SEPx (lseg ls dsh psh s v v2 :: R)))) - c Post. -Proof. -intros. -rewrite lseg_neq by auto. -unfold lseg_cons. -apply semax_pre0 with - (EX h: elemtype ls, EX r: list val, EX y: val, - !!(s = v :: r /\ is_pointer_or_null y) && - PROPx P (LOCALx Q (SEPx (list_token dsh v :: list_cell ls dsh h v :: - field_at psh list_struct (StructField list_link :: nil) - (valinject - (nested_field_type list_struct - (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y v2 :: R)))). -go_lowerx; entailer. -Exists h r y. -rewrite <- ?sepcon_assoc. -normalize. - autorewrite with subst norm1 norm2; normalize. -Intros h r y. -apply semax_extract_prop; intros [? ?]. -eauto. -Qed. - - -Lemma semax_lseg_nonnull (ls: listspec list_structid list_link list_token): - forall (Espec: OracleKind) - Delta P Q dsh psh s v R c Post, - ~ (readable_share dsh) -> - ENTAIL Delta, PROPx P (LOCALx Q - (SEPx (lseg ls dsh psh s v nullval :: R))) |-- - !!(typed_true (tptr list_struct) v) -> - (forall (h: elemtype ls) (r: list val) (y: val), - s=v::r -> is_pointer_or_null y -> - semax Delta - (PROPx P (LOCALx Q - (SEPx (list_token dsh v :: list_cell ls dsh h v :: - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) v :: - lseg ls dsh psh r y nullval :: - R)))) c Post) -> - semax Delta - (PROPx P (LOCALx Q (SEPx (lseg ls dsh psh s v nullval :: R)))) - c Post. -Proof. -intros. -assert_PROP (~ ptr_eq v nullval). -eapply derives_trans; [eapply H0 |]. -normalize. -apply semax_lseg_neq; auto. -Qed. - -Lemma lseg_nil_eq (ls: listspec list_structid list_link list_token): - forall dsh psh p q, - lseg ls dsh psh nil p q = !! (ptr_eq p q) && emp. -Proof. intros. - reflexivity. -Qed. - -Lemma lseg_cons_eq (ls: listspec list_structid list_link list_token): - forall dsh psh h r x z , - ~ (readable_share dsh) -> - lseg ls dsh psh (h::r) x z = - !!(x = h /\ ~ ptr_eq x z) && - (EX y : val, - !!(is_pointer_or_null y) && - list_token dsh x * list_cell ls dsh (vund ls) x * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x * - lseg ls dsh psh r y z). -Proof. - intros. rewrite lseg_unroll by auto. - apply pred_ext. - apply orp_left. - rewrite andp_assoc. - apply bi.pure_elim_l; intro. - apply bi.pure_elim_l; intro. - inv H1. - unfold lseg_cons. - normalize. - symmetry in H1; inv H1. - apply bi.exist_intro with y. normalize. - autorewrite with subst norm1 norm2; normalize. - repeat (apply sepcon_derives; auto). - apply derives_refl'; apply nonreadable_list_cell_eq; auto. - apply bi.or_intro_r. - normalize. - unfold lseg_cons. - rewrite prop_true_andp by auto. - apply bi.exist_intro with (vund ls). apply bi.exist_intro with r. apply bi.exist_intro with y. - normalize. - autorewrite with subst norm1 norm2; normalize. -Qed. - -Definition lseg_cons_right (ls: listspec list_structid list_link list_token) - dsh psh (l: list val) (x z: val) : mpred := - !! (~ ptr_eq x z) && - EX r:list val , EX y:val, - !!(l=r++y::nil /\ is_pointer_or_null y) && - list_cell ls dsh (vund ls) y * - field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls dsh psh r x y. - -Lemma lseg_cons_right_neq (ls: listspec list_structid list_link list_token): - forall dsh psh l x h y w z, - sepalg.nonidentity psh -> - ~ (readable_share dsh) -> - list_token dsh y * list_cell ls dsh h y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls dsh psh l x y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) w) z - |-- lseg ls dsh psh (l++y::nil) x z * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) w) z. -Proof. -intros. rename H into SH. rename H0 into NR. -assert (SZ: 0 < sizeof (nested_field_type list_struct (DOT list_link))) - by (rewrite list_link_type; unfold sizeof; simpl; destruct Archi.ptr64; computable). -rewrite (field_at_isptr _ _ _ _ z). -normalize. -revert x; induction l; simpl; intros. -* -unfold lseg. -simpl. -normalize. - autorewrite with subst norm1 norm2; normalize. -apply bi.exist_intro with z. -entailer. - apply derives_refl'; f_equal. f_equal. f_equal. - apply (nonreadable_list_cell_eq); auto. -* -unfold lseg; simpl. -normalize. -apply bi.exist_intro with x0. -rewrite <- ?sepcon_assoc. -normalize. - autorewrite with subst norm1 norm2; normalize. -specialize (IHl x0). -entailer. -pull_right (list_token dsh x); pull_right (list_cell ls dsh (vund ls) x). -apply sepcon_derives; auto. -apply sepcon_derives; auto. -pull_right (field_at psh list_struct (StructField list_link :: nil) - (valinject - (nested_field_type list_struct (StructField list_link :: nil)) x0) - x). -apply sepcon_derives; auto. -Qed. - -Lemma lseg_cons_right_null (ls: listspec list_structid list_link list_token): forall dsh psh l x h y, - ~ (readable_share dsh) -> - list_token dsh y * list_cell ls dsh h y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) nullval) y * - lseg ls dsh psh l x y - |-- lseg ls dsh psh (l++y::nil) x nullval. -Proof. -intros. rename H into NR. -unfold lseg. -revert x; induction l; simpl; intros. -* -normalize. - autorewrite with subst norm1 norm2; normalize. -apply bi.exist_intro with nullval. -apply andp_right. -apply not_prop_right; intro. -apply ptr_eq_e in H. subst y. -entailer!. -destruct H. contradiction H. -rewrite prop_true_andp by reflexivity. -rewrite prop_true_andp by apply ptr_eq_nullval. -normalize. -apply derives_refl'; f_equal. f_equal. -apply nonreadable_list_cell_eq; auto. -* -normalize. -apply bi.exist_intro with x0. -normalize. - autorewrite with subst norm1 norm2; normalize. -specialize (IHl x0). -apply andp_right. -rewrite prop_and. -apply andp_right; [ | apply prop_right; auto]. -apply not_prop_right; intro. -apply ptr_eq_e in H0. subst x. -entailer. -destruct H2; contradiction H2. -eapply derives_trans. -2: apply sepcon_derives; [ | eassumption]; apply derives_refl. -clear IHl. -cancel. -Qed. - - -Lemma lseg_cons_right_list (ls: listspec list_structid list_link list_token): - forall dsh psh l l' x h y z, - sepalg.nonidentity psh -> - ~ (readable_share dsh) -> - list_token dsh y * list_cell ls dsh h y * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) z) y * - lseg ls dsh psh l x y * lseg ls dsh psh l' z nullval - |-- lseg ls dsh psh (l++y::nil) x z * lseg ls dsh psh l' z nullval. -Proof. -intros. -destruct l'. -rewrite lseg_nil_eq. -normalize. -rewrite prop_true_andp by apply ptr_eq_nullval. -apply lseg_cons_right_null; auto. -rewrite lseg_cons_eq; auto. -Intros u. Exists u. subst. -rewrite !prop_true_andp by auto. -rewrite <- !sepcon_assoc. -apply sepcon_derives; auto. -pull_right (list_cell ls dsh (vund ls) v). -apply sepcon_derives; auto. -pull_right (list_token dsh v). -apply sepcon_derives; auto. -apply lseg_cons_right_neq; auto. -Qed. - -Lemma lseg_unroll_right (ls: listspec list_structid list_link list_token): forall sh sh' l x z , - lseg ls sh sh' l x z = (!! (ptr_eq x z) && !! (l=nil) && emp) || lseg_cons_right ls sh sh' l x z. -Abort. (* not likely true *) - -Lemma lseg_local_facts: - forall ls dsh psh contents p q, - lseg ls dsh psh contents p q |-- - !! (is_pointer_or_null p /\ (p=q <-> contents=nil)). -Proof. -intros. -rewrite lseg_unfold. -destruct contents. -apply bi.pure_elim_l; intro. -unfold ptr_eq in H. -apply prop_right. -destruct p; try contradiction; simpl; auto. -destruct q; try contradiction; auto. -destruct H as [? [? ?]]. rewrite H. -unfold Int.cmpu in *. -apply int_eq_e in H0. -apply int_eq_e in H1. subst. -split; auto; split; auto. -destruct q; try contradiction; auto. -destruct H as [? [? ?]]. rewrite H. -unfold Int64.cmpu in *. -apply int64_eq_e in H0. -apply int64_eq_e in H1. subst. -split; auto; split; auto. -destruct q; try contradiction; auto. -destruct H; subst. -unfold Ptrofs.cmpu in *. -apply ptrofs_eq_e in H0. subst. -intuition. -normalize. -rewrite field_at_isptr. -normalize. - autorewrite with subst norm1 norm2; normalize. -apply prop_right. -split. intro; subst q. -contradiction H. normalize. -intros. discriminate. -Qed. - -Definition lseg_cell (ls: listspec list_structid list_link list_token) - (dsh psh : share) - (v: elemtype ls) (x y: val) := - !!is_pointer_or_null y && list_token dsh x * list_cell ls dsh v x * field_at psh list_struct (StructField list_link :: nil) (valinject (nested_field_type list_struct (StructField list_link :: nil)) y) x. - -Lemma lseg_cons_eq2: forall - (ls : listspec list_structid list_link list_token) (dsh psh : share) (h : elemtype ls) - (r : list val ) (x z : val), - ~ (readable_share dsh) -> - lseg ls dsh psh (x :: r) x z = - !!(~ ptr_eq x z) && (EX y : val, lseg_cell ls dsh psh h x y * lseg ls dsh psh r y z). -Proof. - intros. - rewrite -> lseg_cons_eq by auto. - unfold lseg_cell. - normalize. - autorewrite with subst norm1 norm2; normalize. - f_equal. extensionality y. - f_equal. f_equal. f_equal. f_equal. - apply nonreadable_list_cell_eq; auto. -Qed. - -Lemma list_append: forall {dsh psh: share} - {ls : listspec list_structid list_link list_token} (hd mid tl:val) ct1 ct2 P, - (forall tl', lseg_cell ls dsh psh (vund ls) tl tl' * P tl |-- FF) -> - (lseg ls dsh psh ct1 hd mid) * lseg ls dsh psh ct2 mid tl * P tl|-- - (lseg ls dsh psh (ct1 ++ ct2) hd tl) * P tl. -Proof. - intros. - unfold lseg. - revert hd; induction ct1; simpl; intros; auto. -* - normalize. -* - normalize. - progress (autorewrite with subst norm1 norm2); normalize. - apply bi.exist_intro with y. - apply andp_right. - + - apply not_prop_right; intro. apply ptr_eq_e in H1; subst hd. - clear IHct1. - specialize (H y). - unfold lseg_cell in H. - rewrite prop_true_andp in H by auto. - change (LsegGeneral.lseg ls dsh psh (map (fun v : val => (v, vund ls)) ct1)) - with (lseg ls dsh psh ct1). - change (LsegGeneral.lseg ls dsh psh (map (fun v : val => (v, vund ls)) ct2)) - with (lseg ls dsh psh ct2). - apply derives_trans with - (lseg ls dsh psh ct1 y mid * lseg ls dsh psh ct2 mid tl * FF). - cancel. auto. - rewrite sepcon_FF; auto. - + - normalize. - specialize (IHct1 y). clear H. - do 2 rewrite sepcon_assoc. - eapply derives_trans. - apply sepcon_derives. - apply derives_refl. - rewrite <- !sepcon_assoc; eassumption. - cancel. -Qed. - -Lemma list_append_null: - forall - (ls: listspec list_structid list_link list_token) - (dsh psh: share) - (hd mid: val) ct1 ct2, - lseg ls dsh psh ct1 hd mid * lseg ls dsh psh ct2 mid nullval |-- - lseg ls dsh psh (ct1++ct2) hd nullval. -Proof. -intros. - rewrite <- sepcon_emp. - eapply derives_trans; [ | apply (list_append hd mid nullval ct1 ct2 (fun _ => emp))]. - normalize. - intros. - unfold lseg_cell. simpl. saturate_local. destruct H. contradiction H. -Qed. - -Lemma list_cell_valid_pointer: - forall (LS: listspec list_structid list_link list_token) (dsh psh: Share.t) v p, - sepalg.nonidentity dsh -> - sepalg.join_sub dsh psh -> - field_offset cenv_cs list_link list_fields + sizeof (field_type list_link list_fields) - = field_offset_next cenv_cs list_link list_fields (co_sizeof (get_co list_structid)) -> - list_cell LS dsh v p * field_at_ psh list_struct (StructField list_link::nil) p - |-- valid_pointer p. -Proof. - intros ? ? ? ? ? NON_ID ? ?. - destruct H as [bsh ?]. - rewrite <- (field_at__share_join _ _ _ _ _ _ H). - rewrite <- sepcon_assoc. - rewrite list_cell_link_join_nospacer; auto. - apply sepcon_valid_pointer1. - unfold data_at_, field_at_, data_at. - eapply derives_trans; [ apply field_at_valid_ptr; auto | ]. - change (nested_field_type list_struct nil) with list_struct. - apply LsegGeneral.sizeof_list_struct_pos. - unfold field_address. - if_tac; auto. - change (Int.repr (nested_field_offset list_struct nil)) with Int.zero. - rewrite valid_pointer_offset_val_zero; auto. - simpl. - change predicates_hered.FF with FF. apply FF_left. -Qed. - -Lemma list_cell_valid_pointerx: - forall (ls : listspec list_structid list_link list_token) sh v p, - sh <> Share.bot -> - list_cell ls sh v p |-- valid_pointer p. -Proof. - intros. - unfold list_cell. -Abort. (* probably not true; would be true with a direct (non-magic-wand) - definition of list_cell *) - -Lemma lseg_valid_pointer: - forall (ls : listspec list_structid list_link list_token) dsh psh contents p q R, - sepalg.nonidentity dsh -> - dsh <> Share.bot -> - sepalg.join_sub dsh psh -> - field_offset cenv_cs list_link list_fields + sizeof (field_type list_link list_fields) - = field_offset_next cenv_cs list_link list_fields (co_sizeof (get_co list_structid)) -> - (R |-- valid_pointer q) -> - R * lseg ls dsh psh contents p q |-- valid_pointer p. -Proof. -intros. -destruct contents. -rewrite lseg_nil_eq. normalize. -unfold lseg; simpl. -normalize. -apply sepcon_valid_pointer2. -rewrite !sepcon_assoc. -apply sepcon_valid_pointer2. -rewrite <- !sepcon_assoc. -apply sepcon_valid_pointer1. -eapply derives_trans with - (list_cell ls dsh (vund ls) p * field_at_ psh list_struct (StructField list_link :: nil) p). -cancel. -apply list_cell_valid_pointer; auto. -Qed. - -End LIST2. - -Lemma join_sub_Tsh: - forall sh, sepalg.join_sub sh Tsh. -Admitted. (* easy *) -#[export] Hint Resolve join_sub_Tsh: valid_pointer. - -#[export] Hint Rewrite @lseg_nil_eq : norm. - -#[export] Hint Rewrite @lseg_eq using reflexivity: norm. - -#[export] Hint Resolve lseg_local_facts : saturate_local. - -#[export] Hint Resolve denote_tc_test_eq_split : valid_pointer. - -Ltac resolve_lseg_valid_pointer := -match goal with - | |- ?Q |-- valid_pointer ?p => - match Q with context [lseg ?A ?B ?C ?D p ?q] => - repeat rewrite <- sepcon_assoc; - pull_right (lseg A B C D p q); - apply lseg_valid_pointer; [auto | | | reflexivity | ]; - auto 50 with valid_pointer - end - end. - -#[export] Hint Extern 10 (_ |-- valid_pointer _) => - resolve_lseg_valid_pointer : valid_pointer. - -Ltac resolve_list_cell_valid_pointer := - match goal with |- ?A |-- valid_pointer ?p => - match A with context [@list_cell ?cs ?sid ?lid ?tok ?LS ?dsh ?v p] => - match A with context [field_at ?psh ?t (StructField lid::nil) ?v' p] => - apply derives_trans with - (@list_cell cs sid lid tok LS dsh v p * - field_at_ psh t (StructField lid::nil) p * TT); - [cancel - | apply sepcon_valid_pointer1; - apply list_cell_valid_pointer; [auto | | reflexivity]; auto with valid_pointer] - end - end - end. - -#[export] Hint Extern 10 (_ |-- valid_pointer _) => - resolve_list_cell_valid_pointer : valid_pointer. - -End Links. - -Arguments elemtype {cs} {list_structid} {list_link} {list_token} ls / .