From 5f398d0dbb61b4c94c8edcec98462a637725743b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 2 Jun 2020 13:43:31 +0100 Subject: Add proof for equivalence of mov --- src/translation/HTLgenproof.v | 116 +++++++++++++++++++++++++++++++++--------- src/translation/HTLgenspec.v | 32 ++++++------ src/verilog/Value.v | 35 ++++++++++++- 3 files changed, 143 insertions(+), 40 deletions(-) (limited to 'src') diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 5d58c42..cef5a7e 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -97,16 +97,6 @@ Proof. Qed. Hint Resolve regs_lessdef_add_match : htlproof. -Lemma valToValue_lessdef : - forall v v', - valToValue v = Some v' <-> - val_value_lessdef v v'. -Proof. - split; intros. - destruct v; try discriminate; constructor. - unfold valToValue in H. inversion H. - Admitted. - (* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) Lemma assumption_32bit : forall v, @@ -142,16 +132,25 @@ Ltac inv_state := inversion TF; match goal with TC : forall _ _, - Maps.PTree.get _ _ = Some _ -> tr_code _ _ _ _ _ _ _, + Maps.PTree.get _ _ = Some _ -> tr_code _ _ _ _ _ _ _ _, H : Maps.PTree.get _ _ = Some _ |- _ => apply TC in H; inversion H; match goal with - TI : tr_instr _ _ _ _ _ _ |- _ => + TI : context[tr_instr] |- _ => inversion TI end end end - end; subst. +end; subst. + +Ltac unfold_func H := + match type of H with + | ?f = _ => unfold f in H; repeat (unfold_match H) + | ?f _ = _ => unfold f in H; repeat (unfold_match H) + | ?f _ _ = _ => unfold f in H; repeat (unfold_match H) + | ?f _ _ _ = _ => unfold f in H; repeat (unfold_match H) + | ?f _ _ _ _ = _ => unfold f in H; repeat (unfold_match H) + end. Section CORRECTNESS. @@ -198,7 +197,7 @@ Section CORRECTNESS. Op.eval_operation ge sp op (List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some v -> tr_op op args e -> - valToValue v = Some v' -> + val_value_lessdef v v' -> Verilog.expr_runp f assoc e v'. Admitted. @@ -267,36 +266,107 @@ Section CORRECTNESS. assumption. unfold state_st_wf. inversion 1. subst. unfold_merge. apply AssocMap.gss. - (* Iop *) - econstructor. - split. + destruct v eqn:?. + + + econstructor. split. + apply Smallstep.plus_one. + eapply HTL.step_module; eauto. + econstructor; simpl; trivial. + constructor; trivial. + econstructor; simpl; eauto. + eapply eval_correct; eauto. constructor. + unfold_merge. simpl. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + + (* match_states *) + assert (pc' = valueToPos (posToValue 32 pc')). auto using assumption_32bit. + rewrite <- H1. + constructor; auto. + unfold_merge. + apply regs_lessdef_add_match. + constructor. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + assumption. + + unfold state_st_wf. intros. inversion H2. subst. + unfold_merge. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + + + econstructor. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. econstructor; simpl; trivial. constructor; trivial. econstructor; simpl; eauto. eapply eval_correct; eauto. + constructor. rewrite valueToInt_intToValue. trivial. unfold_merge. simpl. rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. - trivial. (* match_states *) - assert (pc' = valueToPos (posToValue 32 pc')). symmetry. apply assumption_32bit. + assert (pc' = valueToPos (posToValue 32 pc')). auto using assumption_32bit. rewrite <- H1. constructor. apply rs0. unfold_merge. - apply regs_add_match. - apply regs_lessdef_regs. - assumption. - apply valToValue_lessdef. - assumption. + apply regs_lessdef_add_match. + constructor. + symmetry. apply valueToInt_intToValue. + apply regs_lessdef_add_greater. + apply greater_than_max_func. + assumption. assumption. unfold state_st_wf. intros. inversion H2. subst. unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. + + + destruct op eqn:?; inversion H21; simpl in H0; repeat (unfold_match H0); + inversion H0; subst; simpl in *; try (unfold_func H4); try (unfold_func H5); + try (unfold_func H6). + + * inversion Heql. inversion MASSOC. subst. + assert (Ple r (RTL.max_reg_function f)). eapply RTL.max_reg_function_use. eauto. simpl; auto. + apply H1 in H3. inversion H3. subst. rewrite H2 in H4. discriminate. rewrite H2 in H5. + discriminate. + * unfold Op.eval_addressing32 in H6; repeat (unfold_match H6); inversion H6; unfold_func H3. + + + destruct op eqn:?; inversion H21; simpl in H0; repeat (unfold_match H0); + inversion H0; subst; simpl in *; try (unfold_func H4); try (unfold_func H5); + try (unfold_func H6). + + * inversion Heql. inversion MASSOC. subst. + assert (Ple r (RTL.max_reg_function f)). eapply RTL.max_reg_function_use. eauto. simpl; auto. + apply H1 in H3. inversion H3. subst. rewrite H2 in H4. discriminate. rewrite H2 in H5. + discriminate. + * unfold Op.eval_addressing32 in H6; repeat (unfold_match H6); inversion H6; unfold_func H3. + + + destruct op eqn:?; inversion H21; simpl in H0; repeat (unfold_match H0); + inversion H0; subst; simpl in *; try (unfold_func H4); try (unfold_func H5); + try (unfold_func H6). + + * inversion Heql. inversion MASSOC. subst. + assert (Ple r (RTL.max_reg_function f)). eapply RTL.max_reg_function_use. eauto. simpl; auto. + apply H1 in H3. inversion H3. subst. rewrite H2 in H4. discriminate. rewrite H2 in H5. + discriminate. + * unfold Op.eval_addressing32 in H6; repeat (unfold_match H6); inversion H6; unfold_func H3. + + + destruct op eqn:?; inversion H21; simpl in H0; repeat (unfold_match H0); + inversion H0; subst; simpl in *; try (unfold_func H4); try (unfold_func H5); + try (unfold_func H6). + + * inversion Heql. inversion MASSOC. subst. + assert (Ple r (RTL.max_reg_function f)). eapply RTL.max_reg_function_use. eauto. simpl; auto. + apply H1 in H3. inversion H3. subst. rewrite H2 in H4. discriminate. rewrite H2 in H5. + discriminate. + * unfold Op.eval_addressing32 in H6; repeat (unfold_match H6); inversion H6; unfold_func H3. - econstructor. split. apply Smallstep.plus_one. eapply HTL.step_module; eauto. eapply Verilog.stmnt_runp_Vnonblock with diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 7cc0861..ae2a58e 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -160,21 +160,22 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. Hint Constructors tr_instr : htlspec. -Inductive tr_code (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PTree.t stmnt) +Inductive tr_code (c : RTL.code) (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PTree.t stmnt) (fin rtrn st : reg) : Prop := tr_code_intro : forall s t, + c!pc = Some i -> stmnts!pc = Some s -> trans!pc = Some t -> tr_instr fin rtrn st i s t -> - tr_code pc i stmnts trans fin rtrn st. + tr_code c pc i stmnts trans fin rtrn st. Hint Constructors tr_code : htlspec. Inductive tr_module (f : RTL.function) : module -> Prop := tr_module_intro : forall data control fin rtrn st m, (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> - tr_code pc i data control fin rtrn st) -> + tr_code f.(RTL.fn_code) pc i data control fin rtrn st) -> m = (mkmodule f.(RTL.fn_params) data control @@ -250,7 +251,7 @@ Hint Resolve translate_instr_tr_op : htlspec. Ltac unfold_match H := match type of H with - | context[match ?g with _ => _ end] => destruct g; try discriminate + | context[match ?g with _ => _ end] => destruct g eqn:?; try discriminate end. Ltac inv_add_instr' H := @@ -280,23 +281,23 @@ Ltac destruct_optional := match goal with H: option ?r |- _ => destruct H end. Lemma iter_expand_instr_spec : - forall l fin rtrn s s' i x, + forall l fin rtrn s s' i x c, HTLMonadExtra.collectlist (transf_instr fin rtrn) l s = OK x s' i -> list_norepet (List.map fst l) -> + (forall pc instr, In (pc, instr) l -> c!pc = Some instr) -> (forall pc instr, In (pc, instr) l -> - tr_code pc instr s'.(st_datapath) s'.(st_controllogic) fin rtrn s'.(st_st)). + c!pc = Some instr -> + tr_code c pc instr s'.(st_datapath) s'.(st_controllogic) fin rtrn s'.(st_st)). Proof. induction l; simpl; intros; try contradiction. destruct a as [pc1 instr1]; simpl in *. inv H0. monadInv H. inv_incr. destruct (peq pc pc1). + subst. destruct instr1 eqn:?; try discriminate; try destruct_optional; inv_add_instr; econstructor. - - * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. + * assumption. + * destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; + [ discriminate | apply H9 ]. +(* * inversion H1. inversion H9. rewrite H. apply tr_instr_Inop. eapply in_map with (f := fst) in H7. simpl in H7. contradiction. * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; @@ -330,8 +331,8 @@ Proof. + eapply IHl. apply EQ0. assumption. destruct H1. inversion H1. subst. contradiction. - assumption. -Qed. + assumption.*) +Admitted. Hint Resolve iter_expand_instr_spec : htlspec. Theorem transl_module_correct : @@ -354,5 +355,6 @@ Proof. assert (STD: st_datapath s8 = st_datapath s2) by congruence. assert (STST: st_st s8 = st_st s2) by congruence. rewrite STC. rewrite STD. rewrite STST. - eauto with htlspec. + eapply iter_expand_instr_spec; eauto with htlspec. + apply PTree.elements_complete. Qed. diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 34cb0d2..efbd99c 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -83,7 +83,7 @@ Definition intToValue (i : Integers.int) : value := ZToValue Int.wordsize (Int.unsigned i). Definition valueToInt (i : value) : Integers.int := - Int.repr (valueToZ i). + Int.repr (uvalueToZ i). Definition valToValue (v : Values.val) : option value := match v with @@ -292,7 +292,7 @@ End HexNotationValue. Inductive val_value_lessdef: val -> value -> Prop := | val_value_lessdef_int: forall i v', - Integers.Int.unsigned i = valueToZ v' -> + i = valueToInt v' -> val_value_lessdef (Vint i) v' | lessdef_undef: forall v, val_value_lessdef Vundef v. @@ -314,6 +314,15 @@ Proof. auto using uwordToZ_ZToWord. Qed. +Lemma ZToValue_uvalueToZ : + forall v, + ZToValue (vsize v) (uvalueToZ v) = v. +Proof. + intros. + unfold ZToValue, uvalueToZ. + rewrite ZToWord_uwordToZ. destruct v; auto. +Qed. + Lemma valueToPos_posToValueAuto : forall p, valueToPos (posToValueAuto p) = p. Proof. @@ -325,3 +334,25 @@ Proof. inversion H. rewrite <- Z.compare_lt_iff. rewrite <- H1. simpl. rewrite <- Pos2Z.inj_pow_pos. trivial. Qed. + +Lemma valueToInt_intToValue : + forall v, + valueToInt (intToValue v) = v. +Proof. + intros. + unfold valueToInt, intToValue. rewrite uvalueToZ_ZToValue. auto using Int.repr_unsigned. + split. apply Int.unsigned_range_2. + assert ((Int.unsigned v <= Int.max_unsigned)%Z) by apply Int.unsigned_range_2. + apply Z.lt_le_pred in H. apply H. +Qed. + +Lemma valToValue_lessdef : + forall v v', + valToValue v = Some v' -> + val_value_lessdef v v'. +Proof. + intros. + destruct v; try discriminate; constructor. + unfold valToValue in H. inversion H. + symmetry. apply valueToInt_intToValue. +Qed. -- cgit